├── LICENSE ├── README.md ├── benchmarks ├── fmls_k.tar └── fmls_s4.tar ├── leanpkg.toml └── src ├── K ├── README.md ├── full_language.lean ├── jump.lean ├── marking.lean ├── ops.lean ├── rules.lean ├── semantics.lean ├── size.lean ├── tableau.lean └── test.lean ├── KT ├── KT_defs.lean ├── ops.lean ├── rules.lean ├── semantics.lean ├── size.lean └── vanilla.lean ├── S4 ├── S4_defs.lean ├── data.lean ├── full_language.lean ├── ops.lean ├── rules.lean ├── semantics.lean ├── size.lean ├── tableau.lean ├── test.lean └── vanilla.lean ├── apps ├── examples.lean ├── topo_translation.lean └── topological_semantics.lean └── defs.lean /LICENSE: -------------------------------------------------------------------------------- 1 | MIT License 2 | 3 | Copyright (c) 2018 Minchao Wu 4 | 5 | Permission is hereby granted, free of charge, to any person obtaining a copy 6 | of this software and associated documentation files (the "Software"), to deal 7 | in the Software without restriction, including without limitation the rights 8 | to use, copy, modify, merge, publish, distribute, sublicense, and/or sell 9 | copies of the Software, and to permit persons to whom the Software is 10 | furnished to do so, subject to the following conditions: 11 | 12 | The above copyright notice and this permission notice shall be included in all 13 | copies or substantial portions of the Software. 14 | 15 | THE SOFTWARE IS PROVIDED "AS IS", WITHOUT WARRANTY OF ANY KIND, EXPRESS OR 16 | IMPLIED, INCLUDING BUT NOT LIMITED TO THE WARRANTIES OF MERCHANTABILITY, 17 | FITNESS FOR A PARTICULAR PURPOSE AND NONINFRINGEMENT. IN NO EVENT SHALL THE 18 | AUTHORS OR COPYRIGHT HOLDERS BE LIABLE FOR ANY CLAIM, DAMAGES OR OTHER 19 | LIABILITY, WHETHER IN AN ACTION OF CONTRACT, TORT OR OTHERWISE, ARISING FROM, 20 | OUT OF OR IN CONNECTION WITH THE SOFTWARE OR THE USE OR OTHER DEALINGS IN THE 21 | SOFTWARE. 22 | -------------------------------------------------------------------------------- /README.md: -------------------------------------------------------------------------------- 1 | # Verified decision procedures for modal logics 2 | 3 | The proofs compile after commit [5173c54](https://github.com/minchaowu/ModalTab/commit/5173c548ac9710d7ec9b259e82c76a1208cf4a41). For neater proof scripts and enhancement of efficiency, please follow the recent commits. 4 | 5 | How to compile: 6 | 7 | + Install [Lean](https://github.com/leanprover-community/lean) using [elan](https://github.com/Kha/elan) or following the [build instructions](https://github.com/leanprover-community/lean/blob/master/doc/make/index.md). 8 | 9 | + Run `leanpkg build` in the root of this repository. Several lines of tests will be printed, but there should be no errors reported by Lean. 10 | 11 | For each implementation, `is_sat` is the function to run. For K, it is in `jump.lean`. For KT and S4, they are in `vanilla.lean`. Alternatively, `tableau` in the same file can be used to inspect the model returned. 12 | -------------------------------------------------------------------------------- /benchmarks/fmls_k.tar: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/minchaowu/ModalTab/9bb0bf17faf0554d907ef7bdd639648742889178/benchmarks/fmls_k.tar -------------------------------------------------------------------------------- /benchmarks/fmls_s4.tar: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/minchaowu/ModalTab/9bb0bf17faf0554d907ef7bdd639648742889178/benchmarks/fmls_s4.tar -------------------------------------------------------------------------------- /leanpkg.toml: -------------------------------------------------------------------------------- 1 | [package] 2 | name = "ModalTab" 3 | version = "0.1" 4 | lean_version = "leanprover-community/lean:3.16.2" 5 | path = "src" 6 | 7 | [dependencies] 8 | mathlib = {git = "https://github.com/leanprover-community/mathlib", rev = "938b73a4cc852b643a31334cf89df4f29ac0784a"} 9 | -------------------------------------------------------------------------------- /src/K/README.md: -------------------------------------------------------------------------------- 1 | # Modal Tableau for K 2 | 3 | A verified proof-producing tableau prover with backjumping for modal logic K. -------------------------------------------------------------------------------- /src/K/full_language.lean: -------------------------------------------------------------------------------- 1 | import .semantics 2 | 3 | inductive fml 4 | | var (n : nat) 5 | | neg (φ : fml) 6 | | and (φ ψ : fml) 7 | | or (φ ψ : fml) 8 | | impl (φ ψ : fml) 9 | | box (φ : fml) 10 | | dia (φ : fml) 11 | 12 | open fml 13 | 14 | @[simp] def fml_size : fml → ℕ 15 | | (var n) := 1 16 | | (neg φ) := fml_size φ + 1 17 | | (and φ ψ) := fml_size φ + fml_size ψ + 1 18 | | (or φ ψ) := fml_size φ + fml_size ψ + 1 19 | | (impl φ ψ) := fml_size φ + fml_size ψ + 2 20 | | (box φ) := fml_size φ + 1 21 | | (dia φ) := fml_size φ + 1 22 | 23 | @[simp] def fml_force {states : Type} (k : kripke states) : states → fml → Prop 24 | | s (var n) := k.val n s 25 | | s (neg φ) := ¬ fml_force s φ 26 | | s (and φ ψ) := fml_force s φ ∧ fml_force s ψ 27 | | s (or φ ψ) := fml_force s φ ∨ fml_force s ψ 28 | | s (impl φ ψ) := ¬ fml_force s φ ∨ fml_force s ψ 29 | | s (box φ) := ∀ s', k.rel s s' → fml_force s' φ 30 | | s (dia φ) := ∃ s', k.rel s s' ∧ fml_force s' φ 31 | 32 | @[simp] def fml.to_nnf : fml → nnf 33 | | (var n) := nnf.var n 34 | | (neg (var n)) := nnf.neg n 35 | | (neg (neg φ)) := fml.to_nnf φ 36 | | (neg (and φ ψ)) := nnf.or (fml.to_nnf (neg φ)) (fml.to_nnf (neg ψ)) 37 | | (neg (or φ ψ)) := nnf.and (fml.to_nnf (neg φ)) (fml.to_nnf (neg ψ)) 38 | | (neg (impl φ ψ)) := nnf.and (fml.to_nnf φ) (fml.to_nnf (neg ψ)) 39 | | (neg (box φ)) := nnf.dia (fml.to_nnf (neg φ)) 40 | | (neg (dia φ)) := nnf.box (fml.to_nnf (neg φ)) 41 | | (and φ ψ) := nnf.and (fml.to_nnf φ) (fml.to_nnf ψ) 42 | | (or φ ψ) := nnf.or (fml.to_nnf φ) (fml.to_nnf ψ) 43 | | (impl φ ψ) := nnf.or (fml.to_nnf (neg φ)) (fml.to_nnf ψ) 44 | | (box φ) := nnf.box (fml.to_nnf φ) 45 | | (dia φ) := nnf.dia (fml.to_nnf φ) 46 | using_well_founded {rel_tac := λ _ _, `[exact ⟨_, measure_wf fml_size⟩]} 47 | 48 | @[simp] def trans_size_left {st} (k : kripke st) : (Σ' (s) (φ : fml), fml_force k s φ) → ℕ := 49 | λ h, fml_size h.snd.fst 50 | 51 | @[simp] def trans_size_right {st} (k : kripke st) : (Σ' (s : st) (φ : fml), force k s (fml.to_nnf φ)) → ℕ := 52 | λ h, fml_size h.snd.fst 53 | 54 | /- This direction requires choice -/ 55 | 56 | theorem trans_left {st} (k : kripke st) : Π s φ, 57 | fml_force k s φ → force k s (fml.to_nnf φ) 58 | | s (var n) h := h 59 | | s (neg (var n)) h := h 60 | | s (neg (neg φ)) h := by dsimp at h; rw classical.not_not at h; exact trans_left _ _ h 61 | | s (neg (and φ ψ)) h := begin 62 | have := trans_left s (neg φ), 63 | dsimp at h, 64 | rw classical.not_and_distrib at h, 65 | cases h with l r, 66 | {simp [trans_left _ (neg φ) l]}, 67 | {simp [trans_left _ (neg ψ) r]} 68 | end 69 | | s (neg (or φ ψ)) h := begin 70 | dsimp at h, 71 | rw not_or_distrib at h, 72 | simp [and.intro (trans_left _ (neg φ) h.1) (trans_left _ (neg ψ) h.2)] 73 | end 74 | | s (neg (impl φ ψ)) h := begin 75 | dsimp at h, 76 | rw not_or_distrib at h, 77 | have : force k s (fml.to_nnf φ), 78 | { apply trans_left, have := h.1, 79 | rw classical.not_not at this, 80 | exact this }, 81 | simp [and.intro this (trans_left _ (neg ψ) h.2)] 82 | end 83 | | s (neg (box φ)) h := begin 84 | dsimp at h, 85 | rw classical.not_forall at h, 86 | cases h with w hw, 87 | rw classical.not_imp at hw, 88 | simp, split, split, 89 | {exact hw.1}, {apply trans_left, exact hw.2} 90 | end 91 | | s (neg (dia φ)) h := begin 92 | dsimp at h, simp, 93 | intros s' hs', 94 | rw not_exists at h, 95 | have := h s', 96 | rw not_and at this, 97 | have hnf := this hs', 98 | apply trans_left, exact hnf 99 | end 100 | | s (and φ ψ) h := by dsimp at h; simp [and.intro (trans_left _ _ h.1) (trans_left _ _ h.2)] 101 | | s (or φ ψ) h := begin 102 | dsimp at h, cases h, 103 | {exact or.inl (trans_left _ _ h)}, 104 | {exact or.inr (trans_left _ _ h)} 105 | end 106 | | s (impl φ ψ) h := begin 107 | dsimp at h, simp, 108 | cases h, 109 | {left, apply trans_left, exact h}, 110 | {right, apply trans_left, exact h} 111 | end 112 | | s (box φ) h := begin 113 | dsimp at h, simp, 114 | intros s' hs', 115 | exact trans_left _ _ (h s' hs') 116 | end 117 | | s (dia φ) h := begin 118 | dsimp at h, simp, 119 | cases h with w hw, 120 | split, split, {exact hw.1}, {exact trans_left _ _ hw.2} 121 | end 122 | using_well_founded {rel_tac := λ _ _, `[exact ⟨_, measure_wf (trans_size_left k)⟩]} 123 | 124 | /- This direction is still constructive -/ 125 | 126 | theorem trans_right {st} (k : kripke st) : Π s φ, 127 | force k s (fml.to_nnf φ) → fml_force k s φ 128 | | s (var n) h := h 129 | | s (neg (var n)) h := h 130 | | s (neg (neg φ)) h := by dsimp; simp at h; exact not_not_intro (trans_right _ _ h) 131 | | s (neg (and φ ψ)) h := begin 132 | dsimp, simp at h, 133 | apply not_and_of_not_or_not, 134 | cases h, 135 | {left, exact trans_right _ _ h}, 136 | {right, exact trans_right _ _ h} 137 | end 138 | | s (neg (or φ ψ)) h := begin 139 | dsimp, simp at h, 140 | apply not_or, 141 | {exact trans_right _ _ h.1}, {exact trans_right _ _ h.2} 142 | end 143 | | s (neg (impl φ ψ)) h := begin 144 | dsimp, simp at h, 145 | apply not_or, 146 | {exact not_not_intro (trans_right _ _ h.1)}, {exact trans_right _ _ h.2} 147 | end 148 | | s (neg (box φ)) h := begin 149 | dsimp, simp at h, 150 | apply not_forall_of_exists_not, 151 | cases h with w hw, 152 | split, 153 | {apply not_imp_of_and_not, split, exact hw.1, exact trans_right _ _ hw.2} 154 | end 155 | | s (neg (dia φ)) h := begin 156 | dsimp, simp at h, 157 | rw not_exists, 158 | intros s' hs', 159 | have := h s' hs'.1, 160 | have hn := trans_right _ _ this, 161 | have := hs'.2, 162 | contradiction 163 | end 164 | | s (and φ ψ) h := begin 165 | dsimp, simp at h, 166 | split, {exact trans_right _ _ h.1}, {exact trans_right _ _ h.2} 167 | end 168 | | s (or φ ψ) h := begin 169 | dsimp, simp at h, 170 | cases h, {left, exact trans_right _ _ h}, {right, exact trans_right _ _ h} 171 | end 172 | | s (impl φ ψ) h := begin 173 | dsimp, simp at h, cases h, 174 | {left, exact trans_right _ _ h}, {right, exact trans_right _ _ h} 175 | end 176 | | s (box φ) h := begin 177 | dsimp, simp at h, 178 | intros s' hs', 179 | exact trans_right _ _ (h s' hs') 180 | end 181 | | s (dia φ) h := begin 182 | dsimp, simp at h, 183 | cases h with w hw, 184 | split, split, {exact hw.1}, {exact trans_right _ _ hw.2} 185 | end 186 | using_well_founded {rel_tac := λ _ _, `[exact ⟨_, measure_wf (trans_size_right k)⟩]} 187 | 188 | /- This is a classical result -/ 189 | 190 | theorem trans_iff {st} (k : kripke st) (s φ) : fml_force k s φ ↔ force k s (fml.to_nnf φ) := 191 | ⟨trans_left k s φ, trans_right k s φ⟩ 192 | 193 | def fml_sat {st} (k : kripke st) (s) (Γ : list fml) : Prop := 194 | ∀ φ ∈ Γ, fml_force k s φ 195 | 196 | theorem fml_sat_of_empty {st} (k : kripke st) (s) : fml_sat k s [] := 197 | λ φ h, absurd h $ list.not_mem_nil _ 198 | 199 | def fml_unsatisfiable (Γ : list fml) : Prop := 200 | ∀ (st) (k : kripke st) s, ¬ fml_sat k s Γ 201 | 202 | theorem trans_sat_iff {st} (k : kripke st) (s) : Π Γ, 203 | fml_sat k s Γ ↔ sat k s (list.map fml.to_nnf Γ) 204 | | [] := by simp [sat_of_empty, fml_sat_of_empty] 205 | | (hd::tl) := 206 | begin 207 | split, 208 | { intro h, dsimp, intros φ hφ, cases hφ, 209 | { rw hφ, rw ←trans_iff, apply h, simp }, 210 | { have : fml_sat k s tl, 211 | { intros ψ hψ, apply h, simp [hψ] }, 212 | rw trans_sat_iff at this, 213 | exact this _ hφ } }, 214 | { intros h φ hφ, cases hφ, 215 | { rw hφ, rw trans_iff, apply h, simp }, 216 | { have : sat k s (list.map fml.to_nnf tl), 217 | { intros ψ hψ, apply h, simp [hψ] }, 218 | rw ←trans_sat_iff at this, 219 | exact this _ hφ } } 220 | end 221 | 222 | theorem trans_unsat_iff (Γ : list fml) : fml_unsatisfiable Γ ↔ unsatisfiable (list.map fml.to_nnf Γ) := 223 | begin 224 | split, 225 | {intros h _ _ _ _, apply h, rw trans_sat_iff, assumption}, 226 | {intros h _ _ _ _, apply h, rw ←trans_sat_iff, assumption} 227 | end 228 | 229 | -------------------------------------------------------------------------------- /src/K/jump.lean: -------------------------------------------------------------------------------- 1 | import .rules 2 | open psum nnf node 3 | 4 | set_option eqn_compiler.zeta true 5 | 6 | def tableau : Π Γ : list nnf, node Γ 7 | | Γ := 8 | match get_contra Γ with 9 | | inl w := contra_rule w.2 10 | | inr no_contra := 11 | match get_and Γ with 12 | | inl p := 13 | let Δ := p.1.1 :: p.1.2 :: Γ.erase (nnf.and p.val.1 p.val.2) in 14 | let inst := and_instance.cons p.2 in 15 | have node_size ((p.val).fst :: (p.val).snd :: list.erase Γ (and ((p.val).fst) ((p.val).snd))) < node_size Γ, 16 | by apply split_lt_and; exact p.2, 17 | let d_delta : node Δ := tableau Δ in and_rule inst d_delta 18 | | inr no_and := 19 | match get_or Γ with 20 | | inl p := 21 | let Γ₁ := p.1.1 :: Γ.erase (nnf.or p.val.1 p.val.2) in 22 | let Γ₂ := p.1.2 :: Γ.erase (nnf.or p.val.1 p.val.2) in 23 | let inst := or_instance.cons p.2 in 24 | have node_size ((p.val).fst :: list.erase Γ (or ((p.val).fst) ((p.val).snd))) < node_size Γ, 25 | by apply split_lt_or_left; exact p.2, 26 | have node_size ((p.val).snd :: list.erase Γ (or ((p.val).fst) ((p.val).snd))) < node_size Γ, 27 | by apply split_lt_or_right; exact p.2, 28 | let d_Γ₁ : node (p.1.1 :: Γ.erase (or p.val.1 p.val.2)) := tableau Γ₁ in 29 | match d_Γ₁ with 30 | | closed m hsat pm := if hprcp : left_prcp inst ∈ m then 31 | or_rule inst (closed m hsat pm) (tableau Γ₂) 32 | else jump_rule inst m hsat pm hprcp 33 | | open_ w := open_rule inst w.2 34 | end 35 | | inr no_or := 36 | match get_dia Γ with 37 | | inl p := 38 | let ma : modal_applicable Γ := 39 | {no_contra := no_contra, 40 | no_and := no_and, 41 | no_or := no_or, 42 | v := get_var Γ, 43 | hv := λ n, get_var_iff, 44 | φ := p.1, 45 | ex := p.2} in 46 | let l := @tmap (λ Δ, node_size Δ < node_size Γ) 47 | (λ x h, tableau x) (unmodal Γ) 48 | (unmodal_size Γ) in 49 | match l with 50 | | inl w := begin 51 | left, 52 | {exact unsat_of_unsat_unmodal ma w.1.1 ⟨w.2.1, w.2.2.1⟩}, 53 | swap, {exact mark_modal Γ w.1.1 w.1.2}, 54 | { exact modal_pmark ma w.1.1 w.1.2 ⟨w.2.1, w.2.2.1⟩ w.2.2.2 } 55 | end 56 | | inr w := begin 57 | right, split, 58 | apply sat_of_batch_sat, 59 | exact ma, exact w.2 60 | end 61 | end 62 | | inr no_dia := 63 | let mc : model_constructible Γ := 64 | {no_contra := no_contra, 65 | no_and := no_and, 66 | no_or := no_or, 67 | v := get_var Γ, 68 | hv := λ n, get_var_iff, 69 | no_dia := no_dia} in 70 | by right; split; apply build_model; exact mc 71 | end 72 | end 73 | end 74 | end 75 | using_well_founded {rel_tac := λ _ _, `[exact ⟨_, measure_wf node_size⟩]} 76 | 77 | def is_sat (Γ : list nnf) : bool := 78 | match tableau Γ with 79 | | closed _ _ _:= ff 80 | | open_ _ := tt 81 | end 82 | 83 | theorem correctness (Γ : list nnf) : is_sat Γ = tt ↔ ∃ (st : Type) (k : kripke st) s, sat k s Γ := 84 | begin 85 | cases h : is_sat Γ, 86 | constructor, 87 | {intro, contradiction}, 88 | {intro hsat, cases eq : tableau Γ, 89 | rcases hsat with ⟨w, k, s, hsat⟩, 90 | apply false.elim, apply a, exact hsat, 91 | {dsimp [is_sat] at h, rw eq at h, contradiction} }, 92 | {split, intro, dsimp [is_sat] at h, 93 | cases eq : tableau Γ, 94 | { rw eq at h, contradiction }, 95 | { split, split, split, exact a_1.2}, 96 | { simp } } 97 | end 98 | 99 | /- The negation of the defining formula of K is unsatisfiable -/ 100 | 101 | def φ : nnf := 102 | and (box $ or (neg 1) (var 2)) (and (box $ var 1) (dia $ neg 2)) 103 | 104 | #eval is_sat [φ] 105 | 106 | def ψ : nnf := and (box (var 1)) (dia (dia (neg 1))) 107 | 108 | #eval is_sat [ψ] 109 | -------------------------------------------------------------------------------- /src/K/marking.lean: -------------------------------------------------------------------------------- 1 | import .semantics 2 | 3 | open nnf subtype list 4 | 5 | def pmark (Γ m : list nnf) := ∀ Δ, (∀ δ ∈ Δ, δ ∉ m) → Δ <+ Γ → unsatisfiable (list.diff Γ Δ) 6 | 7 | @[simp] def mark_and : nnf → list nnf → list nnf 8 | | (nnf.and φ ψ) m := if φ ∈ m ∨ ψ ∈ m then nnf.and φ ψ :: m else 9 | m 10 | | _ m := m 11 | 12 | @[simp] def mark_or : nnf → list nnf → list nnf → list nnf 13 | | (nnf.or φ ψ) ml mr:= if φ ∈ ml ∨ ψ ∈ mr then nnf.or φ ψ :: (ml++mr) else ml ++ mr 14 | | _ ml mr := ml ++ mr 15 | 16 | @[simp] def mark_modal (Γ i m : list nnf) : list nnf := 17 | dia i.head :: rebox (i.tail ∩ m) 18 | 19 | namespace list 20 | universes u v w x 21 | variables {α : Type u} {β : Type v} {γ : Type w} {δ : Type x} 22 | 23 | theorem mem_tail_of_ne_head [inhabited α] {a : α} : Π {l : list α}, a ∈ l → a ≠ l.head → a ∈ l.tail 24 | | [] h₁ h₂ := absurd h₁ $ not_mem_nil _ 25 | | (hd::tl) h₁ h₂ := by cases h₁; {simpa} 26 | 27 | end list 28 | 29 | theorem box_mem_of_mark_modal {φ} (Γ i m : list nnf) (h₁ : φ ∈ m) (h₂ : φ ∈ i) : 30 | φ = i.head ∨ box φ ∈ mark_modal Γ i m := 31 | begin 32 | simp, 33 | by_cases h : φ = i.head, 34 | {left, exact h}, 35 | {right, rw rebox_iff, simp, split, 36 | {apply list.mem_tail_of_ne_head h₂ h}, 37 | {exact h₁} } 38 | end 39 | 40 | theorem subset_mark_and {φ Γ} : Γ ⊆ mark_and φ Γ := 41 | begin 42 | cases heq : φ, 43 | case nnf.and : ψ₁ ψ₂ 44 | { dsimp, 45 | by_cases ψ₁ ∈ Γ ∨ ψ₂ ∈ Γ, 46 | {rw if_pos h, simp}, 47 | {rw if_neg h, simp} }, 48 | all_goals {simp} 49 | end 50 | 51 | theorem unsat_mark_and {φ Γ} (h : unsatisfiable Γ) : 52 | unsatisfiable (mark_and φ Γ) := 53 | begin apply unsat_subset, apply subset_mark_and, exact h end 54 | 55 | theorem subset_mark_or_left {φ Γ₁ Γ₂} : Γ₁ ⊆ mark_or φ Γ₁ Γ₂:= 56 | begin 57 | cases heq : φ, 58 | case nnf.or : ψ₁ ψ₂ 59 | { dsimp, 60 | by_cases ψ₁ ∈ Γ₁ ∨ ψ₂ ∈ Γ₂, 61 | {rw if_pos h, rw ←cons_append, 62 | apply subset_append_of_subset_left, simp}, 63 | {rw if_neg h, simp} }, 64 | all_goals {simp} 65 | end 66 | 67 | theorem subset_mark_or_right {φ Γ₁ Γ₂} : Γ₂ ⊆ mark_or φ Γ₁ Γ₂:= 68 | begin 69 | cases heq : φ, 70 | case nnf.or : ψ₁ ψ₂ 71 | { dsimp, 72 | by_cases ψ₁ ∈ Γ₁ ∨ ψ₂ ∈ Γ₂, 73 | {rw if_pos h, rw ←cons_append, 74 | apply subset_append_of_subset_right, simp}, 75 | {rw if_neg h, simp} }, 76 | all_goals {simp} 77 | end 78 | 79 | theorem unsat_mark_or_left {φ Γ₁ Γ₂} (h : unsatisfiable Γ₁) : 80 | unsatisfiable (mark_or φ Γ₁ Γ₂) := 81 | begin apply unsat_subset, apply subset_mark_or_left, exact h end 82 | 83 | theorem unsat_mark_or_right {φ Γ₁ Γ₂} (h : unsatisfiable Γ₂) : 84 | unsatisfiable (mark_or φ Γ₁ Γ₂) := 85 | begin apply unsat_subset, apply subset_mark_or_right, exact h end 86 | 87 | def pmark_of_closed_and {Γ Δ} (i : and_instance Γ Δ) (m) (h : unsatisfiable Δ) (p : pmark Δ m) : {x // pmark Γ x} := 88 | begin 89 | cases i with φ ψ hin, split, swap, 90 | {exact mark_and (and φ ψ) m}, 91 | { intros Δ' hΔ hsub, 92 | by_cases hm : φ ∈ m ∨ ψ ∈ m, 93 | -- pos hm 94 | { cases hm, 95 | -- left hm 96 | { have marked : nnf.and φ ψ ∈ mark_and (nnf.and φ ψ) m, 97 | {simp [mem_cons_self, hm] }, 98 | have : unsatisfiable (list.diff (φ :: ψ :: list.erase Γ (and φ ψ)) Δ'), 99 | {apply p, 100 | {intros δ hδ hmem, apply hΔ _ hδ, apply subset_mark_and hmem}, 101 | have : Δ'.erase (and φ ψ) <+ φ :: ψ :: list.erase Γ (and φ ψ), 102 | {apply sublist.cons, apply sublist.cons, apply sublist.erase _ hsub}, 103 | have hnin: and φ ψ ∉ Δ', 104 | {intro, apply hΔ _ a marked}, 105 | rw ←erase_of_not_mem hnin, exact this}, 106 | intro, intros, intro hsat, 107 | have hsat : sat k s (list.diff (φ :: ψ :: list.erase Γ (and φ ψ)) Δ'), 108 | {apply sat_subset _ (φ :: (ψ :: list.erase Γ (and φ ψ)).diff Δ'), 109 | apply subset_cons_diff, apply sat_subset _ (φ :: ψ :: (list.erase Γ (and φ ψ)).diff Δ'), apply cons_subset_cons, apply subset_cons_diff, 110 | have hsatp : and φ ψ ∈ Γ.diff Δ', 111 | {apply mem_diff_of_mem hin, intro, apply hΔ _ a marked}, 112 | intros x hx, rcases hx with eq₁ | eq₂ | hin', 113 | rw eq₁, apply and.left, rw ←sat_of_and, apply hsat _ hsatp, 114 | rw eq₂, apply and.right, rw ←sat_of_and, apply hsat _ hsatp, 115 | have hsubd : list.diff (list.erase Γ (and φ ψ)) Δ' ⊆ list.diff Γ Δ', 116 | {apply sublist.subset, apply sublist.diff_right, apply erase_sublist}, 117 | apply hsat, apply hsubd, exact hin'}, 118 | apply this, exact hsat }, 119 | -- right hm 120 | { have marked : nnf.and φ ψ ∈ mark_and (nnf.and φ ψ) m, 121 | { simp [mem_cons_self, hm] }, 122 | have : unsatisfiable (list.diff (φ :: ψ :: list.erase Γ (and φ ψ)) Δ'), 123 | {apply p, 124 | {intros δ hδ hmem, apply hΔ _ hδ, apply subset_mark_and hmem}, 125 | have : Δ'.erase (and φ ψ) <+ φ :: ψ :: list.erase Γ (and φ ψ), 126 | {apply sublist.cons, apply sublist.cons, apply sublist.erase, exact hsub}, 127 | have hnin: and φ ψ ∉ Δ', 128 | {intro, apply hΔ _ a marked}, 129 | rw ←erase_of_not_mem hnin, exact this}, 130 | intro, intros, intro hsat, 131 | have hsat : sat k s (list.diff (φ :: ψ :: list.erase Γ (and φ ψ)) Δ'), 132 | {apply sat_subset _ (φ :: (ψ :: list.erase Γ (and φ ψ)).diff Δ'), 133 | apply subset_cons_diff, apply sat_subset _ (φ :: ψ :: (list.erase Γ (and φ ψ)).diff Δ'), apply cons_subset_cons, apply subset_cons_diff, 134 | have hsatp : and φ ψ ∈ Γ.diff Δ', 135 | {apply mem_diff_of_mem hin, intro, apply hΔ _ a marked}, 136 | intros x hx, rcases hx with eq₁ | eq₂ | hin', 137 | rw eq₁, apply and.left, rw ←sat_of_and, apply hsat _ hsatp, 138 | rw eq₂, apply and.right, rw ←sat_of_and, apply hsat _ hsatp, 139 | have hsubd : list.diff (list.erase Γ (and φ ψ)) Δ' ⊆ list.diff Γ Δ', 140 | {apply sublist.subset, apply sublist.diff_right, apply erase_sublist}, 141 | apply hsat, apply hsubd, exact hin'}, 142 | apply this, exact hsat } }, 143 | -- neg hm 144 | { have : unsatisfiable ((φ :: ψ :: list.erase Γ (and φ ψ)).diff $ φ :: ψ :: Δ'.erase (and φ ψ)), 145 | {apply p, intros δ hδ hdin, rcases hδ with eq₁ | eq₂ | hmem, 146 | {apply hm, left, rw ←eq₁, exact hdin}, 147 | {apply hm, right, rw ←eq₂, exact hdin}, 148 | {apply hΔ, apply erase_subset _ _ hmem, apply subset_mark_and hdin}, 149 | apply sublist.cons2, apply sublist.cons2, apply sublist.erase _ hsub }, 150 | intro, intros, intro hsat, apply this, simp, 151 | apply sat_sublist, apply erase_diff_erase_sublist_of_sublist hsub, exact hsat } } 152 | end 153 | 154 | theorem unsat_of_jump {Γ₁ Γ₂ Δ : list nnf} (i : or_instance Δ Γ₁ Γ₂) 155 | (m : list nnf) (h₁ : unsatisfiable Γ₁) (h₂ : pmark Γ₁ m) 156 | (h₃ : left_prcp i ∉ m) : unsatisfiable Δ := 157 | begin 158 | cases i with φ ψ hin, 159 | have : unsatisfiable ((φ :: list.erase Δ (or φ ψ)).diff [φ]), 160 | {apply h₂, intros, cases H, rw H, exact h₃, intro, exact not_mem_nil δ H, apply sublist.cons2, simp}, 161 | intro, intros, intro hsat, apply this, 162 | apply sat_subset _ Δ _ _ _ hsat, simp [erase_subset] 163 | end 164 | 165 | def pmark_of_jump {Γ₁ Γ₂ Δ : list nnf} (i : or_instance Δ Γ₁ Γ₂) 166 | (m : list nnf) (h₁ : unsatisfiable Γ₁) (h₂ : pmark Γ₁ m) 167 | (h₃ : left_prcp i ∉ m) : {x // pmark Δ x} := 168 | begin 169 | cases i with φ ψ hin, 170 | split, swap, {exact mark_or (or φ ψ) m []}, 171 | { intros Δ' hΔ' hsub, 172 | have : unsatisfiable (list.diff (φ :: Δ.erase (or φ ψ)) (φ :: Δ'.erase (or φ ψ))), 173 | { apply h₂, intros, intro hmem, rcases H with eq | hin, 174 | {apply h₃, simpa [eq] using hmem}, 175 | {apply hΔ' δ, apply erase_subset _ _ hin, apply subset_mark_or_left hmem}, 176 | apply sublist.cons2, apply sublist.erase _ hsub }, 177 | intro, intros, intro hsat, 178 | apply this st k s, simp, apply sat_sublist, apply erase_diff_erase_sublist_of_sublist, exact hsub, exact hsat } 179 | end 180 | 181 | def pmark_of_closed_or {Γ₁ Γ₂ Δ : list nnf} {m₁ m₂ : list nnf} (i : or_instance Δ Γ₁ Γ₂) (h₁ : unsatisfiable Γ₁) (p₁ : pmark Γ₁ m₁) 182 | (h₂ : unsatisfiable Γ₂) (p₂ : pmark Γ₂ m₂) : {x // pmark Δ x} := 183 | begin 184 | cases i with φ ψ hin, split, swap, 185 | {exact mark_or (or φ ψ) m₁ m₂}, 186 | {intros Δ' hΔ hsub, 187 | by_cases hmφ : φ ∈ m₁, 188 | { have marked : or φ ψ ∈ mark_or (or φ ψ) m₁ m₂, 189 | {simp [hmφ]}, 190 | have hnin: or φ ψ ∉ Δ', 191 | {intro, apply hΔ _ a marked}, 192 | have hl : unsatisfiable (list.diff (φ :: list.erase Δ (or φ ψ)) Δ'), 193 | {apply p₁, 194 | {intros δ hδ hin, apply hΔ, exact hδ, apply subset_mark_or_left, exact hin}, 195 | have : Δ'.erase (or φ ψ) <+ φ :: list.erase Δ (or φ ψ), 196 | {apply sublist.cons, apply sublist.erase, exact hsub}, 197 | rw ←erase_of_not_mem hnin, exact this}, 198 | have hr : unsatisfiable (list.diff (ψ :: list.erase Δ (or φ ψ)) Δ'), 199 | {apply p₂, 200 | {intros δ hδ hin, apply hΔ, exact hδ, apply subset_mark_or_right, exact hin}, 201 | have : Δ'.erase (or φ ψ) <+ ψ :: list.erase Δ (or φ ψ), 202 | {apply sublist.cons, apply sublist.erase, exact hsub}, 203 | rw ←erase_of_not_mem hnin, exact this}, 204 | intro, intros, intro hsat, 205 | have : or φ ψ ∈ list.diff Δ Δ', 206 | {apply mem_diff_of_mem hin hnin}, 207 | have hforce := hsat _ this, dsimp at hforce, 208 | cases hforce with l r, 209 | { apply hl st k s, 210 | apply sat_subset, 211 | apply subperm.subset, 212 | apply subperm_cons_diff, 213 | intros x hx, 214 | cases hx, rw hx, exact l, apply hsat, 215 | have hsubd : list.diff (list.erase Δ (or φ ψ)) Δ' ⊆ list.diff Δ Δ', 216 | {apply sublist.subset, apply sublist.diff_right, apply erase_sublist}, 217 | apply hsubd hx }, 218 | { apply hr st k s, 219 | apply sat_subset, 220 | apply subperm.subset, 221 | apply subperm_cons_diff, 222 | intros x hx, 223 | cases hx, rw hx, exact r, apply hsat, 224 | have hsubd : list.diff (list.erase Δ (or φ ψ)) Δ' ⊆ list.diff Δ Δ', 225 | {apply sublist.subset, apply sublist.diff_right, apply erase_sublist}, 226 | apply hsubd hx } }, 227 | {by_cases hmψ : ψ ∈ m₂, 228 | -- marked ψ 229 | { have marked : or φ ψ ∈ mark_or (or φ ψ) m₁ m₂, 230 | {simp [hmψ]}, 231 | have hnin: or φ ψ ∉ Δ', 232 | {intro, apply hΔ _ a marked}, 233 | have hl : unsatisfiable (list.diff (φ :: list.erase Δ (or φ ψ)) Δ'), 234 | {apply p₁, 235 | {intros δ hδ hin, apply hΔ, exact hδ, apply subset_mark_or_left, exact hin}, 236 | have : Δ'.erase (or φ ψ) <+ φ :: list.erase Δ (or φ ψ), 237 | {apply sublist.cons, apply sublist.erase, exact hsub}, 238 | rw ←erase_of_not_mem hnin, exact this}, 239 | have hr : unsatisfiable (list.diff (ψ :: list.erase Δ (or φ ψ)) Δ'), 240 | {apply p₂, 241 | {intros δ hδ hin, apply hΔ, exact hδ, apply subset_mark_or_right, exact hin}, 242 | have : Δ'.erase (or φ ψ) <+ ψ :: list.erase Δ (or φ ψ), 243 | {apply sublist.cons, apply sublist.erase, exact hsub}, 244 | rw ←erase_of_not_mem hnin, exact this}, 245 | intro, intros, intro hsat, 246 | have : or φ ψ ∈ list.diff Δ Δ', 247 | {apply mem_diff_of_mem hin hnin}, 248 | have hforce := hsat _ this, dsimp at hforce, 249 | cases hforce with l r, 250 | { apply hl st k s, 251 | apply sat_subset, 252 | apply subperm.subset, 253 | apply subperm_cons_diff, 254 | intros x hx, 255 | cases hx, rw hx, exact l, apply hsat, 256 | have hsubd : list.diff (list.erase Δ (or φ ψ)) Δ' ⊆ list.diff Δ Δ', 257 | {apply sublist.subset, apply sublist.diff_right, apply erase_sublist}, 258 | apply hsubd hx }, 259 | { apply hr st k s, 260 | apply sat_subset, 261 | apply subperm.subset, 262 | apply subperm_cons_diff, 263 | intros x hx, 264 | cases hx, rw hx, exact r, apply hsat, 265 | have hsubd : list.diff (list.erase Δ (or φ ψ)) Δ' ⊆ list.diff Δ Δ', 266 | {apply sublist.subset, apply sublist.diff_right, apply erase_sublist}, 267 | apply hsubd hx } }, 268 | -- both unmarked 269 | { have : unsatisfiable (list.diff (φ :: Δ.erase (or φ ψ)) (φ :: Δ'.erase (or φ ψ))), 270 | {apply p₁, intros, intro hmem, rcases H with eq | hin, 271 | {apply hmφ, rw ←eq, exact hmem}, 272 | {apply hΔ δ, apply erase_subset _ _ hin, apply subset_mark_or_left, exact hmem}, 273 | apply sublist.cons2, apply sublist.erase _ hsub}, 274 | intro, intros, intro hsat, 275 | apply this st k s, simp, apply sat_sublist, 276 | apply erase_diff_erase_sublist_of_sublist hsub, exact hsat } } } 277 | end 278 | 279 | def unbox_sublist_of_unmodal (Γ : list nnf) : ∀ (i : list nnf), i ∈ unmodal Γ → (∀ Δ, Δ <+ Γ → unbox Δ <+ i) := 280 | list.mapp _ _ 281 | begin 282 | intros φ h Δ hΔ, 283 | apply sublist_cons_of_sublist, 284 | apply unbox_sublist hΔ 285 | end 286 | 287 | theorem modal_pmark {Γ} (h₁ : modal_applicable Γ) (i m) 288 | (h₂ : i ∈ unmodal Γ ∧ unsatisfiable i) 289 | (h₃ : pmark i m) : pmark Γ (mark_modal Γ i m) := 290 | begin 291 | intro, intros hδ hsub, intro, intros, intro hsat, 292 | let B' := filter (≠ i.head) (unbox Δ), 293 | have hsubB' : B' <+ i, 294 | {apply sublist.trans, swap 3, exact unbox Δ, apply filter_sublist, apply unbox_sublist_of_unmodal _ _ h₂.1 _ hsub}, 295 | have hB' : ∀ x, x ∈ B' → x ∉ m, 296 | {intros x hx hin, 297 | have hxi : x ∈ i, { exact sublist.subset hsubB' hx }, 298 | have := box_mem_of_mark_modal _ _ _ hin hxi, 299 | swap, {exact Γ}, 300 | cases this, 301 | {rw [mem_filter] at hx, have := hx.2, contradiction}, 302 | {apply hδ, swap, exact this, rw unbox_iff, rw [mem_filter] at hx, exact hx.1} }, 303 | have hunsat := h₃ B' hB' hsubB', 304 | have hhead : dia i.head ∉ Δ, 305 | {intro hmem, apply hδ, exact hmem, simp}, 306 | have := unmodal_jump _ _ h₂.1 _ _ _ _ hsat hhead, 307 | rcases this with ⟨w, hw⟩, 308 | apply hunsat, exact hw, 309 | end 310 | -------------------------------------------------------------------------------- /src/K/ops.lean: -------------------------------------------------------------------------------- 1 | import .size data.list.perm 2 | open nnf list 3 | 4 | namespace list 5 | universes u v w 6 | 7 | variables {α : Type u} {β : Type v} {γ : Type w} 8 | 9 | theorem mapp {p : β → Prop} (f : α → β) : Π (l : list α) (h : ∀ x∈l, p (f x)) x, x ∈ list.map f l → p x 10 | | [] h x := by simp 11 | | (hd::tl) h x := 12 | begin 13 | intro hmem, cases hmem, 14 | { simp [hmem, h] }, 15 | { apply mapp tl _ _ hmem, intros a ha, simp [h, ha] } 16 | end 17 | 18 | def ne_empty_head : Π l : list α, l ≠ [] → α 19 | | [] h := by contradiction 20 | | (a :: l) h := a 21 | 22 | end list 23 | 24 | @[simp] def unbox : list nnf → list nnf 25 | | [] := [] 26 | | ((box φ) :: l) := φ :: unbox l 27 | | (e :: l) := unbox l 28 | 29 | theorem unbox_sublist_cons {Γ : list nnf} {φ} : unbox Γ <+ unbox (φ :: Γ) := 30 | by cases heq : φ ; { simp } 31 | 32 | theorem unbox_sublist : Π {Γ₁ Γ₂ : list nnf} (h : Γ₁ <+ Γ₂), 33 | unbox Γ₁ <+ unbox Γ₂ 34 | | [] Γ₂ h := by simp 35 | | (hd::tl) [] h := by have := eq_nil_of_sublist_nil h; contradiction 36 | | (hd₁::tl₁) (hd₂::tl₂) h := 37 | begin 38 | cases h with _ _ _ c₁ _ _ _ c₂, 39 | { simp [sublist.trans (unbox_sublist c₁), unbox_sublist_cons] }, 40 | { cases heq : hd₁, 41 | case nnf.box : ψ { simp [sublist.cons2, unbox_sublist c₂] }, 42 | all_goals { simp [unbox_sublist c₂] } } 43 | end 44 | 45 | theorem unbox_erase : Π {Γ : list nnf} {φ}, 46 | (unbox Γ).erase φ = unbox (Γ.erase (box φ)) 47 | | [] φ := by simp 48 | | (hd::tl) φ := 49 | begin 50 | cases heq : hd, 51 | case nnf.box : ψ 52 | { dsimp, by_cases hφ : ψ = φ, 53 | { simp [hφ] }, 54 | { rw [erase_cons_tail, erase_cons_tail], simp, 55 | apply unbox_erase, intro hbe, simp at hbe, apply hφ hbe, exact hφ} }, 56 | all_goals { simp [unbox_erase] } 57 | end 58 | 59 | theorem unbox_erase_of_ne_box : Π {Γ : list nnf} {φ : nnf} (h : ∀ ψ, φ ≠ box ψ), unbox (Γ.erase φ) = unbox Γ 60 | | [] φ h := by simp 61 | | (hd::tl) φ h := 62 | begin 63 | cases heq : hd, 64 | case nnf.box : ψ 65 | { dsimp, by_cases hφ : φ = box ψ, 66 | { exfalso, apply h _ hφ }, 67 | { rw erase_cons_tail, simp, 68 | apply unbox_erase_of_ne_box h, 69 | intro, rw a at hφ, contradiction } }, 70 | all_goals 71 | { dsimp, by_cases hφ : φ = hd, 72 | { simp [hφ, heq] }, 73 | { rw heq at hφ, rw erase_cons_tail, simp, 74 | apply unbox_erase_of_ne_box h, 75 | intro, rw a at hφ, contradiction } } 76 | end 77 | 78 | theorem unbox_diff : Π {Γ₁ Γ₂ : list nnf}, 79 | (unbox Γ₂).diff (unbox Γ₁) = unbox (Γ₂.diff Γ₁) 80 | | [] Γ₂ := by simp 81 | | (hd::tl) Γ₂ := 82 | begin 83 | cases heq : hd, 84 | case nnf.box : ψ { simp [unbox_diff, unbox_erase] }, 85 | all_goals 86 | { by_cases hin : hd ∈ Γ₂, 87 | { simp, rw [←heq, ←(@unbox_diff tl (Γ₂.erase hd)), unbox_erase_of_ne_box], 88 | intros ψ hψ, rw heq at hψ, contradiction }, 89 | { simp, rw [←heq, erase_of_not_mem hin], apply unbox_diff } } 90 | end 91 | 92 | theorem unbox_iff : Π {Γ φ}, box φ ∈ Γ ↔ φ ∈ unbox Γ 93 | | [] φ := begin split, repeat {intro h, simpa using h} end 94 | | (hd::tl) φ := 95 | begin 96 | split, 97 | { intro h, cases h₁ : hd, 98 | case nnf.box : ψ 99 | { cases h, 100 | {left, rw h₁ at h, injection h}, 101 | {right, exact (@unbox_iff tl φ).1 h} }, 102 | all_goals 103 | { cases h, 104 | {rw h₁ at h, contradiction}, 105 | {exact (@unbox_iff tl φ).1 h} } }, 106 | { intro h, cases h₁ : hd, 107 | case nnf.box : ψ 108 | { rw h₁ at h, cases h, 109 | {simp [h]}, {right, exact (@unbox_iff tl φ).2 h} }, 110 | all_goals 111 | { rw h₁ at h, right, exact (@unbox_iff tl φ).2 h } } 112 | end 113 | 114 | theorem unbox_size_aux : Π {Γ}, node_size (unbox Γ) ≤ node_size Γ 115 | | [] := by simp 116 | | (hd::tl) := 117 | begin 118 | cases h : hd, 119 | case nnf.box : ψ 120 | { apply add_le_add, 121 | { dsimp [sizeof, has_sizeof.sizeof, nnf.sizeof], 122 | rw add_comm, apply nat.le_succ }, 123 | { apply unbox_size_aux } }, 124 | all_goals 125 | { dsimp, apply le_add_of_nonneg_of_le, 126 | { dsimp [sizeof, has_sizeof.sizeof, nnf.sizeof], rw add_comm, apply nat.zero_le }, 127 | { apply unbox_size_aux } } 128 | end 129 | 130 | theorem unbox_size : Π {Γ φ}, dia φ ∈ Γ → 131 | node_size (unbox Γ) + sizeof φ < node_size Γ 132 | | [] := by simp 133 | | (hd::tl) := 134 | begin 135 | intros φ h, 136 | cases h₁ : hd, 137 | case nnf.box : ψ 138 | { dsimp, cases h, 139 | { rw h₁ at h, contradiction }, 140 | { rw add_assoc, apply add_lt_add, 141 | { dsimp [sizeof, has_sizeof.sizeof, nnf.sizeof], rw add_comm, apply nat.lt_succ_self }, 142 | { apply unbox_size h } } }, 143 | case nnf.dia : ψ 144 | { dsimp, rw add_comm, cases h, 145 | { rw h₁ at h, have : φ = ψ, { injection h }, rw this, 146 | apply add_lt_add_of_lt_of_le, 147 | { dsimp [sizeof, has_sizeof.sizeof, nnf.sizeof], rw add_comm, apply nat.lt_succ_self }, 148 | { apply unbox_size_aux } }, 149 | { apply lt_add_of_pos_of_lt, 150 | { dsimp [sizeof, has_sizeof.sizeof, nnf.sizeof], rw add_comm, apply nat.succ_pos }, 151 | { rw add_comm, apply unbox_size h } } }, 152 | all_goals 153 | { dsimp, apply nat.lt_add_left, apply unbox_size, 154 | cases h, rw h₁ at h, contradiction, exact h } 155 | end 156 | 157 | @[simp] def rebox : list nnf → list nnf 158 | | [] := [] 159 | | (hd::tl) := box hd :: rebox tl 160 | 161 | theorem rebox_unbox_of_mem : Π {Γ} (h : ∀ {φ}, φ ∈ unbox Γ → box φ ∈ Γ), rebox (unbox Γ) ⊆ Γ 162 | | [] h := by simp 163 | | (hd::tl) h := 164 | begin 165 | cases hψ : hd, 166 | case nnf.box : φ {simp [cons_subset_cons, subset_cons_of_subset, rebox_unbox_of_mem, unbox_iff]}, 167 | all_goals {simp [subset_cons_of_subset, rebox_unbox_of_mem, unbox_iff]} 168 | end 169 | 170 | theorem unbox_rebox : Π {Γ}, unbox (rebox Γ) = Γ 171 | | [] := by simp 172 | | (hd::tl) := by simp [unbox_rebox] 173 | 174 | def box_only_rebox : Π {Γ}, box_only (rebox Γ) 175 | | [] := {no_var := by simp, 176 | no_neg := by simp, 177 | no_and := by simp, 178 | no_or := by simp, 179 | no_dia := by simp} 180 | | (hd::tl) := 181 | begin 182 | cases h : hd, 183 | all_goals { 184 | exact { no_var := begin 185 | intros n hn, cases hn, contradiction, 186 | apply (@box_only_rebox tl).no_var, exact hn 187 | end, 188 | no_neg := begin 189 | intros n hn, cases hn, contradiction, 190 | apply (@box_only_rebox tl).no_neg, exact hn 191 | end, 192 | no_and := begin 193 | intros φ ψ hand, cases hand, contradiction, 194 | apply (@box_only_rebox tl).no_and, exact hand 195 | end, 196 | no_or := begin 197 | intros φ ψ hor, cases hor, contradiction, 198 | apply (@box_only_rebox tl).no_or, exact hor 199 | end, 200 | no_dia := begin 201 | intros φ hdia, cases hdia, contradiction, 202 | apply (@box_only_rebox tl).no_dia, exact hdia 203 | end} } 204 | end 205 | 206 | theorem rebox_iff : Π {φ Γ}, box φ ∈ rebox Γ ↔ φ ∈ Γ 207 | | φ [] := by simp 208 | | φ (hd::tl) := 209 | begin 210 | split, 211 | { intro h, cases h₁ : hd, 212 | all_goals { cases h, 213 | { left, rw ←h₁, injection h }, 214 | { right, exact (@rebox_iff φ tl).1 h } } }, 215 | { intro h, cases h₁ : hd, 216 | all_goals { dsimp, cases h, 217 | { left, rw [←h₁, h]}, 218 | { right, exact (@rebox_iff φ tl).2 h } } } 219 | end 220 | 221 | @[simp] def undia : list nnf → list nnf 222 | | [] := [] 223 | | ((dia φ) :: l) := φ :: undia l 224 | | (e :: l) := undia l 225 | 226 | theorem undia_iff : Π {Γ φ}, dia φ ∈ Γ ↔ φ ∈ undia Γ 227 | | [] φ := begin split, repeat {intro h, simpa using h} end 228 | | (hd::tl) φ := 229 | begin 230 | split, 231 | { intro h, cases h₁ : hd, 232 | case nnf.dia : ψ 233 | { cases h, 234 | { left, rw h₁ at h, injection h }, 235 | { right, exact (@undia_iff tl φ).1 h} }, 236 | all_goals 237 | { cases h, 238 | { rw h₁ at h, contradiction }, 239 | { exact (@undia_iff tl φ).1 h } } }, 240 | { intro h, cases h₁ : hd, 241 | case nnf.dia : ψ 242 | { rw h₁ at h, cases h, 243 | { simp [h] }, {right, exact (@undia_iff tl φ).2 h} }, 244 | all_goals 245 | { rw h₁ at h, right, exact (@undia_iff tl φ).2 h } } 246 | end 247 | 248 | theorem undia_size : Π {Γ φ}, φ ∈ undia Γ → 249 | sizeof φ + node_size (unbox Γ) < node_size Γ 250 | := 251 | begin intros Γ φ h, rw add_comm, apply unbox_size, rw undia_iff, exact h end 252 | 253 | @[simp] def get_modal : list nnf → list nnf 254 | | [] := [] 255 | | (φ@(dia ψ) :: l) := φ :: get_modal l 256 | | (φ@(box ψ) :: l) := φ :: get_modal l 257 | | (e :: l) := get_modal l 258 | 259 | theorem get_modal_iff_dia : Π {Γ} {φ : nnf}, 260 | dia φ ∈ get_modal Γ ↔ dia φ ∈ Γ 261 | | [] φ := by simp 262 | | (hd::tl) φ := 263 | begin 264 | split, 265 | { intro h, cases heq : hd, 266 | case nnf.dia : ψ 267 | { rw heq at h, cases h, left, exact h, right, exact (@get_modal_iff_dia tl φ).1 h }, 268 | case nnf.box : ψ 269 | { rw heq at h, cases h, contradiction, right, exact (@get_modal_iff_dia tl φ).1 h }, 270 | all_goals 271 | { rw heq at h, right, exact (@get_modal_iff_dia tl φ).1 h } }, 272 | { intro h, cases heq : hd, 273 | case nnf.dia : ψ 274 | { rw heq at h, cases h, left, exact h, right, exact (@get_modal_iff_dia tl φ).2 h }, 275 | case nnf.box : ψ 276 | { rw heq at h, cases h, contradiction, right, exact (@get_modal_iff_dia tl φ).2 h }, 277 | all_goals 278 | { rw heq at h, cases h, contradiction, exact (@get_modal_iff_dia tl φ).2 h } } 279 | end 280 | 281 | theorem get_modal_iff_box : Π {Γ} {φ : nnf}, 282 | box φ ∈ get_modal Γ ↔ box φ ∈ Γ 283 | | [] φ := by simp 284 | | (hd::tl) φ := 285 | begin 286 | split, 287 | { intro h, cases heq : hd, 288 | case nnf.box : ψ 289 | { rw heq at h, cases h, left, exact h, right, exact (@get_modal_iff_box tl φ).1 h }, 290 | case nnf.dia : ψ 291 | { rw heq at h, cases h, contradiction, right, exact (@get_modal_iff_box tl φ).1 h }, 292 | all_goals 293 | {rw heq at h, right, exact (@get_modal_iff_box tl φ).1 h } }, 294 | { intro h, cases heq : hd, 295 | case nnf.box : ψ 296 | { rw heq at h, cases h, left, exact h, right, exact (@get_modal_iff_box tl φ).2 h }, 297 | case nnf.dia : ψ 298 | { rw heq at h, cases h, contradiction, right, exact (@get_modal_iff_box tl φ).2 h }, 299 | all_goals 300 | { rw heq at h, cases h, contradiction, exact (@get_modal_iff_box tl φ).2 h } } 301 | end 302 | 303 | def get_contra : Π Γ : list nnf, 304 | psum {p : nat // var p ∈ Γ ∧ neg p ∈ Γ} 305 | (∀ n, var n ∈ Γ → neg n ∉ Γ) 306 | | [] := psum.inr $ λ _ h, absurd h $ not_mem_nil _ 307 | | (hd :: tl) := 308 | begin 309 | cases h : hd, 310 | case nnf.var : n 311 | {apply dite (neg n ∈ tl), 312 | { intro t, exact psum.inl ⟨n, ⟨mem_cons_self _ _, mem_cons_of_mem _ t⟩⟩ }, 313 | {intro e, cases (get_contra tl), 314 | {left, split, split, 315 | apply mem_cons_of_mem _ val.2.1, 316 | apply mem_cons_of_mem _ val.2.2}, 317 | {right, 318 | intros m hm hin, 319 | by_cases eq : m=n, 320 | {apply e, cases hin, contradiction, rw ←eq, assumption}, 321 | {cases hm, apply eq, injection hm, apply val _ hm, 322 | cases hin, contradiction, assumption} } } 323 | }, 324 | case nnf.neg : n 325 | { apply dite (var n ∈ tl), 326 | { intro t, exact psum.inl ⟨n, ⟨mem_cons_of_mem _ t, mem_cons_self _ _⟩⟩ }, 327 | { intro e, 328 | cases (get_contra tl), 329 | {left, split, split, 330 | apply mem_cons_of_mem _ val.2.1, 331 | apply mem_cons_of_mem _ val.2.2 }, 332 | { right, 333 | intros m hm hin, 334 | by_cases eq : m=n, 335 | { apply e, cases hm, contradiction, rw ←eq, assumption }, 336 | { cases hin, apply eq, injection hin, apply val, 337 | swap, exact hin, cases hm, contradiction, assumption } } } }, 338 | all_goals 339 | { cases (get_contra tl), 340 | { left, constructor, constructor, 341 | apply mem_cons_of_mem, exact val.2.1, 342 | apply mem_cons_of_mem, exact val.2.2 }, 343 | { right, 344 | intros m hm hin, 345 | {apply val, swap 3, exact m, 346 | cases hm, contradiction, assumption, 347 | cases hin, contradiction, assumption} } } 348 | end 349 | 350 | def get_and : Π Γ : list nnf, 351 | psum {p : nnf × nnf // and p.1 p.2 ∈ Γ} 352 | (∀ φ ψ, nnf.and φ ψ ∉ Γ) 353 | | [] := psum.inr $ λ _ _, not_mem_nil _ 354 | | (hd :: tl) := 355 | begin 356 | cases h : hd, 357 | case nnf.and : φ ψ { left, constructor,swap, 358 | constructor, exact φ, exact ψ, simp }, 359 | all_goals 360 | { cases (get_and tl), 361 | { left, constructor, 362 | apply mem_cons_of_mem _ val.2}, 363 | { right, intros γ ψ h, 364 | cases h, contradiction, 365 | apply val, assumption } } 366 | end 367 | 368 | def get_or : Π Γ : list nnf, 369 | psum {p : nnf × nnf // or p.1 p.2 ∈ Γ} 370 | (∀ φ ψ, nnf.or φ ψ ∉ Γ) 371 | | [] := psum.inr $ λ _ _, not_mem_nil _ 372 | | (hd :: tl) := 373 | begin 374 | cases h : hd, 375 | case nnf.or : φ ψ { left, constructor,swap, 376 | constructor, exact φ, exact ψ, simp }, 377 | all_goals 378 | { cases (get_or tl), 379 | { left, constructor, 380 | apply mem_cons_of_mem _ val.2}, 381 | { right, intros γ ψ h, 382 | cases h, contradiction, 383 | apply val, assumption } } 384 | end 385 | 386 | def get_dia : Π Γ : list nnf, 387 | psum {p : nnf // dia p ∈ Γ} 388 | (∀ φ, nnf.dia φ ∉ Γ) 389 | | [] := psum.inr $ λ _, not_mem_nil _ 390 | | (hd :: tl) := 391 | begin 392 | cases h : hd, 393 | case nnf.dia : φ { left, constructor, swap, exact φ, simp }, 394 | all_goals 395 | { cases (get_dia tl), 396 | { left, constructor, 397 | apply mem_cons_of_mem _ val.2}, 398 | { right, intros γ h, 399 | cases h, contradiction, 400 | apply val, assumption } } 401 | end 402 | 403 | @[simp] def get_var : list nnf → list ℕ 404 | | [] := [] 405 | | ((var n) :: l) := n :: get_var l 406 | | (e :: l) := get_var l 407 | 408 | theorem get_var_iff : Π {Γ n}, var n ∈ Γ ↔ n ∈ get_var Γ 409 | | [] φ := begin split, repeat {intro h, simpa using h} end 410 | | (hd::tl) φ := 411 | begin 412 | split, 413 | { intro h, cases h₁ : hd, 414 | case nnf.var : n 415 | { cases h, 416 | { left, rw h₁ at h, injection h }, 417 | { right, exact (@get_var_iff tl φ).1 h } }, 418 | all_goals 419 | { cases h, 420 | { rw h₁ at h, contradiction }, 421 | { exact (@get_var_iff tl φ).1 h } } }, 422 | { intro h, cases h₁ : hd, 423 | case nnf.var : n 424 | { rw h₁ at h, cases h, 425 | { simp [h] }, { right, exact (@get_var_iff tl φ).2 h } }, 426 | all_goals 427 | { rw h₁ at h, right, exact (@get_var_iff tl φ).2 h } } 428 | end 429 | -------------------------------------------------------------------------------- /src/K/rules.lean: -------------------------------------------------------------------------------- 1 | import .marking 2 | open nnf subtype list batch_sat 3 | 4 | inductive node (Γ : list nnf) : Type 5 | | closed : Π m, unsatisfiable Γ → pmark Γ m → node 6 | | open_ : {s // sat builder s Γ} → node 7 | 8 | open node 9 | 10 | def and_rule {Γ Δ} (i : and_instance Γ Δ) : node Δ → node Γ 11 | | (closed m h p) := begin 12 | left, 13 | {apply unsat_of_closed_and, repeat {assumption}}, 14 | exact (pmark_of_closed_and i m h p).2 15 | end 16 | 17 | | (open_ w) := begin 18 | right, 19 | clear and_rule, 20 | cases i with φ ψ h, 21 | constructor, 22 | exact sat_and_of_sat_split _ _ _ _ _ h w.2 23 | end 24 | 25 | def jump_rule {Γ₁ Γ₂ Δ : list nnf} (i : or_instance Δ Γ₁ Γ₂) 26 | (m : list nnf) (h₁ : unsatisfiable Γ₁) (h₂ : pmark Γ₁ m) 27 | (h₃ : left_prcp i ∉ m) : node Δ := 28 | begin 29 | left, 30 | {apply unsat_of_jump, repeat {assumption}}, 31 | exact (pmark_of_jump i m h₁ h₂ h₃).2 32 | end 33 | 34 | def or_rule {Γ₁ Γ₂ Δ} (i : or_instance Δ Γ₁ Γ₂) : 35 | node Γ₁ → node Γ₂ → node Δ 36 | | (open_ w) _ := begin 37 | right, 38 | clear or_rule, 39 | cases i with φ ψ h, 40 | constructor, 41 | exact sat_or_of_sat_split_left _ _ _ _ _ h w.2 42 | end 43 | | _ (open_ w) := begin 44 | right, 45 | clear or_rule, 46 | cases i with φ ψ h, 47 | constructor, 48 | exact sat_or_of_sat_split_right _ _ _ _ _ h w.2 49 | end 50 | | (closed m₁ h₁ p₁) (closed m₂ h₂ p₂):= 51 | begin 52 | left, 53 | {apply unsat_of_closed_or, repeat {assumption}}, 54 | exact (pmark_of_closed_or i h₁ p₁ h₂ p₂).2 55 | end 56 | 57 | def open_rule {Γ₁ Γ₂ Δ} {s} (i : or_instance Δ Γ₁ Γ₂) (h : sat builder s Γ₁) : node Δ := 58 | begin 59 | right, constructor, 60 | cases i with φ ψ hin, 61 | swap, exact s, 62 | exact sat_or_of_sat_split_left _ _ _ _ _ hin h 63 | end 64 | 65 | def contra_rule {Δ n} (h : var n ∈ Δ ∧ neg n ∈ Δ) : node Δ := 66 | begin 67 | left, 68 | {exact unsat_contra h.1 h.2}, 69 | swap, 70 | {exact [var n, neg n]}, 71 | {intros Δ' hΔ' hsub, 72 | apply unsat_contra, 73 | {apply mem_diff_of_mem, 74 | exact h.1, intro, apply hΔ' (var n) a, simp}, 75 | {apply mem_diff_of_mem, 76 | exact h.2, intro, apply hΔ' (neg n) a, simp} } 77 | end 78 | 79 | /- This is essentially the modal rule. -/ 80 | 81 | def tmap {p : list nnf → Prop} (f : Π Γ, p Γ → node Γ): Π Γ : list $ list nnf, (∀ i∈Γ, p i) → 82 | psum 83 | ({ c : list nnf × list nnf // c.1 ∈ Γ ∧ unsatisfiable c.1 ∧ pmark c.1 c.2}) 84 | {x // batch_sat x Γ} 85 | | [] h := psum.inr ⟨[], bs_nil⟩ 86 | | (hd :: tl) h := 87 | match f hd (h hd (by simp)) with 88 | | (node.closed m pr pm) := psum.inl ⟨⟨hd,m⟩, by simp, pr, pm⟩ 89 | | (node.open_ w₁) := 90 | match tmap tl (λ x hx, h x $ by simp [hx]) with 91 | | (psum.inl uw) := begin 92 | left, rcases uw with ⟨w, hin, h⟩, 93 | split, split, swap, exact h, simp [hin] 94 | end 95 | | (psum.inr w₂) := psum.inr ⟨(w₁.1::w₂), bs_cons _ _ _ _ w₁.2 w₂.2⟩ 96 | end 97 | end 98 | -------------------------------------------------------------------------------- /src/K/semantics.lean: -------------------------------------------------------------------------------- 1 | import .ops 2 | open subtype nnf 3 | 4 | @[simp] def force {states : Type} (k : kripke states) : states → nnf → Prop 5 | | s (var n) := k.val n s 6 | | s (neg n) := ¬ k.val n s 7 | | s (and φ ψ) := force s φ ∧ force s ψ 8 | | s (or φ ψ) := force s φ ∨ force s ψ 9 | | s (box φ) := ∀ s', k.rel s s' → force s' φ 10 | | s (dia φ) := ∃ s', k.rel s s' ∧ force s' φ 11 | 12 | def sat {st} (k : kripke st) (s) (Γ : list nnf) : Prop := 13 | ∀ φ ∈ Γ, force k s φ 14 | 15 | def unsatisfiable (Γ : list nnf) : Prop := 16 | ∀ (st) (k : kripke st) s, ¬ sat k s Γ 17 | 18 | theorem unsat_singleton {φ} (h : unsatisfiable [φ]) : ∀ {st} (k : kripke st) s, ¬ force k s φ 19 | := 20 | begin 21 | intros st k s hf, apply h, intros ψ hψ, rw list.mem_singleton at hψ, rw hψ, exact hf 22 | end 23 | 24 | theorem sat_of_empty {st} (k : kripke st) (s) : sat k s [] := 25 | λ φ h, absurd h $ list.not_mem_nil _ 26 | 27 | theorem ne_empty_of_unsat {Γ} (h : unsatisfiable Γ): Γ ≠ [] := 28 | begin 29 | intro heq, rw heq at h, 30 | apply h, apply @sat_of_empty nat, 31 | apply inhabited_kripke.1, exact 0 32 | end 33 | 34 | class val_constructible (Γ : list nnf) extends saturated Γ:= 35 | (no_contra : ∀ {n}, var n ∈ Γ → neg n ∉ Γ) 36 | (v : list ℕ) 37 | (hv : ∀ {n}, var n ∈ Γ ↔ n ∈ v) 38 | 39 | class modal_applicable (Γ : list nnf) extends val_constructible Γ := 40 | (φ : nnf) 41 | (ex : dia φ ∈ Γ) 42 | 43 | class model_constructible (Γ : list nnf) extends val_constructible Γ := 44 | (no_dia : ∀ {φ}, nnf.dia φ ∉ Γ) 45 | 46 | -- unbox and undia take a list of formulas, and 47 | -- get rid of the outermost box or diamond of each formula respectively 48 | def unmodal (Γ : list nnf) : list $ list nnf := 49 | list.map (λ d, d :: (unbox Γ)) (undia Γ) 50 | 51 | theorem unmodal_size (Γ : list nnf) : ∀ (i : list nnf), i ∈ unmodal Γ → (node_size i < node_size Γ) := 52 | list.mapp _ _ begin intros φ h, apply undia_size h end 53 | 54 | def unmodal_mem_box (Γ : list nnf) : ∀ (i : list nnf), i ∈ unmodal Γ → (∀ φ, box φ ∈ Γ → φ ∈ i) := 55 | list.mapp _ _ begin intros φ h ψ hψ, right, apply (@unbox_iff Γ ψ).1 hψ end 56 | 57 | def unmodal_sat_of_sat (Γ : list nnf) : ∀ (i : list nnf), i ∈ unmodal Γ → 58 | (∀ {st : Type} (k : kripke st) s Δ 59 | (h₁ : ∀ φ, box φ ∈ Γ → box φ ∈ Δ) 60 | (h₂ : ∀ φ, dia φ ∈ Γ → dia φ ∈ Δ), 61 | sat k s Δ → ∃ s', sat k s' i) := 62 | list.mapp _ _ 63 | begin 64 | intros φ hmem st k s Δ h₁ h₂ h, 65 | have : force k s (dia φ), 66 | { apply h, apply h₂, rw undia_iff, exact hmem }, 67 | rcases this with ⟨w, hrel, hforce⟩, 68 | split, swap, { exact w }, 69 | { intro ψ, intro hψ, cases hψ, 70 | { rw hψ, exact hforce }, 71 | { apply (h _ (h₁ _ ((@unbox_iff Γ ψ).2 hψ))) _ hrel} }, 72 | end 73 | 74 | def unmodal_mem_head (Γ : list nnf) : ∀ (i : list nnf), i ∈ unmodal Γ → dia (list.head i) ∈ Γ := 75 | list.mapp _ _ begin intros φ hmem, rw undia_iff, exact hmem end 76 | 77 | def unmodal_unsat_of_unsat (Γ : list nnf) : ∀ (i : list nnf), i ∈ unmodal Γ → Π h : unsatisfiable i, unsatisfiable (dia (list.head i) :: rebox (unbox Γ)) := 78 | list.mapp _ _ 79 | begin 80 | intros φ _, 81 | { intro h, intro, intros k s hsat, 82 | have ex := hsat (dia φ) (by simp), 83 | cases ex with s' hs', 84 | apply h st k s', intros ψ hmem, 85 | cases hmem, 86 | { rw hmem, exact hs'.2 }, 87 | { have := (@rebox_iff ψ (unbox Γ)).2 hmem, 88 | apply hsat (box ψ) (by right; assumption) s' hs'.1 } } 89 | end 90 | 91 | def mem_unmodal (Γ : list nnf) (φ) (h : φ ∈ undia Γ) : (φ :: unbox Γ) ∈ unmodal Γ := 92 | begin apply list.mem_map_of_mem (λ φ, φ :: unbox Γ) h end 93 | 94 | def unsat_of_unsat_unmodal {Γ : list nnf} (h : modal_applicable Γ) (i) : i ∈ unmodal Γ ∧ unsatisfiable i → unsatisfiable Γ := 95 | begin 96 | intro hex, intros st k s h, 97 | have := unmodal_sat_of_sat Γ i hex.1 k s Γ (λ x hx, hx) (λ x hx, hx) h, 98 | cases this with w hw, 99 | exact hex.2 _ k w hw 100 | end 101 | 102 | namespace list 103 | universes u v w x 104 | variables {α : Type u} {β : Type v} {γ : Type w} {δ : Type x} 105 | 106 | theorem cons_diff_of_ne_mem [decidable_eq α] {a : α} : Π {l₁ l₂ : list α} (h : a ∉ l₂), (a::l₁).diff l₂ = a :: l₁.diff l₂ 107 | | l₁ [] h := by simp 108 | | l₁ (hd::tl) h := 109 | begin 110 | simp, 111 | rw erase_cons_tail, 112 | apply cons_diff_of_ne_mem, 113 | {intro hin, apply h, simp [hin]}, 114 | {intro heq, apply h, simp [heq]} 115 | end 116 | 117 | -- TODO: Can we strengthen this? 118 | theorem subset_of_diff_filter [decidable_eq α] {a : α} : Π {l₁ l₂ : list α}, l₁.diff (filter (≠ a) l₂) ⊆ a :: l₁.diff l₂ 119 | | l₁ [] := by simp 120 | | l₁ (hd::tl) := 121 | begin 122 | by_cases heq : hd = a, 123 | {rw heq, simp, 124 | by_cases ha : a ∈ l₁, 125 | {have hsub₁ := @subset_of_diff_filter l₁ tl, 126 | have hsp := @subperm_cons_diff _ _ a (l₁.erase a) tl, 127 | have hsub₂ := subperm.subset hsp, 128 | have hsub := (perm.subset (perm.diff_right tl (perm_cons_erase ha))), 129 | intros x hx, 130 | cases hsub₁ hx with hxa, 131 | {left, exact hxa}, 132 | {exact hsub₂ (hsub h)}}, 133 | {rw erase_of_not_mem ha, apply subset_of_diff_filter}}, 134 | {simp [heq], apply subset_of_diff_filter} 135 | end 136 | 137 | end list 138 | 139 | 140 | /- Regular lemmas for the propositional part. -/ 141 | 142 | section 143 | variables (φ ψ : nnf) (Γ₁ Γ₂ Δ : list nnf) {st : Type} 144 | variables (k : kripke st) (s : st) 145 | open list 146 | 147 | theorem sat_subset (h₁ : Γ₁ ⊆ Γ₂) (h₂ : sat k s Γ₂) : sat k s Γ₁ := 148 | λ x hx, h₂ _ (h₁ hx) 149 | 150 | theorem unsat_subset (h₁ : Γ₁ ⊆ Γ₂) (h₂ : unsatisfiable Γ₁) : unsatisfiable Γ₂ := 151 | λ st k s h, (h₂ st k s (sat_subset _ Γ₂ _ _ h₁ h)) 152 | 153 | theorem sat_sublist (h₁ : Γ₁ <+ Γ₂) (h₂ :sat k s Γ₂) : sat k s Γ₁ := 154 | sat_subset _ _ _ _ (sublist.subset h₁) h₂ 155 | 156 | theorem unsat_sublist (h₁ : Γ₁ <+ Γ₂) (h₂ : unsatisfiable Γ₁) : unsatisfiable Γ₂ := 157 | λ st k s h, (h₂ st k s (sat_sublist _ Γ₂ _ _ h₁ h)) 158 | 159 | theorem unsat_contra {Δ n} : var n ∈ Δ → neg n ∈ Δ → unsatisfiable Δ:= 160 | begin 161 | intros h₁ h₂, intros v hsat, intros s hsat, 162 | have := hsat _ h₁, have := hsat _ h₂, simpa 163 | end 164 | 165 | theorem sat_of_and : force k s (and φ ψ) ↔ (force k s φ) ∧ (force k s ψ) := 166 | by split; {intro, simpa}; {intro, simpa} 167 | 168 | theorem sat_of_sat_erase (h₁ : sat k s $ Δ.erase φ) (h₂ : force k s φ) : sat k s Δ := 169 | begin 170 | intro ψ, intro h, 171 | by_cases (ψ = φ), 172 | {rw h, assumption}, 173 | {have : ψ ∈ Δ.erase φ, 174 | rw mem_erase_of_ne, assumption, exact h, 175 | apply h₁, assumption} 176 | end 177 | 178 | theorem unsat_and_of_unsat_split 179 | (h₁ : and φ ψ ∈ Δ) 180 | (h₂ : unsatisfiable $ φ :: ψ :: Δ.erase (and φ ψ)) : 181 | unsatisfiable Δ := 182 | begin 183 | intro st, intros, intro h, 184 | apply h₂, swap 3, exact k, swap, exact s, 185 | intro e, intro he, 186 | cases he, 187 | {rw he, have := h _ h₁, rw sat_of_and at this, exact this.1}, 188 | {cases he, 189 | {rw he, have := h _ h₁, rw sat_of_and at this, exact this.2}, 190 | {have := h _ h₁, apply h, apply mem_of_mem_erase he} } 191 | end 192 | 193 | theorem sat_and_of_sat_split 194 | (h₁ : and φ ψ ∈ Δ) 195 | (h₂ : sat k s $ φ :: ψ :: Δ.erase (and φ ψ)) : 196 | sat k s Δ := 197 | begin 198 | intro e, intro he, 199 | by_cases (e = and φ ψ), 200 | { rw h, split, repeat {apply h₂, simp} }, 201 | { have : e ∈ Δ.erase (and φ ψ), 202 | { rw mem_erase_of_ne, repeat { assumption } }, 203 | apply h₂, simp [this] } 204 | end 205 | 206 | theorem unsat_or_of_unsat_split 207 | (h : or φ ψ ∈ Δ) 208 | (h₁ : unsatisfiable $ φ :: Δ.erase (nnf.or φ ψ)) 209 | (h₂ : unsatisfiable $ ψ :: Δ.erase (nnf.or φ ψ)) : 210 | unsatisfiable $ Δ := 211 | begin 212 | intro, intros, intro hsat, 213 | have := hsat _ h, 214 | cases this, 215 | {apply h₁, swap 3, exact k, swap, exact s, intro e, intro he, 216 | cases he, rw he, exact this, apply hsat, apply mem_of_mem_erase he}, 217 | {apply h₂, swap 3, exact k, swap, exact s, intro e, intro he, 218 | cases he, rw he, exact this, apply hsat, apply mem_of_mem_erase he} 219 | end 220 | 221 | theorem sat_or_of_sat_split_left 222 | (h : or φ ψ ∈ Δ) 223 | (hl : sat k s $ φ :: Δ.erase (nnf.or φ ψ)) : 224 | sat k s Δ := 225 | begin 226 | intros e he, 227 | by_cases (e = or φ ψ), 228 | { rw h, left, apply hl, simp}, 229 | {have : e ∈ Δ.erase (or φ ψ), 230 | { rw mem_erase_of_ne, repeat { assumption } }, 231 | apply hl, simp [this]} 232 | end 233 | 234 | theorem sat_or_of_sat_split_right 235 | (h : or φ ψ ∈ Δ) 236 | (hl : sat k s $ ψ :: Δ.erase (nnf.or φ ψ)) : 237 | sat k s Δ := 238 | begin 239 | intros e he, 240 | by_cases (e = or φ ψ), 241 | { rw h, right, apply hl, simp}, 242 | { have : e ∈ Δ.erase (or φ ψ), 243 | { rw mem_erase_of_ne, repeat { assumption } }, 244 | apply hl, simp [this] } 245 | end 246 | 247 | end 248 | 249 | def unmodal_jump (Γ : list nnf) : ∀ (i : list nnf), i ∈ unmodal Γ → Π Δ st (k : kripke st) s 250 | (hsat : sat k s (Γ.diff Δ)) (hdia : dia i.head ∉ Δ), 251 | ∃ s, sat k s (i.diff (list.filter (≠ i.head) (unbox Δ))) := 252 | list.mapp _ _ 253 | begin 254 | intros x hx Δ st k s hsat hdia, 255 | rw list.cons_diff_of_ne_mem, swap, 256 | {intro hmem, rw [list.mem_filter] at hmem, have := hmem.2, 257 | simp at this, exact this}, 258 | { rw [←undia_iff] at hx, 259 | have := hsat _ (list.mem_diff_of_mem hx hdia), 260 | rcases this with ⟨w, hw⟩, 261 | split, swap, {exact w}, 262 | { apply sat_subset, swap 3, 263 | { exact x::list.diff (unbox Γ) (unbox Δ) }, 264 | { intros b hb, cases hb, 265 | { simp [hb] }, 266 | { apply list.subset_of_diff_filter, exact hb } }, 267 | { rw unbox_diff, intros c hc, cases hc, 268 | {rw hc, exact hw.2}, 269 | {have := (@unbox_iff (list.diff Γ Δ) c).2 hc, 270 | have hforce := hsat _ this, apply hforce, exact hw.1} } } } 271 | end 272 | 273 | /- Part of the soundness -/ 274 | 275 | theorem unsat_of_closed_and {Γ Δ} (i : and_instance Γ Δ) (h : unsatisfiable Δ) : unsatisfiable Γ := 276 | by cases i; { apply unsat_and_of_unsat_split, repeat {assumption} } 277 | 278 | theorem unsat_of_closed_or {Γ₁ Γ₂ Δ : list nnf} (i : or_instance Δ Γ₁ Γ₂) (h₁ : unsatisfiable Γ₁) (h₂ : unsatisfiable Γ₂) : unsatisfiable Δ := 279 | by cases i; {apply unsat_or_of_unsat_split, repeat {assumption} } 280 | 281 | /- Tree models -/ 282 | 283 | inductive model 284 | | cons : list ℕ → list model → model 285 | 286 | instance : decidable_eq model := by tactic.mk_dec_eq_instance 287 | 288 | instance inhabited_model : inhabited model := ⟨model.cons [] []⟩ 289 | 290 | open model 291 | 292 | @[simp] def mval : ℕ → model → bool 293 | | p (cons v r) := p ∈ v 294 | 295 | @[simp] def mrel : model → model → bool 296 | | (cons v r) m := m ∈ r 297 | 298 | theorem mem_of_mrel_tt : Π {v r m}, mrel (cons v r) m = tt → m ∈ r := 299 | begin 300 | intros v r m h, by_contradiction, 301 | simpa using h 302 | end 303 | 304 | @[simp] def builder : kripke model := 305 | {val := λ n s, mval n s, rel := λ s₁ s₂, mrel s₁ s₂} 306 | 307 | inductive batch_sat : list model → list (list nnf) → Prop 308 | | bs_nil : batch_sat [] [] 309 | | bs_cons (m Γ l₁ l₂) : sat builder m Γ → batch_sat l₁ l₂ → 310 | batch_sat (m::l₁) (Γ::l₂) 311 | 312 | open batch_sat 313 | 314 | theorem bs_ex : Π l Γ, 315 | batch_sat l Γ → ∀ m ∈ l, ∃ i ∈ Γ, sat builder m i 316 | | l Γ bs_nil := λ m hm, by simpa using hm 317 | | l Γ (bs_cons m Δ l₁ l₂ h hbs) := 318 | begin 319 | intros n hn, 320 | cases hn, 321 | {split, swap, exact Δ, split, simp, rw hn, exact h}, 322 | {have : ∃ (i : list nnf) (H : i ∈ l₂), sat builder n i, 323 | {apply bs_ex, exact hbs, exact hn}, 324 | {rcases this with ⟨w, hw, hsat⟩, split, swap, exact w, split, 325 | {simp [hw]}, {exact hsat} } } 326 | end 327 | 328 | theorem bs_forall : Π l Γ, 329 | batch_sat l Γ → ∀ i ∈ Γ, ∃ m ∈ l, sat builder m i 330 | | l Γ bs_nil := λ m hm, by simpa using hm 331 | | l Γ (bs_cons m Δ l₁ l₂ h hbs) := 332 | begin 333 | intros i hi, 334 | cases hi, {split, swap, exact m, split, simp, rw hi, exact h}, 335 | {have : ∃ (n : model) (H : n ∈ l₁), sat builder n i, 336 | {apply bs_forall, exact hbs, exact hi}, 337 | {rcases this with ⟨w, hw, hsat⟩, split, swap, exact w, split, 338 | {simp [hw]}, {exact hsat} } } 339 | end 340 | 341 | theorem sat_of_batch_sat : Π l Γ (h : modal_applicable Γ), 342 | batch_sat l (unmodal Γ) → sat builder (cons h.v l) Γ := 343 | begin 344 | intros l Γ h hbs φ hφ, 345 | cases hfml : φ, 346 | case nnf.var : n {rw hfml at hφ, simp, rw ←h.hv, exact hφ}, 347 | case nnf.box : ψ 348 | {intros s' hs', have hmem := mem_of_mrel_tt hs', 349 | have := bs_ex l (unmodal Γ) hbs s' hmem, 350 | rcases this with ⟨w, hw, hsat⟩, 351 | have := unmodal_mem_box Γ w hw ψ _, 352 | swap, {rw ←hfml, exact hφ}, {apply hsat, exact this} }, 353 | case nnf.dia : ψ 354 | {dsimp, 355 | have := bs_forall l (unmodal Γ) hbs (ψ :: unbox Γ) _, swap, 356 | {apply mem_unmodal, rw [←undia_iff, ←hfml], exact hφ}, 357 | {rcases this with ⟨w, hw, hsat⟩, split, swap, exact w, split, 358 | {simp [hw]}, {apply hsat, simp} } }, 359 | case nnf.neg : n 360 | {rw hfml at hφ, have : var n ∉ Γ, 361 | {intro hin, have := h.no_contra, have := this hin, contradiction}, 362 | simp, rw ←h.hv, exact this }, 363 | case nnf.and : φ ψ 364 | {rw hfml at hφ, have := h.no_and, have := @this φ ψ, contradiction}, 365 | case nnf.or : φ ψ 366 | {rw hfml at hφ, have := h.no_or, have := @this φ ψ, contradiction} 367 | end 368 | 369 | theorem build_model : Π Γ (h : model_constructible Γ), 370 | sat builder (cons h.v []) Γ := 371 | begin 372 | intros, intro, intro hmem, 373 | cases heq : φ, 374 | case nnf.var : n {rw [heq,h.hv] at hmem, simp [hmem]}, 375 | case nnf.neg : n {rw heq at hmem, simp, rw ←h.hv, intro hin, exfalso, apply h.no_contra, repeat {assumption} }, 376 | case nnf.box : φ {simp}, 377 | case nnf.and : φ ψ { rw heq at hmem, exfalso, apply h.no_and, assumption}, 378 | case nnf.or : φ ψ { rw heq at hmem, exfalso, apply h.no_or, assumption}, 379 | case nnf.dia : φ { rw heq at hmem, exfalso, apply h.no_dia, assumption}, 380 | end 381 | -------------------------------------------------------------------------------- /src/K/size.lean: -------------------------------------------------------------------------------- 1 | import defs 2 | open list 3 | 4 | lemma size_erase_add_eq_size_sum {φ} : Π (Γ : list nnf) (h : φ ∈ Γ), node_size (Γ.erase φ) + sizeof φ = node_size Γ 5 | | [] h := absurd h $ not_mem_nil _ 6 | | (hd :: tl) h := 7 | begin 8 | by_cases eq : φ = hd, 9 | {dsimp [list.erase], rw if_pos, rw eq, apply add_comm, rw eq}, 10 | {dsimp [list.erase], rw if_neg, dsimp, rw [add_assoc], have : node_size (list.erase tl φ) + sizeof φ = node_size tl, 11 | {apply size_erase_add_eq_size_sum, cases h, contradiction, assumption}, 12 | rw this, intro h', rw h' at eq, contradiction} 13 | end 14 | 15 | theorem split_lt_and {φ ψ} (Γ : list nnf) (h : nnf.and φ ψ ∈ Γ) : 16 | node_size (φ :: ψ :: Γ.erase (nnf.and φ ψ)) < node_size Γ := 17 | begin 18 | dsimp [node_size], 19 | rw ←size_erase_add_eq_size_sum Γ, 20 | swap, exact h, 21 | rw [←add_assoc, add_comm], 22 | apply add_lt_add_left, dsimp [sizeof, has_sizeof.sizeof, nnf.sizeof], 23 | rw [add_assoc], 24 | rw [nat.one_add], 25 | apply nat.lt_succ_self 26 | end 27 | 28 | theorem split_lt_or_left {φ ψ} (Γ : list nnf) (h : nnf.or φ ψ ∈ Γ) : 29 | node_size (φ :: Γ.erase (nnf.or φ ψ)) < node_size Γ := 30 | begin 31 | dsimp, 32 | rw ←size_erase_add_eq_size_sum Γ, 33 | swap, exact h, 34 | rw [add_comm], 35 | apply add_lt_add_left, dsimp [sizeof, has_sizeof.sizeof, nnf.sizeof], 36 | rw [add_comm, ←add_assoc], 37 | apply nat.lt_add_of_pos_left, 38 | apply nat.succ_pos 39 | end 40 | 41 | theorem split_lt_or_right {φ ψ} (Γ : list nnf) (h : nnf.or φ ψ ∈ Γ) : 42 | node_size (ψ :: Γ.erase (nnf.or φ ψ)) < node_size Γ := 43 | begin 44 | dsimp, 45 | rw ←size_erase_add_eq_size_sum Γ, 46 | swap, exact h, 47 | rw [add_comm], 48 | apply add_lt_add_left, dsimp [sizeof, has_sizeof.sizeof, nnf.sizeof], 49 | apply nat.lt_add_of_pos_left, 50 | rw nat.one_add, 51 | apply nat.succ_pos 52 | end 53 | -------------------------------------------------------------------------------- /src/K/tableau.lean: -------------------------------------------------------------------------------- 1 | import .full_language .jump 2 | 3 | def fml_is_sat (Γ : list fml) : bool := 4 | match tableau (list.map fml.to_nnf Γ) with 5 | | node.closed _ _ _ := ff 6 | | node.open_ _ := tt 7 | end 8 | 9 | theorem classical_correctness (Γ : list fml) : 10 | fml_is_sat Γ = tt ↔ ∃ (st : Type) (k : kripke st) s, fml_sat k s Γ := 11 | begin 12 | cases h : fml_is_sat Γ, 13 | constructor, 14 | { intro, contradiction }, 15 | { intro hsat, cases eq : tableau (list.map fml.to_nnf Γ), 16 | rcases hsat with ⟨w, k, s, hsat⟩, 17 | apply false.elim, apply a, rw trans_sat_iff at hsat, exact hsat, 18 | { dsimp [fml_is_sat] at h, rw eq at h, contradiction } }, 19 | { split, intro, dsimp [fml_is_sat] at h, 20 | cases eq : tableau (list.map fml.to_nnf Γ), 21 | { rw eq at h, contradiction }, 22 | { split, split, split, have := a_1.2, rw ←trans_sat_iff at this, 23 | exact this }, 24 | { simp } } 25 | end 26 | 27 | open fml 28 | 29 | def fml.exm : fml := or (var 1) (neg (var 1)) 30 | 31 | def K : fml := impl (box (impl (var 1) (var 2))) (impl (box $ var 1) (box $ var 2)) 32 | 33 | def tst : list fml := [dia $ and (var 0) (neg $ var 1), dia $ var 1, box $ var 2, and (var 1) (var 2)] 34 | 35 | #eval fml_is_sat tst 36 | -------------------------------------------------------------------------------- /src/K/test.lean: -------------------------------------------------------------------------------- 1 | import .tableau 2 | open fml 3 | 4 | -- A simple workaround for fml.iff. 5 | -- Unlike fml.impl, this is really a trusted string-level translation 6 | -- similar to a bash script modifiying the benchmark text. 7 | -- Only k_poly_p and k_poly_n contain iffs. 8 | 9 | def fml.iff (φ ψ : fml) := and (impl φ ψ) (impl ψ φ) 10 | 11 | def test (ϕ : fml) := fml_is_sat [neg ϕ] 12 | 13 | local infix ` ∨ `:1000 := fml.or 14 | local infix ` ∧ `:1000 := fml.and 15 | local infix ` -> `:1000 := fml.impl 16 | local infix ` <-> `:1000 := fml.iff 17 | 18 | local prefix `~` := fml.neg 19 | 20 | precedence `p`:max 21 | local prefix ` p ` := fml.var 22 | 23 | set_option profiler true 24 | 25 | -- #eval test () 26 | -------------------------------------------------------------------------------- /src/KT/KT_defs.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2018-2019 Minchao Wu. All rights reserved. 3 | Released under MIT license as described in the file LICENSE. 4 | Author: Minchao Wu 5 | -/ 6 | 7 | import defs 8 | open nnf tactic 9 | 10 | meta def frame_restriction : tactic unit := 11 | do intro `a >> `[simp] 12 | 13 | structure KT (states : Type) extends kripke states := 14 | (refl : reflexive rel . frame_restriction) 15 | 16 | instance inhabited_KT : inhabited (KT ℕ) := 17 | ⟨{ val := λ a b, tt, rel := λ a b, tt }⟩ 18 | 19 | @[simp] def force {states : Type} (k : KT states) : states → nnf → Prop 20 | | s (var n) := k.val n s 21 | | s (neg n) := ¬ k.val n s 22 | | s (and φ ψ) := force s φ ∧ force s ψ 23 | | s (or φ ψ) := force s φ ∨ force s ψ 24 | | s (box φ) := ∀ s', k.rel s s' → force s' φ 25 | | s (dia φ) := ∃ s', k.rel s s' ∧ force s' φ 26 | 27 | def sat {st} (k : KT st) (s) (Γ : list nnf) : Prop := 28 | ∀ φ ∈ Γ, force k s φ 29 | 30 | def unsatisfiable (Γ : list nnf) : Prop := 31 | ∀ (st) (k : KT st) s, ¬ sat k s Γ 32 | 33 | theorem unsat_singleton {φ} : unsatisfiable [φ] → ∀ (st) (k : KT st) s, ¬ force k s φ 34 | := 35 | begin 36 | intro h, intros, intro hf, 37 | apply h, intros ψ hψ, rw list.mem_singleton at hψ, rw hψ, exact hf 38 | end 39 | 40 | theorem sat_of_empty {st} (k : KT st) (s) : sat k s [] := 41 | λ φ h, absurd h $ list.not_mem_nil _ 42 | 43 | theorem ne_empty_of_unsat {Γ} (h : unsatisfiable Γ): Γ ≠ [] := 44 | begin 45 | intro heq, rw heq at h, 46 | apply h, apply sat_of_empty, exact nat, 47 | apply inhabited_KT.1, exact 0 48 | end 49 | 50 | inductive model 51 | | cons : list ℕ → list model → model 52 | 53 | instance : decidable_eq model := by tactic.mk_dec_eq_instance 54 | 55 | open model 56 | 57 | @[simp] def mval : ℕ → model → bool 58 | | p (cons v r) := p ∈ v 59 | 60 | @[simp] def mrel : model → model → bool 61 | | m₁@(cons v r) m₂ := m₂ ∈ r ∨ m₁ = m₂ 62 | 63 | theorem refl_mrel (s : model) : mrel s s := by cases s with v r; simp 64 | 65 | theorem mem_of_mrel_tt : Π {v r m}, mrel (cons v r) m = tt → m ∈ r ∨ cons v r = m := 66 | begin 67 | intros v r m h, by_cases hc : cons v r = m, 68 | {right, exact hc},{left, by_contradiction hn, 69 | have : mrel (cons v r) m = ff, { simp [hc, hn] }, 70 | rw h at this, contradiction} 71 | end 72 | 73 | -- TODO : make this neater. Currently it's just a copy-paste 74 | theorem mem_of_mrel_empty : Π {v m}, mrel (cons v []) m = tt → cons v [] = m := 75 | begin 76 | intros v m h, by_cases hc : cons v [] = m, 77 | {exact hc}, 78 | {by_contradiction hn, 79 | have : mrel (cons v []) m = ff, { simp [hc, hn] }, 80 | rw h at this, contradiction} 81 | end 82 | 83 | @[simp] def builder : KT model := 84 | {val := λ n s, mval n s, rel := λ s₁ s₂, mrel s₁ s₂, refl := refl_mrel} 85 | 86 | theorem force_box_of_leaf {v φ} (h : force builder (cons v []) φ): 87 | force builder (cons v []) (box φ) := 88 | begin 89 | dsimp, intros s' hs', 90 | have : cons v [] = s', {apply mem_of_mrel_empty, exact hs'}, 91 | rw ←this, assumption 92 | end 93 | 94 | def srefl (m h : list nnf) := 95 | ∀ {v l φ}, sat builder (cons v l) m → 96 | box φ ∈ h → 97 | (∀ ψ, box ψ ∈ h → ∀ m ∈ l, force builder m ψ) → 98 | force builder (cons v l) φ 99 | 100 | structure seqt : Type := 101 | (main : list nnf) 102 | (hdld : list nnf) 103 | -- srefl main hdld says that sequent hdld | main satisfies theorem 3.7(2) 104 | (pmain : srefl main hdld) 105 | -- box_only says there are only boxed formulas in hdld 106 | (phdld : box_only hdld) 107 | 108 | class val_constructible (Γ : seqt) := 109 | (satu : saturated Γ.main) 110 | (no_box_main : ∀ {φ}, box φ ∉ Γ.main) 111 | (no_contra_main : ∀ {n}, var n ∈ Γ.main → neg n ∉ Γ.main) 112 | (v : list ℕ) 113 | (hv : ∀ n, var n ∈ Γ.main ↔ n ∈ v) 114 | 115 | class modal_applicable (Γ : seqt) extends val_constructible Γ := 116 | (φ : nnf) 117 | (ex : dia φ ∈ Γ.main) 118 | 119 | class model_constructible (Γ : seqt) extends val_constructible Γ := 120 | (no_dia : ∀ {φ}, nnf.dia φ ∉ Γ.main) 121 | 122 | theorem build_model : Π Γ (h : model_constructible Γ), 123 | sat builder (cons h.v []) Γ.main := 124 | begin 125 | intros, intro, intro hmem, 126 | cases heq : φ, 127 | case nnf.var : n {dsimp, have := h.hv, rw heq at hmem, rw this at hmem, simp [hmem]}, 128 | case nnf.neg : n {dsimp, have h₁ := h.hv, rw heq at hmem, have := h.no_contra_main, simp, rw ←h₁, intro hvar, apply this, swap, exact hmem, exact hvar}, 129 | case nnf.box : ψ {dsimp, intros, have := h.no_box_main, exfalso, rw heq at hmem, exact this hmem}, 130 | case nnf.and : φ ψ { rw heq at hmem, have := h.satu.no_and, have := @this φ ψ, contradiction}, 131 | case nnf.or : φ ψ { rw heq at hmem, have := h.satu.no_or, have := @this φ ψ, contradiction}, 132 | case nnf.dia : φ { rw heq at hmem, have := h.no_dia, have := @this φ, contradiction}, 133 | end 134 | 135 | 136 | /- Regular lemmas for the propositional part. -/ 137 | 138 | section 139 | variables (φ ψ : nnf) (Γ₁ Γ₂ Δ Λ: list nnf) {st : Type} 140 | variables (k : KT st) (s : st) 141 | open list 142 | 143 | theorem sat_subset (h₁ : Γ₁ ⊆ Γ₂) (h₂ : sat k s Γ₂) : sat k s Γ₁ := 144 | λ x hx, h₂ _ (h₁ hx) 145 | 146 | theorem sat_sublist (h₁ : Γ₁ <+ Γ₂) (h₂ :sat k s Γ₂) : sat k s Γ₁ := 147 | sat_subset _ _ _ _ (sublist.subset h₁) h₂ 148 | 149 | theorem sat_append (h₁ : sat k s Γ₁) (h₂ : sat k s Γ₂) : sat k s (Γ₁ ++ Γ₂) := 150 | begin 151 | intros φ h, rw mem_append at h, cases h, 152 | apply h₁ _ h, apply h₂ _ h 153 | end 154 | 155 | theorem unsat_contra {Δ n} : var n ∈ Δ → neg n ∈ Δ → unsatisfiable Δ:= 156 | begin 157 | intros h₁ h₂, intros v hsat, intros s hsat, 158 | have := hsat _ h₁, have := hsat _ h₂, simpa 159 | end 160 | 161 | theorem unsat_contra_seqt {Δ : seqt} {n} : var n ∈ Δ.main → neg n ∈ Δ.main → unsatisfiable (Δ.main ++ Δ.hdld):= 162 | begin 163 | intros h₁ h₂, intros st m, intros s hsat, 164 | have := unsat_contra h₁ h₂, 165 | have := this _ m s, 166 | apply this, 167 | apply sat_subset _ _ _ _ _ hsat, 168 | simp 169 | end 170 | 171 | theorem sat_of_and : force k s (and φ ψ) ↔ (force k s φ) ∧ (force k s ψ) := 172 | by split; {intro, simpa} 173 | 174 | theorem sat_of_sat_erase (h₁ : sat k s $ Δ.erase φ) (h₂ : force k s φ) : sat k s Δ := 175 | begin 176 | intro ψ, intro h, 177 | by_cases (ψ = φ), 178 | {rw h, assumption}, 179 | {have : ψ ∈ Δ.erase φ, 180 | rw mem_erase_of_ne, assumption, exact h, 181 | apply h₁, assumption} 182 | end 183 | 184 | theorem unsat_and_of_unsat_split 185 | (h₁ : and φ ψ ∈ Δ) 186 | (h₂ : unsatisfiable $ φ :: ψ :: Δ.erase (and φ ψ)) : 187 | unsatisfiable Δ := 188 | begin 189 | intro st, intros, intro h, 190 | apply h₂, swap 3, exact k, swap, exact s, 191 | intro e, intro he, 192 | cases he, 193 | {rw he, have := h _ h₁, rw sat_of_and at this, exact this.1}, 194 | {cases he, 195 | {rw he, have := h _ h₁, rw sat_of_and at this, exact this.2}, 196 | {have := h _ h₁, apply h, apply mem_of_mem_erase he} } 197 | end 198 | 199 | theorem unsat_and_of_unsat_split_seqt {Γ} 200 | (h₁ : and φ ψ ∈ Δ) 201 | (h₂ : unsatisfiable $ (φ :: ψ :: Δ.erase (and φ ψ)++Γ)) : 202 | unsatisfiable (Δ++Γ) := 203 | begin 204 | intro st, intros, intro h, 205 | apply h₂, swap 3, exact k, swap, exact s, 206 | intro e, intro he, 207 | cases he, 208 | {rw he, have := h _ (mem_append_left _ h₁), rw sat_of_and at this, exact this.1}, 209 | {cases he, 210 | {rw he, have := h _ (mem_append_left _ h₁), rw sat_of_and at this, exact this.2}, 211 | {have := h _ (mem_append_left _ h₁), apply h, apply mem_of_mem_erase, rw erase_append_left, exact he, exact h₁} } 212 | end 213 | 214 | theorem sat_and_of_sat_split 215 | (h₁ : and φ ψ ∈ Δ) 216 | (h₂ : sat k s $ φ :: ψ :: Δ.erase (and φ ψ)) : 217 | sat k s Δ := 218 | begin 219 | intro e, intro he, 220 | by_cases (e = and φ ψ), 221 | { rw h, dsimp, split, repeat {apply h₂, simp} }, 222 | { have : e ∈ Δ.erase (and φ ψ), 223 | { rw mem_erase_of_ne, repeat { assumption } }, 224 | apply h₂, simp [this] } 225 | end 226 | 227 | theorem sat_and_of_sat_split_seqt {Γ} 228 | (h₁ : and φ ψ ∈ Δ) 229 | (h₂ : sat k s $ (φ :: ψ :: Δ.erase (and φ ψ)++Γ)) : 230 | sat k s (Δ++Γ) := 231 | begin 232 | intro e, intro he, 233 | by_cases (e = and φ ψ), 234 | { rw h, dsimp, split, repeat {apply h₂, simp} }, 235 | { have : e ∈ Δ.erase (and φ ψ) ++ Γ, 236 | { rw ←erase_append_left, rw mem_erase_of_ne, repeat {assumption} }, 237 | apply h₂, simp [this] } 238 | end 239 | 240 | theorem sat_split_of_sat_and_seqt {Γ} 241 | (h₁ : and φ ψ ∈ Δ) 242 | (h₂ : sat k s (Δ++Γ)) : 243 | sat k s $ (φ :: ψ :: Δ.erase (and φ ψ)++Γ) := 244 | begin 245 | intros e he, rw mem_append at he, cases he, 246 | have : force k s (and φ ψ), {apply h₂, simp [h₁]}, rw sat_of_and at this, 247 | {cases he, 248 | {rw he, exact this.left}, 249 | {cases he, rw he, exact this.right, apply h₂, rw mem_append, left, apply mem_of_mem_erase he} 250 | }, 251 | {apply h₂, rw mem_append, right, exact he} 252 | end 253 | 254 | theorem unsat_or_of_unsat_split_seqt {Γ} 255 | (h : or φ ψ ∈ Δ) 256 | (h₁ : unsatisfiable $ (φ :: Δ.erase (nnf.or φ ψ)++Γ)) 257 | (h₂ : unsatisfiable $ (ψ :: Δ.erase (nnf.or φ ψ)++Γ)) : 258 | unsatisfiable $ (Δ++Γ) := 259 | begin 260 | intro, intros, intro hsat, 261 | have := hsat _ (mem_append_left _ h), 262 | dsimp at this, 263 | cases this, 264 | {apply h₁, swap 3, exact k, swap, exact s, intro e, intro he, 265 | cases he, rw he, exact this, apply hsat, 266 | apply mem_of_mem_erase, rw erase_append_left, exact he, exact h}, 267 | {apply h₂, swap 3, exact k, swap, exact s, intro e, intro he, 268 | cases he, rw he, exact this, apply hsat, apply mem_of_mem_erase, rw erase_append_left, exact he, exact h} 269 | end 270 | 271 | theorem sat_or_of_sat_split_left 272 | (h : or φ ψ ∈ Δ) 273 | (hl : sat k s $ φ :: Δ.erase (nnf.or φ ψ)) : 274 | sat k s Δ := 275 | begin 276 | intros e he, 277 | by_cases (e = or φ ψ), 278 | { rw h, dsimp, left, apply hl, simp}, 279 | {have : e ∈ Δ.erase (or φ ψ), 280 | { rw mem_erase_of_ne, repeat { assumption } }, 281 | apply hl, simp [this]} 282 | end 283 | 284 | theorem sat_or_of_sat_split_right 285 | (h : or φ ψ ∈ Δ) 286 | (hl : sat k s $ ψ :: Δ.erase (nnf.or φ ψ)) : 287 | sat k s Δ := 288 | begin 289 | intros e he, 290 | by_cases (e = or φ ψ), 291 | { rw h, dsimp, right, apply hl, simp}, 292 | { have : e ∈ Δ.erase (or φ ψ), 293 | { rw mem_erase_of_ne, repeat { assumption } }, 294 | apply hl, simp [this] } 295 | end 296 | 297 | /- KT-specific lemmas -/ 298 | 299 | theorem force_of_force_box (h : force k s $ box φ) : force k s φ 300 | := begin dsimp at h, apply h, apply k.refl end 301 | 302 | theorem unsat_copy_of_unsat_box 303 | (h₁ : box φ ∈ Δ) 304 | (h₂ : unsatisfiable $ (φ :: Δ.erase (box φ)) ++ box φ :: Λ) : 305 | unsatisfiable (Δ ++ Λ) := 306 | begin 307 | intros st k s h, 308 | apply h₂, swap 3, exact k, swap, exact s, 309 | intros e he, 310 | rw [mem_append] at he, 311 | cases he, 312 | {cases he, 313 | {rw he, apply force_of_force_box, apply h (box φ), simp [h₁]}, 314 | {apply h, rw mem_append, left, apply mem_of_mem_erase he }}, 315 | {cases he, 316 | {rw ←he at h₁, apply h, rw mem_append, left, exact h₁}, 317 | {apply h, rw mem_append, right, assumption}} 318 | end 319 | 320 | theorem sat_copy_of_sat_box 321 | (h₁ : box φ ∈ Δ) 322 | (h₂ : sat k s $ (φ :: Δ.erase (box φ)) ++ box φ :: Λ) : 323 | sat k s (Δ ++ Λ) := 324 | begin 325 | intros ψ hφ, 326 | rw mem_append at hφ, 327 | cases hφ, 328 | {by_cases heq : ψ = box φ, 329 | {rw heq, apply h₂ (box φ), simp}, 330 | {have := mem_erase_of_ne heq, rw ←this at hφ, apply h₂, simp, right, left, exact hφ}}, 331 | {apply h₂, simp, repeat {right}, exact hφ} 332 | end 333 | 334 | theorem sat_box_of_sat_copy 335 | (h₁ : box φ ∈ Δ) 336 | (h₂ : sat k s (Δ ++ Λ)) : 337 | sat k s $ (φ :: Δ.erase (box φ)) ++ box φ :: Λ := 338 | begin 339 | apply sat_append, 340 | {intros ψ h, cases h, have : sat k s Δ, {apply sat_subset Δ (Δ++Λ), simp, assumption}, have := this _ h₁, apply force_of_force_box, rw h, exact this, apply h₂, have := mem_of_mem_erase h, simp [this]}, 341 | {intros ψ h, cases h, rw h, have := h₂ (box φ), apply this, simp [h₁], apply h₂, rw mem_append, right, exact h} 342 | end 343 | 344 | end 345 | 346 | def and_child {φ ψ} (Γ : seqt) (h : nnf.and φ ψ ∈ Γ.main) : seqt := 347 | ⟨φ :: ψ :: Γ.main.erase (and φ ψ), Γ.hdld, 348 | begin 349 | intros k s γ hsat hin hall, 350 | by_cases heq : γ = and φ ψ, 351 | {rw heq, split, apply hsat, simp, apply hsat, simp}, 352 | {apply Γ.pmain _ hin hall, apply sat_and_of_sat_split _ _ _ _ _ h hsat} 353 | end, 354 | Γ.phdld⟩ 355 | 356 | inductive and_instance_seqt (Γ : seqt) : seqt → Type 357 | | cons : Π {φ ψ} (h : nnf.and φ ψ ∈ Γ.main), 358 | and_instance_seqt $ and_child Γ h 359 | 360 | def or_child_left {φ ψ} (Γ : seqt) (h : nnf.or φ ψ ∈ Γ.main) : seqt := 361 | ⟨φ :: Γ.main.erase (or φ ψ), Γ.hdld, 362 | begin 363 | intros k s γ hsat hin hall, 364 | by_cases heq : γ = or φ ψ, 365 | {rw heq, dsimp, left, apply hsat, simp}, 366 | {apply Γ.pmain, apply sat_or_of_sat_split_left, exact h, exact hsat, exact hin, exact hall} 367 | end, 368 | Γ.phdld⟩ 369 | 370 | def or_child_right {φ ψ} (Γ : seqt) (h : nnf.or φ ψ ∈ Γ.main) : seqt := 371 | ⟨ψ :: Γ.main.erase (or φ ψ), Γ.hdld, 372 | begin 373 | intros k s γ hsat hin hall, 374 | by_cases heq : γ = or φ ψ, 375 | {rw heq, dsimp, right, apply hsat, simp}, 376 | {apply Γ.pmain, apply sat_or_of_sat_split_right, exact h, exact hsat, exact hin, exact hall} 377 | end, 378 | Γ.phdld⟩ 379 | 380 | inductive or_instance_seqt (Γ : seqt) : seqt → seqt → Type 381 | | cons : Π {φ ψ} (h : nnf.or φ ψ ∈ Γ.main), 382 | or_instance_seqt (or_child_left Γ h) (or_child_right Γ h) 383 | 384 | def box_child {φ} (Γ : seqt) (h : nnf.box φ ∈ Γ.main) : seqt := 385 | ⟨φ :: Γ.main.erase (box φ), box φ :: Γ.hdld, 386 | begin 387 | intros k s γ hsat hin hall, 388 | cases hin, 389 | {simp at hin, rw hin, apply hsat, simp}, 390 | {apply Γ.pmain, intros ψ hψ, 391 | by_cases heq : ψ = box φ, 392 | {rw heq, dsimp, intros s' hs', 393 | have := mem_of_mrel_tt hs', cases this, 394 | {apply hall, simp, assumption}, 395 | {rw ←this, apply hsat, simp} }, 396 | {apply hsat, right, 397 | have := list.mem_erase_of_ne heq, swap, exact Γ.main, 398 | rw ←this at hψ, assumption}, 399 | {assumption},{intros, apply hall, simp [a], exact H} } 400 | end, 401 | cons_box_only Γ.phdld⟩ 402 | 403 | inductive copy_instance_seqt (Γ : seqt) : seqt → Type 404 | | cons : Π {φ} (h : nnf.box φ ∈ Γ.main), 405 | copy_instance_seqt $ box_child Γ h 406 | 407 | theorem build_model_seqt : Π Γ (h : model_constructible Γ), 408 | sat builder (cons h.v []) (Γ.main ++ Γ.hdld) := 409 | begin 410 | intros Γ h, apply sat_append, apply build_model, 411 | intros ψ hψ, 412 | have := box_only_ex Γ.phdld hψ, cases this with w hw, 413 | rw hw, apply force_box_of_leaf, apply Γ.pmain _ _, 414 | {intros ψ hin m hm, exfalso, apply list.not_mem_nil, exact hm}, 415 | {apply build_model}, 416 | {rw hw at hψ, assumption} 417 | end 418 | -------------------------------------------------------------------------------- /src/KT/ops.lean: -------------------------------------------------------------------------------- 1 | import .size 2 | open nnf list 3 | 4 | namespace list 5 | universes u v w 6 | 7 | variables {α : Type u} {β : Type v} {γ : Type w} 8 | 9 | theorem mapp {p : β → Prop} (f : α → β) : Π (l : list α) (h : ∀ x∈l, p (f x)) x, x ∈ list.map f l → p x 10 | | [] h x := by simp 11 | | (hd::tl) h x := 12 | begin 13 | intro hmem, cases hmem, 14 | {rw hmem, apply h, simp}, 15 | {apply mapp tl, intros a ha, apply h, simp [ha], exact hmem} 16 | end 17 | 18 | @[simp] def ne_empty_head : Π l : list α, l ≠ [] → α 19 | | [] h := by contradiction 20 | | (a :: l) h := a 21 | 22 | end list 23 | 24 | @[simp] def unbox : list nnf → list nnf 25 | | [] := [] 26 | | ((box φ) :: l) := φ :: unbox l 27 | | (e :: l) := unbox l 28 | 29 | theorem unbox_iff : Π {Γ φ}, box φ ∈ Γ ↔ φ ∈ unbox Γ 30 | | [] φ := begin split, repeat {intro h, simpa using h} end 31 | | (hd::tl) φ := 32 | begin 33 | split, 34 | { intro h, cases h₁ : hd, 35 | case nnf.box : ψ 36 | { dsimp [unbox], cases h, 37 | {left, rw h₁ at h, injection h}, 38 | {right, exact (@unbox_iff tl φ).1 h} }, 39 | all_goals 40 | { dsimp [unbox], cases h, 41 | {rw h₁ at h, contradiction}, 42 | {exact (@unbox_iff tl φ).1 h} } }, 43 | { intro h, cases h₁ : hd, 44 | case nnf.box : ψ 45 | { rw h₁ at h, dsimp [unbox] at h, cases h, 46 | {simp [h]}, {right, exact (@unbox_iff tl φ).2 h} }, 47 | all_goals 48 | { rw h₁ at h, dsimp [unbox] at h, right, exact (@unbox_iff tl φ).2 h } } 49 | end 50 | 51 | theorem unbox_size_aux : Π {Γ}, node_size (unbox Γ) ≤ node_size Γ 52 | | [] := by simp 53 | | (hd::tl) := 54 | begin 55 | cases h : hd, 56 | case nnf.box : ψ 57 | { dsimp, apply add_le_add, 58 | { dsimp [sizeof, has_sizeof.sizeof, nnf.sizeof], 59 | rw add_comm, apply nat.le_succ }, 60 | { apply unbox_size_aux } }, 61 | all_goals 62 | { dsimp, apply le_add_of_nonneg_of_le, 63 | { dsimp [sizeof, has_sizeof.sizeof, nnf.sizeof], rw add_comm, apply nat.zero_le }, 64 | { apply unbox_size_aux } } 65 | end 66 | 67 | theorem le_of_unbox_degree : Π {Γ}, modal_degree (unbox Γ) ≤ modal_degree Γ 68 | | [] := by simp 69 | | (hd::tl) := 70 | begin 71 | cases heq : hd, 72 | case nnf.box : ψ 73 | {dsimp, rw cons_degree, rw cons_degree, 74 | apply max_le_max, { simp }, { apply le_of_unbox_degree } }, 75 | all_goals 76 | {dsimp, rw cons_degree, apply le_max_right_of_le, apply le_of_unbox_degree} 77 | end 78 | 79 | theorem zero_degree_of_eq_unbox : Π {Γ}, modal_degree (unbox Γ) = modal_degree Γ → modal_degree Γ = 0 80 | | [] h := begin dsimp [maximum, modal_degree], refl end 81 | | (hd::tl) h := 82 | begin 83 | cases heq : hd, 84 | case nnf.box : ψ 85 | {rw heq at h, dsimp at h, rw cons_degree at h, rw cons_degree at h, 86 | have := @le_of_unbox_degree tl, 87 | cases (lt_or_eq_of_le this) with h₁ h₂, 88 | {have hne : max (count_modal ψ) (modal_degree (unbox tl)) ≠ max (count_modal (box ψ)) (modal_degree tl), 89 | { apply ne_of_lt, apply max_lt_max, {simp}, {exact h₁} }, 90 | contradiction }, 91 | {rw [h₂, zero_degree_of_eq_unbox h₂] at h, 92 | have : max (count_modal ψ) 0 < max (count_modal (box ψ)) 0, 93 | {simp [_root_.max], by_cases hz : count_modal ψ = 0, 94 | repeat {simp [hz]} }, 95 | have hne := ne_of_lt this, contradiction } }, 96 | case nnf.dia : ψ 97 | {rw heq at h, dsimp at h, rw cons_degree at h, 98 | have := @le_of_unbox_degree tl, 99 | cases (lt_or_eq_of_le this) with h₁ h₂, 100 | {have hne : modal_degree (unbox tl) ≠ max (count_modal (box ψ)) (modal_degree tl), 101 | { apply ne_of_lt, rw lt_max_iff, right, exact h₁}, 102 | contradiction }, 103 | {have : modal_degree (unbox tl) < max (count_modal (dia ψ)) (modal_degree tl), 104 | {rw [h₂, zero_degree_of_eq_unbox h₂], 105 | simp [_root_.max] }, 106 | have hne := ne_of_lt this, contradiction } }, 107 | case nnf.var : ψ 108 | {rw heq at h, dsimp at h, rw cons_degree at h, 109 | have hle : 0 ≤ modal_degree tl, { simp }, 110 | have hz : modal_degree tl = 0, {apply zero_degree_of_eq_unbox, dsimp [_root_.max] at h, rw h, apply if_pos hle }, 111 | rw cons_degree, simp [_root_.max, hz] }, 112 | case nnf.neg : ψ 113 | {rw heq at h, dsimp at h, rw cons_degree at h, 114 | have hle : 0 ≤ modal_degree tl, { simp }, 115 | have hz : modal_degree tl = 0, {apply zero_degree_of_eq_unbox, dsimp [_root_.max] at h, rw h, apply if_pos hle }, 116 | rw cons_degree, simp [_root_.max, hz] }, 117 | case nnf.and : ψ₁ ψ₂ 118 | {rw heq at h, dsimp at h, rw cons_degree at h, 119 | by_cases hc : count_modal (and ψ₁ ψ₂) = 0, 120 | {rw hc at h, 121 | have hle : 0 ≤ modal_degree tl, { simp }, 122 | have hz : modal_degree tl = 0, {apply zero_degree_of_eq_unbox, dsimp [_root_.max] at h, rw h, apply if_pos hle }, 123 | rw cons_degree, simp [_root_.max, hc, hz] }, -- eq case end 124 | {have := @le_of_unbox_degree tl, 125 | cases (lt_or_eq_of_le this) with h₁ h₂, 126 | {have hne : modal_degree (unbox tl) ≠ max (count_modal (and ψ₁ ψ₂)) (modal_degree tl), 127 | { apply ne_of_lt, rw lt_max_iff, right, exact h₁}, 128 | contradiction }, 129 | {have : modal_degree (unbox tl) < max (count_modal (and ψ₁ ψ₂)) (modal_degree tl), 130 | {rw [h₂, zero_degree_of_eq_unbox h₂], 131 | rw lt_max_iff, left, exact nat.pos_of_ne_zero hc}, 132 | have hne := ne_of_lt this, contradiction} } }, -- lt case end 133 | case nnf.or : ψ₁ ψ₂ 134 | {rw heq at h, dsimp at h, rw cons_degree at h, 135 | by_cases hc : count_modal (or ψ₁ ψ₂) = 0, 136 | {rw hc at h, 137 | have hle : 0 ≤ modal_degree tl, { simp }, 138 | have hz : modal_degree tl = 0, {apply zero_degree_of_eq_unbox, dsimp [_root_.max] at h, rw h, apply if_pos hle }, 139 | rw cons_degree, simp [_root_.max, hc, hz] }, 140 | {have := @le_of_unbox_degree tl, 141 | cases (lt_or_eq_of_le this) with h₁ h₂, 142 | {have hne : modal_degree (unbox tl) ≠ max (count_modal (and ψ₁ ψ₂)) (modal_degree tl), 143 | { apply ne_of_lt, rw lt_max_iff, right, exact h₁}, 144 | contradiction }, 145 | {have : modal_degree (unbox tl) < max (count_modal (and ψ₁ ψ₂)) (modal_degree tl), 146 | {rw [h₂, zero_degree_of_eq_unbox h₂], 147 | rw lt_max_iff, left, exact nat.pos_of_ne_zero hc}, 148 | have hne := ne_of_lt this, contradiction} } } 149 | end 150 | 151 | theorem unbox_degree_aux : Π {Γ φ}, dia φ ∈ Γ → modal_degree (unbox Γ) < modal_degree Γ 152 | | [] φ h := absurd h $ not_mem_nil _ 153 | | (hd::tl) φ h := 154 | begin 155 | cases heq : hd, 156 | case nnf.box : ψ 157 | {dsimp, rw cons_degree, rw cons_degree, 158 | apply max_lt_max, 159 | { simp }, 160 | { apply unbox_degree_aux, swap, exact φ, 161 | cases h, {rw heq at h, contradiction}, {exact h} } }, 162 | case nnf.dia : ψ 163 | {dsimp [-modal_degree], rw cons_degree, 164 | have := @le_of_unbox_degree tl, 165 | cases (lt_or_eq_of_le this) with h₁ h₂, 166 | {rw lt_max_iff, right, exact h₁}, 167 | {rw [h₂, zero_degree_of_eq_unbox h₂], simp [_root_.max]} }, 168 | all_goals 169 | {dsimp [-modal_degree], rw cons_degree, 170 | rw lt_max_iff, right, 171 | apply unbox_degree_aux, swap, exact φ, 172 | cases h, { rw heq at h, contradiction}, {exact h} } 173 | end 174 | 175 | theorem unbox_degree : Π {Γ φ}, dia φ ∈ Γ → 176 | modal_degree (φ :: unbox Γ) < modal_degree Γ 177 | | [] φ h := absurd h $ not_mem_nil _ 178 | | (hd::tl) φ h := 179 | begin 180 | rw cons_degree, apply max_lt, 181 | {apply lt_of_lt_of_le, swap 3, exact count_modal (dia φ), simp, apply le_of_mem_degree, exact h}, 182 | {apply unbox_degree_aux, exact h} 183 | end 184 | 185 | @[simp] def rebox : list nnf → list nnf 186 | | [] := [] 187 | | (hd::tl) := box hd :: rebox tl 188 | 189 | theorem rebox_unbox_of_mem : Π {Γ} (h : ∀ {φ}, φ ∈ unbox Γ → box φ ∈ Γ), rebox (unbox Γ) ⊆ Γ 190 | | [] h := by simp 191 | | (hd::tl) h := 192 | begin 193 | cases hψ : hd, 194 | case nnf.box : φ {dsimp, simp [cons_subset_cons], apply subset_cons_of_subset, apply rebox_unbox_of_mem, simp [unbox_iff]}, 195 | all_goals {dsimp, apply subset_cons_of_subset, apply rebox_unbox_of_mem, simp [unbox_iff]} 196 | end 197 | 198 | theorem unbox_rebox : Π {Γ}, unbox (rebox Γ) = Γ 199 | | [] := by simp 200 | | (hd::tl) := by simp [unbox_rebox] 201 | 202 | -- Just that I don't want to say ∃ φ s.t. ... 203 | def box_only_rebox : Π {Γ}, box_only (rebox Γ) 204 | | [] := {no_var := by simp, 205 | no_neg := by simp, 206 | no_and := by simp, 207 | no_or := by simp, 208 | no_dia := by simp} 209 | | (hd::tl) := 210 | begin 211 | cases h : hd, 212 | all_goals { 213 | exact { no_var := begin 214 | intros n h, cases h, contradiction, 215 | apply (@box_only_rebox tl).no_var, assumption 216 | end, 217 | no_neg := begin 218 | intros n h, cases h, contradiction, 219 | apply (@box_only_rebox tl).no_neg, assumption 220 | end, 221 | no_and := begin 222 | intros φ ψ h, cases h, contradiction, 223 | apply (@box_only_rebox tl).no_and, assumption 224 | end, 225 | no_or := begin 226 | intros φ ψ h, cases h, contradiction, 227 | apply (@box_only_rebox tl).no_or, assumption 228 | end, 229 | no_dia := begin 230 | intros φ h, cases h, contradiction, 231 | apply (@box_only_rebox tl).no_dia, assumption 232 | end} } 233 | end 234 | 235 | theorem rebox_iff : Π {φ Γ}, box φ ∈ rebox Γ ↔ φ ∈ Γ 236 | | φ [] := by simp 237 | | φ (hd::tl) := 238 | begin 239 | split, 240 | {intro h, cases h₁ : hd, 241 | all_goals { cases h, 242 | {left, rw ←h₁, injection h}, 243 | {right, have := (@rebox_iff φ tl).1, exact this h } }}, 244 | {intro h, cases h₁ : hd, 245 | all_goals { dsimp, cases h, 246 | {left, rw ←h₁, rw h}, 247 | {right, have := (@rebox_iff φ tl).2, exact this h } } } 248 | end 249 | 250 | @[simp] def undia : list nnf → list nnf 251 | | [] := [] 252 | | ((dia φ) :: l) := φ :: undia l 253 | | (e :: l) := undia l 254 | 255 | theorem undia_iff : Π {Γ φ}, dia φ ∈ Γ ↔ φ ∈ undia Γ 256 | | [] φ := begin split, repeat {intro h, simpa using h} end 257 | | (hd::tl) φ := 258 | begin 259 | split, 260 | { intro h, cases h₁ : hd, 261 | case nnf.dia : ψ 262 | { dsimp [undia], cases h, 263 | {left, rw h₁ at h, injection h}, 264 | {right, exact (@undia_iff tl φ).1 h} }, 265 | all_goals 266 | { dsimp [undia], cases h, 267 | {rw h₁ at h, contradiction}, 268 | {exact (@undia_iff tl φ).1 h} } }, 269 | { intro h, cases h₁ : hd, 270 | case nnf.dia : ψ 271 | { rw h₁ at h, dsimp [undia] at h, cases h, 272 | {simp [h]}, {right, exact (@undia_iff tl φ).2 h} }, 273 | all_goals 274 | { rw h₁ at h, dsimp [undia] at h, right, exact (@undia_iff tl φ).2 h } } 275 | end 276 | 277 | theorem undia_degree : Π {Γ φ}, φ ∈ undia Γ → 278 | modal_degree (φ :: unbox Γ) < modal_degree Γ 279 | := 280 | begin 281 | intros Γ φ h, 282 | rw ←undia_iff at h, 283 | apply unbox_degree h 284 | end 285 | 286 | def get_contra : Π Γ : list nnf, 287 | psum {p : nat // var p ∈ Γ ∧ neg p ∈ Γ} 288 | (∀ n, var n ∈ Γ → neg n ∉ Γ) 289 | | [] := psum.inr $ λ _ h, absurd h $ not_mem_nil _ 290 | | (hd :: tl) := 291 | begin 292 | cases h : hd, 293 | case nnf.var : n 294 | {apply dite (neg n ∈ tl), 295 | {intro t, 296 | exact psum.inl ⟨n, ⟨mem_cons_self _ _, mem_cons_of_mem _ t⟩⟩}, 297 | {intro e, 298 | cases (get_contra tl), 299 | {left, constructor, constructor, 300 | apply mem_cons_of_mem, exact val.2.1, 301 | apply mem_cons_of_mem, exact val.2.2}, 302 | {right, 303 | intros m hm hin, 304 | by_cases eq : m=n, 305 | {apply e, cases hin, contradiction, rw ←eq, assumption}, 306 | {cases hm, apply eq, injection hm, apply val, exact hm, 307 | cases hin, contradiction, assumption} } } 308 | }, 309 | case nnf.neg : n 310 | { apply dite (var n ∈ tl), 311 | { intro t, 312 | exact psum.inl ⟨n, ⟨mem_cons_of_mem _ t, mem_cons_self _ _⟩⟩ }, 313 | { intro e, 314 | cases (get_contra tl), 315 | {left, constructor, constructor, 316 | apply mem_cons_of_mem, exact val.2.1, 317 | apply mem_cons_of_mem, exact val.2.2 }, 318 | { right, 319 | intros m hm hin, 320 | by_cases eq : m=n, 321 | { apply e, cases hm, contradiction, rw ←eq, assumption }, 322 | { cases hin, apply eq, injection hin, apply val, 323 | swap, exact hin, cases hm, contradiction, assumption } 324 | } 325 | } 326 | }, 327 | all_goals 328 | { 329 | cases (get_contra tl), 330 | { left, constructor, constructor, 331 | apply mem_cons_of_mem, exact val.2.1, 332 | apply mem_cons_of_mem, exact val.2.2 }, 333 | { right, 334 | intros m hm hin, 335 | {apply val, swap 3, exact m, 336 | cases hm, contradiction, assumption, 337 | cases hin, contradiction, assumption} } 338 | } 339 | end 340 | 341 | def get_contra_seqt : Π Γ : seqt, 342 | psum {p : nat // var p ∈ Γ.main ∧ neg p ∈ Γ.main} 343 | (∀ n, var n ∈ Γ.main → neg n ∉ Γ.main) 344 | := λ Γ, get_contra Γ.main 345 | 346 | def get_and : Π Γ : list nnf, 347 | psum {p : nnf × nnf // and p.1 p.2 ∈ Γ} 348 | (∀ φ ψ, nnf.and φ ψ ∉ Γ) 349 | | [] := psum.inr $ λ _ _, not_mem_nil _ 350 | | (hd :: tl) := 351 | begin 352 | cases h : hd, 353 | case nnf.and : φ ψ { left, constructor,swap, 354 | constructor, exact φ, exact ψ, simp 355 | }, 356 | all_goals 357 | { cases (get_and tl), 358 | {left, 359 | constructor, 360 | apply mem_cons_of_mem, 361 | exact val.2}, 362 | {right, intros γ ψ h, 363 | cases h, contradiction, 364 | apply val, assumption } 365 | } 366 | end 367 | 368 | def get_and_seqt : Π Γ : seqt, 369 | psum {p : nnf × nnf // and p.1 p.2 ∈ Γ.main} 370 | (∀ φ ψ, nnf.and φ ψ ∉ Γ.main) 371 | := λ Γ, get_and Γ.main 372 | 373 | def get_or : Π Γ : list nnf, 374 | psum {p : nnf × nnf // or p.1 p.2 ∈ Γ} 375 | (∀ φ ψ, nnf.or φ ψ ∉ Γ) 376 | | [] := psum.inr $ λ _ _, not_mem_nil _ 377 | | (hd :: tl) := 378 | begin 379 | cases h : hd, 380 | case nnf.or : φ ψ { left, constructor,swap, 381 | constructor, exact φ, exact ψ, simp }, 382 | all_goals 383 | { cases (get_or tl), 384 | {left, 385 | constructor, 386 | apply mem_cons_of_mem, 387 | exact val.2}, 388 | {right, intros γ ψ h, 389 | cases h, contradiction, 390 | apply val, assumption} 391 | } 392 | end 393 | 394 | def get_or_seqt : Π Γ : seqt, 395 | psum {p : nnf × nnf // or p.1 p.2 ∈ Γ.main} 396 | (∀ φ ψ, nnf.or φ ψ ∉ Γ.main) 397 | := λ Γ, get_or Γ.main 398 | 399 | def get_dia : Π Γ : list nnf, 400 | psum {p : nnf // dia p ∈ Γ} 401 | (∀ φ, nnf.dia φ ∉ Γ) 402 | | [] := psum.inr $ λ _, not_mem_nil _ 403 | | (hd :: tl) := 404 | begin 405 | cases h : hd, 406 | case nnf.dia : φ { left, constructor, swap, exact φ, simp }, 407 | all_goals 408 | { cases (get_dia tl), 409 | {left, 410 | constructor, 411 | apply mem_cons_of_mem, 412 | exact val.2}, 413 | {right, intros γ h, 414 | cases h, contradiction, 415 | apply val, assumption } } 416 | end 417 | 418 | def get_dia_seqt : Π Γ : seqt, 419 | psum {p : nnf // dia p ∈ Γ.main} 420 | (∀ φ, nnf.dia φ ∉ Γ.main) 421 | := λ Γ, get_dia Γ.main 422 | 423 | def get_box : Π Γ : list nnf, 424 | psum {l : nnf // box l ∈ Γ} 425 | (∀ φ, nnf.box φ ∉ Γ) 426 | | [] := psum.inr $ λ _, not_mem_nil _ 427 | | (hd :: tl) := 428 | begin 429 | cases h : hd, 430 | case nnf.box : φ { left, constructor, swap, exact φ, simp }, 431 | all_goals 432 | { cases (get_box tl), 433 | {left, 434 | constructor, 435 | apply mem_cons_of_mem, 436 | exact val.2}, 437 | {right, intros γ h, 438 | cases h, contradiction, 439 | apply val, assumption } } 440 | end 441 | 442 | def get_box_seqt : Π Γ : seqt, 443 | psum {p : nnf // box p ∈ Γ.main} 444 | (∀ φ, nnf.box φ ∉ Γ.main) 445 | := λ Γ, get_box Γ.main 446 | 447 | @[simp] def get_var : list nnf → list ℕ 448 | | [] := [] 449 | | ((var n) :: l) := n :: get_var l 450 | | (e :: l) := get_var l 451 | 452 | theorem get_var_iff : Π {Γ n}, var n ∈ Γ ↔ n ∈ get_var Γ 453 | | [] φ := begin split, repeat {intro h, simpa using h} end 454 | | (hd::tl) φ := 455 | begin 456 | split, 457 | { intro h, cases h₁ : hd, 458 | case nnf.var : n 459 | { dsimp, cases h, 460 | {left, rw h₁ at h, injection h}, 461 | {right, exact (@get_var_iff tl φ).1 h} }, 462 | all_goals 463 | { dsimp, cases h, 464 | {rw h₁ at h, contradiction}, 465 | {exact (@get_var_iff tl φ).1 h} } }, 466 | { intro h, cases h₁ : hd, 467 | case nnf.var : n 468 | { rw h₁ at h, dsimp at h, cases h, 469 | {simp [h]}, {right, exact (@get_var_iff tl φ).2 h} }, 470 | all_goals 471 | { rw h₁ at h, dsimp [undia] at h, right, exact (@get_var_iff tl φ).2 h } } 472 | end 473 | -------------------------------------------------------------------------------- /src/KT/rules.lean: -------------------------------------------------------------------------------- 1 | import .semantics 2 | open nnf subtype list 3 | 4 | -- TODO: we could redefine satisfiability. 5 | inductive node (Γ : seqt) : Type 6 | | closed : unsatisfiable (Γ.main ++ Γ.hdld) → node 7 | | open_ : {s // sat builder s (Γ.main ++ Γ.hdld)} → node 8 | 9 | open node 10 | 11 | def contra_rule_seqt {Δ : seqt} {n} (h : var n ∈ Δ.main ∧ neg n ∈ Δ.main) : node Δ := 12 | begin 13 | left, 14 | {exact unsat_contra_seqt h.1 h.2} 15 | end 16 | 17 | def and_rule_seqt {Γ Δ : seqt} (i : and_instance_seqt Γ Δ) : node Δ → node Γ 18 | | (closed h) := begin 19 | left, 20 | {apply unsat_of_closed_and_seqt, repeat {assumption}}, 21 | end 22 | 23 | | (open_ w) := begin 24 | right, 25 | clear and_rule_seqt, 26 | cases i with φ ψ h, 27 | constructor, 28 | exact sat_and_of_sat_split_seqt _ _ _ _ _ h w.2 29 | end 30 | 31 | def or_rule_seqt {Γ₁ Γ₂ Δ} (i : or_instance_seqt Δ Γ₁ Γ₂) : 32 | node Γ₁ → node Γ₂ → node Δ 33 | | (open_ w) _ := begin 34 | right, 35 | clear or_rule_seqt, 36 | cases i, 37 | constructor, 38 | apply sat_or_of_sat_split_left, 39 | apply subset_append_left, 40 | assumption, 41 | apply sat_subset, swap, exact w.2, 42 | rw erase_append_left _ i_h, 43 | rw ←cons_append, simp [subset_cons], 44 | end 45 | | _ (open_ w) := begin 46 | right, 47 | clear or_rule_seqt, 48 | cases i, 49 | constructor, 50 | apply sat_or_of_sat_split_right, 51 | apply subset_append_left, 52 | assumption, 53 | apply sat_subset, swap, exact w.2, 54 | rw erase_append_left _ i_h, 55 | rw ←cons_append, simp [subset_cons] 56 | end 57 | | (closed h₁) (closed h₂):= begin 58 | left, 59 | clear or_rule_seqt, 60 | cases i, 61 | apply unsat_or_of_unsat_split_seqt, 62 | repeat {assumption} 63 | end 64 | 65 | def open_rule_seqt {Γ₁ Γ₂ Δ} {s:model} (i : or_instance_seqt Δ Γ₁ Γ₂) (h : sat builder s (Γ₁.main++Γ₁.hdld)) : node Δ := 66 | begin 67 | right, 68 | cases i, 69 | constructor, 70 | apply sat_or_of_sat_split_left, 71 | apply subset_append_left, 72 | assumption, 73 | apply sat_subset, swap, exact h, 74 | rw erase_append_left _ i_h, 75 | rw ←cons_append, simp [subset_cons] 76 | end 77 | 78 | def copy_rule_seqt {Γ Δ : seqt} (i : copy_instance_seqt Γ Δ) : node Δ → node Γ 79 | | (closed h) := begin 80 | left, 81 | clear copy_rule_seqt, 82 | cases i with φ hmem, 83 | apply unsat_copy_of_unsat_box, 84 | repeat {assumption} 85 | end 86 | 87 | | (open_ w) := begin 88 | right, 89 | clear copy_rule_seqt, 90 | cases i with φ h, 91 | constructor, 92 | apply sat_copy_of_sat_box, exact h, exact w.2 93 | end 94 | 95 | open batch_sat_seqt 96 | 97 | def dia_rule_seqt {p : seqt → Prop} (f : Π Γ, p Γ → node Γ): Π Γ : list seqt, (∀ i∈Γ, p i) → 98 | psum ({i // i ∈ Γ ∧ unsatisfiable (seqt.main i ++ seqt.hdld i)}) {x : list model // batch_sat_seqt x Γ} 99 | | [] h := psum.inr ⟨[], bs_nil⟩ 100 | | (hd :: tl) h := 101 | match f hd (h hd (by simp)) with 102 | | (node.closed pr) := psum.inl ⟨hd, by simp, pr⟩ 103 | | (node.open_ w₁) := 104 | match dia_rule_seqt tl (λ x hx, h x $ by simp [hx]) with 105 | | (psum.inl uw) := begin 106 | left, rcases uw with ⟨w, hin, h⟩, 107 | split, split, swap, exact h, simp [hin] 108 | end 109 | | (psum.inr w₂) := psum.inr ⟨(w₁.1::w₂), bs_cons _ _ _ _ w₁.2 w₂.2⟩ 110 | end 111 | end 112 | -------------------------------------------------------------------------------- /src/KT/semantics.lean: -------------------------------------------------------------------------------- 1 | import .ops 2 | open subtype nnf model 3 | 4 | def unmodal_seqt (Γ : seqt) : list seqt := 5 | list.map 6 | (λ d, 7 | {main := d :: (unbox (Γ.main ++ Γ.hdld)), hdld := [], 8 | pmain := begin intros v hv φ hsat hin hall, exfalso, apply list.not_mem_nil, exact hin end, 9 | phdld := box_only_nil}) 10 | (undia (Γ.main ++ Γ.hdld)) 11 | 12 | def unmodal_nil_hdld (Γ : seqt) : ∀ (i : seqt), i ∈ unmodal_seqt Γ → (i.hdld = []) := 13 | list.mapp _ _ 14 | begin intros φ h, constructor end 15 | 16 | def unmodal_seqt_size (Γ : seqt) : ∀ (i : seqt), i ∈ unmodal_seqt Γ → (prod.measure_lex seqt_size i Γ) := 17 | list.mapp _ _ 18 | begin 19 | intros φ h, 20 | constructor, simp [-list.map_append], 21 | apply undia_degree h 22 | end 23 | 24 | def unmodal_mem_box (Γ : seqt) : ∀ (i : seqt), i ∈ unmodal_seqt Γ → (∀ φ, box φ ∈ Γ.main ++ Γ.hdld → φ ∈ i.main) := 25 | list.mapp _ _ begin intros φ h ψ hψ, right, apply (@unbox_iff (Γ.main++Γ.hdld) ψ).1 hψ end 26 | 27 | def unmodal_sat_of_sat (Γ : seqt) : ∀ (i : seqt), i ∈ unmodal_seqt Γ → 28 | (∀ {st : Type} (k : KT st) s Δ 29 | (h₁ : ∀ φ, box φ ∈ Γ.main++Γ.hdld → box φ ∈ Δ) 30 | (h₂ : ∀ φ, dia φ ∈ Γ.main++Γ.hdld → dia φ ∈ Δ), 31 | sat k s Δ → ∃ s', sat k s' i.main) := 32 | list.mapp _ _ 33 | begin 34 | intros φ hmem st k s Δ h₁ h₂ h, 35 | have : force k s (dia φ), 36 | { apply h, apply h₂, rw undia_iff, exact hmem }, 37 | rcases this with ⟨w, hrel, hforce⟩, 38 | split, swap, {exact w}, 39 | { intro ψ, intro hψ, cases hψ, 40 | {rw hψ, exact hforce}, 41 | {have := h₁ _ ((@unbox_iff (Γ.main++Γ.hdld) ψ).2 hψ), 42 | have := h _ this, 43 | apply this _ hrel} }, 44 | end 45 | 46 | -- TODO: Rewrite this ugly one. 47 | def mem_unmodal_seqt (Γ : seqt) (φ) (h : φ ∈ undia (Γ.main++Γ.hdld)) : 48 | (⟨φ :: unbox (Γ.main++Γ.hdld), [], begin intros v hv φ hsat hin hall, exfalso, apply list.not_mem_nil, exact hin end, box_only_nil⟩ : seqt) ∈ unmodal_seqt Γ := 49 | begin 50 | dsimp [unmodal_seqt], 51 | apply list.mem_map_of_mem (λ φ, (⟨φ :: unbox (Γ.main++Γ.hdld), _,_,_⟩:seqt)) h 52 | end 53 | 54 | def unsat_of_unsat_unmodal {Γ : seqt} (h : modal_applicable Γ) (i : seqt) : i ∈ unmodal_seqt Γ ∧ unsatisfiable (i.main ++ i.hdld) → unsatisfiable (Γ.main ++ Γ.hdld) := 55 | begin 56 | intro hex, intros st k s h, 57 | have := unmodal_sat_of_sat Γ i hex.1 k s (Γ.main ++ Γ.hdld) (λ x hx, hx) (λ x hx, hx) h, 58 | cases this with w hw, 59 | have := hex.2, 60 | rw (unmodal_nil_hdld _ _ hex.1) at this, 61 | simp at this, 62 | exact this _ _ _ hw 63 | end 64 | 65 | /- Part of the soundness -/ 66 | 67 | theorem unsat_of_closed_and {Γ Δ} (i : and_instance Γ Δ) (h : unsatisfiable Δ) : unsatisfiable Γ := 68 | by cases i; { apply unsat_and_of_unsat_split, repeat {assumption} } 69 | 70 | theorem unsat_of_closed_and_seqt {Γ Δ} (i : and_instance_seqt Γ Δ) (h : unsatisfiable (Δ.main++Δ.hdld)) : unsatisfiable (Γ.main++Γ.hdld) := 71 | begin 72 | cases i, {apply unsat_and_of_unsat_split_seqt, repeat {assumption} } 73 | end 74 | 75 | theorem unsat_of_closed_or_seqt {Γ₁ Γ₂ Δ : seqt} (i : or_instance_seqt Δ Γ₁ Γ₂) 76 | (h₁ : unsatisfiable (Γ₁.main++Γ₁.hdld)) 77 | (h₂ : unsatisfiable (Γ₂.main++Γ₂.hdld)) : 78 | unsatisfiable (Δ.main++Δ.hdld) := 79 | begin 80 | cases i, {apply unsat_or_of_unsat_split_seqt, repeat {assumption}} 81 | end 82 | 83 | /- Tree models -/ 84 | 85 | inductive batch_sat_seqt : list model → list seqt → Prop 86 | | bs_nil : batch_sat_seqt [] [] 87 | | bs_cons (m) (Γ:seqt) (l₁ l₂) : sat builder m (Γ.main++Γ.hdld) → 88 | batch_sat_seqt l₁ l₂ → 89 | batch_sat_seqt (m::l₁) (Γ::l₂) 90 | 91 | open batch_sat_seqt 92 | 93 | theorem bs_ex_seqt : Π l Γ, 94 | batch_sat_seqt l Γ → ∀ m ∈ l, ∃ i ∈ Γ, sat builder m (seqt.main i ++ i.hdld) 95 | | l Γ bs_nil := λ m hm, by simpa using hm 96 | | l Γ (bs_cons m Δ l₁ l₂ h hbs) := 97 | begin 98 | intros n hn, 99 | cases hn, 100 | {split, swap, exact Δ, split, simp, rw hn, exact h}, 101 | {have : ∃ (i : seqt) (H : i ∈ l₂), sat builder n (i.main ++ i.hdld), 102 | {apply bs_ex_seqt, exact hbs, exact hn}, 103 | {rcases this with ⟨w, hw, hsat⟩, split, swap, exact w, split, 104 | {simp [hw]}, {exact hsat} } } 105 | end 106 | 107 | theorem bs_forall_seqt : Π l Γ, 108 | batch_sat_seqt l Γ → ∀ i ∈ Γ, ∃ m ∈ l, sat builder m (seqt.main i ++ i.hdld) 109 | | l Γ bs_nil := λ m hm, by simpa using hm 110 | | l Γ (bs_cons m Δ l₁ l₂ h hbs) := 111 | begin 112 | intros i hi, 113 | cases hi, {split, swap, exact m, split, simp, rw hi, exact h}, 114 | {have : ∃ (n : model) (H : n ∈ l₁), sat builder n (seqt.main i ++ i.hdld), 115 | {apply bs_forall_seqt, exact hbs, exact hi}, 116 | {rcases this with ⟨w, hw, hsat⟩, split, swap, exact w, split, 117 | {simp [hw]}, {exact hsat} } } 118 | end 119 | 120 | theorem sat_of_batch_sat_main : Π l Γ (h : modal_applicable Γ), 121 | batch_sat_seqt l (unmodal_seqt Γ) → 122 | sat builder (cons h.v l) (Γ.main) := 123 | begin 124 | intros l Γ h hbs, 125 | intros φ hφ, 126 | cases hfml : φ, 127 | case nnf.var : n 128 | {rw hfml at hφ, simp, rw ←h.hv, 129 | exact hφ }, 130 | case nnf.box : ψ 131 | {exfalso, rw hfml at hφ, apply h.no_box_main, exact hφ}, 132 | case nnf.dia : ψ 133 | {dsimp, 134 | let se : seqt := 135 | {main := ψ :: (unbox (Γ.main ++ Γ.hdld)), hdld := [], 136 | pmain := begin intros v hv φ hsat hin hall, exfalso, apply list.not_mem_nil, exact hin end, 137 | phdld := box_only_nil}, 138 | have := bs_forall_seqt l (unmodal_seqt Γ) hbs se _, swap, 139 | {dsimp [unmodal_seqt], apply list.mem_map_of_mem (λ φ, (⟨φ :: unbox (Γ.main++Γ.hdld),_,_,_⟩:seqt)), rw ←undia_iff, rw ←hfml, simp [hφ] }, 140 | {rcases this with ⟨w, hw, hsat⟩, split, swap, exact w, split, 141 | {simp [hw]}, {apply hsat, simp} } }, 142 | case nnf.neg : n 143 | {dsimp, rw hfml at hφ, have : var n ∉ Γ.main, 144 | {intro hin, have := h.no_contra_main, have := this hin, contradiction}, 145 | simp, rw ←h.hv, exact this }, 146 | case nnf.and : φ ψ 147 | {rw hfml at hφ, have := h.satu.no_and, have := @this φ ψ, contradiction}, 148 | case nnf.or : φ ψ 149 | {rw hfml at hφ, have := h.satu.no_or, have := @this φ ψ, contradiction} 150 | end 151 | 152 | theorem sat_of_batch_sat_box : Π l Γ (h : modal_applicable Γ), 153 | batch_sat_seqt l (unmodal_seqt Γ) → ∀ m∈l, ∀ φ, box φ ∈ Γ.hdld → force builder m φ := 154 | begin 155 | intros l Γ h hbs m hm φ hφ, 156 | have := bs_ex_seqt _ _ hbs _ hm, 157 | rcases this with ⟨i,hin,hi⟩, 158 | have := unmodal_mem_box _ _ hin φ, 159 | apply hi, 160 | rw list.mem_append, left, 161 | apply this, simp [hφ] 162 | end 163 | 164 | theorem sat_of_batch_sat : Π l Γ (h : modal_applicable Γ), 165 | batch_sat_seqt l (unmodal_seqt Γ) → 166 | sat builder (cons h.v l) (Γ.main++Γ.hdld) := 167 | begin 168 | intros l Γ h hbs, 169 | apply sat_append, 170 | {apply sat_of_batch_sat_main, assumption}, 171 | {intros φ hφ, have := box_only_ex Γ.phdld hφ, 172 | cases this with w hw, 173 | rw hw, dsimp, intros s' hs', 174 | have := mem_of_mrel_tt hs', cases this, 175 | {apply sat_of_batch_sat_box, exact h, exact hbs, exact this, rw hw at hφ, exact hφ}, 176 | { rw ←this, apply Γ.pmain, 177 | apply sat_of_batch_sat_main, exact hbs, rw hw at hφ, exact hφ, 178 | { intros, apply sat_of_batch_sat_box, repeat {assumption} } } 179 | } 180 | end 181 | -------------------------------------------------------------------------------- /src/KT/size.lean: -------------------------------------------------------------------------------- 1 | import .KT_defs algebra.order_functions 2 | open nnf prod 3 | 4 | namespace prod 5 | 6 | universe u 7 | 8 | def measure_lex {α : Sort u} : (α → ℕ × ℕ) → α → α → Prop := 9 | inv_image $ lex (<) (<) 10 | 11 | def measure_lex_wf {α : Sort u} (f : α → ℕ × ℕ) : well_founded (measure_lex f) := 12 | inv_image.wf f (lex_wf nat.lt_wf nat.lt_wf) 13 | 14 | end prod 15 | 16 | namespace list 17 | universes u 18 | variables {α : Type u} [inhabited α] [decidable_linear_order α] 19 | 20 | def maximum (l : list α) : α := l.foldr max l.head 21 | 22 | theorem le_of_foldr_max : Π {a b : α} {l}, a ∈ l → a ≤ foldr max b l 23 | | a b [] h := absurd h $ not_mem_nil _ 24 | | a b (hd::tl) h := 25 | begin 26 | cases h, 27 | { simp [h, le_max_left] }, 28 | { simp [le_max_right_of_le, le_of_foldr_max h] } 29 | end 30 | 31 | theorem mem_of_foldr_max : Π {a : α} {l}, foldr max a l ∈ a :: l 32 | | a [] := by simp 33 | | a (hd::tl) := 34 | begin 35 | simp only [foldr_cons], 36 | have h := by exact @max_choice _ _ hd (foldr max a tl), 37 | cases h, 38 | { simp [h] }, 39 | { rw h, 40 | have hmem := @mem_of_foldr_max a tl, 41 | cases hmem, { simp [hmem] }, { right, right, exact hmem } } 42 | end 43 | 44 | theorem mem_maximum : Π {l : list α}, l ≠ [] → maximum l ∈ l 45 | | [] h := by contradiction 46 | | (hd::tl) h := 47 | begin 48 | dsimp [maximum], 49 | have hc := by exact @max_choice _ _ hd (foldr max hd tl), 50 | cases hc, { simp [hc] }, { simp [hc, mem_of_foldr_max] } 51 | end 52 | 53 | theorem le_maximum_of_mem : Π {a : α} {l}, a ∈ l → a ≤ maximum l 54 | | a [] h := absurd h $ not_mem_nil _ 55 | | a (hd::tl) h := 56 | begin 57 | cases h, 58 | { rw h, apply le_of_foldr_max, simp }, 59 | { dsimp [maximum], apply le_max_right_of_le, apply le_of_foldr_max h } 60 | end 61 | 62 | def maximum_cons : Π {a : α} {l}, l ≠ [] → maximum (a :: l) = max a (maximum l) 63 | | a [] h := by contradiction 64 | | a (hd::tl) h := 65 | begin 66 | apply le_antisymm, 67 | { have : a :: hd :: tl ≠ [], { simp [h] }, 68 | have hle := mem_maximum this, 69 | cases hle, 70 | { simp [hle, le_max_left] }, 71 | { apply le_max_right_of_le, apply le_maximum_of_mem, exact hle } }, 72 | { have hc := by exact @max_choice _ _ a (maximum $ hd :: tl), 73 | cases hc, 74 | { simp [hc, le_maximum_of_mem] }, 75 | { simp [hc, le_maximum_of_mem, mem_maximum h] } } 76 | end 77 | 78 | @[simp] def maximum_singleton : Π {a : α}, maximum [a] = a := by simp [maximum] 79 | 80 | section 81 | variables hdf : ∀ {a}, default α ≤ a 82 | include hdf 83 | 84 | def maximum_nil : Π {a : α}, max a (maximum []) = a := 85 | λ a, by dsimp [maximum]; apply max_eq_left; apply hdf 86 | 87 | def maximum_cons_default : Π {a : α} {l}, maximum (a :: l) = max a (maximum l) 88 | | a [] := begin rw maximum_nil, simp, assumption end 89 | | a (hd::tl) := 90 | begin 91 | apply le_antisymm, 92 | { have : a :: hd :: tl ≠ [], { simp }, 93 | have hle := mem_maximum this, 94 | cases hle, 95 | { rw hle, apply le_max_left }, 96 | { apply le_max_right_of_le, apply le_maximum_of_mem, exact hle } }, 97 | { have hc := by exact @max_choice _ _ a (maximum $ hd :: tl), 98 | cases hc, 99 | { rw hc, apply le_maximum_of_mem, simp }, 100 | { rw hc, apply le_maximum_of_mem, right, 101 | have : hd :: tl ≠ [], { simp }, 102 | exact mem_maximum this } } 103 | end 104 | 105 | def maximum_erase : Π {a : α} {l}, a ∈ l → maximum l = max a (maximum $ l.erase a) 106 | | a [] h := absurd h $ not_mem_nil _ 107 | | a (hd::tl) h := 108 | begin 109 | by_cases heq : a = hd, 110 | { rw heq, simp, apply maximum_cons_default, assumption }, 111 | { rw maximum_cons_default @hdf, 112 | rw maximum_erase (mem_of_ne_of_mem heq h), 113 | rw erase_cons_tail, 114 | rw maximum_cons_default @hdf, 115 | repeat {rw ←max_assoc}, 116 | simp [max_comm], 117 | intro hc, rw hc at heq, contradiction } 118 | end 119 | 120 | def maximum_subset : Π {l₁ l₂ : list α}, l₁ ⊆ l₂ → maximum l₁ ≤ maximum l₂ 121 | | [] l₂ h := by apply hdf 122 | | (hd::tl) l₂ h := 123 | begin 124 | rw maximum_cons_default, apply max_le, 125 | { apply le_maximum_of_mem, apply h, simp }, 126 | { apply maximum_subset, rw cons_subset at h, exact h.right }, 127 | assumption 128 | end 129 | 130 | def maximum_sublist {l₁ l₂ : list α} (h : l₁ <+ l₂) : maximum l₁ ≤ maximum l₂ := 131 | maximum_subset (by assumption) (sublist.subset h) 132 | 133 | end 134 | 135 | end list 136 | 137 | open list 138 | 139 | @[simp] def count_modal : nnf → ℕ 140 | | (var n) := 0 141 | | (neg n) := 0 142 | | (and φ ψ) := count_modal φ + count_modal ψ 143 | | (or φ ψ) := count_modal φ + count_modal ψ 144 | | (box φ) := count_modal φ + 1 145 | | (dia φ) := count_modal φ + 1 146 | 147 | def modal_degree (Γ : list nnf) : ℕ := maximum $ Γ.map count_modal 148 | 149 | theorem cons_degree : Π (φ : nnf) (Γ : list nnf), 150 | modal_degree (φ :: Γ) = max (count_modal φ) (modal_degree Γ) 151 | | φ [] := begin dsimp [modal_degree, maximum], simp, rw max_eq_left, apply zero_le end 152 | | φ (hd::tl) := begin simp only [modal_degree], apply maximum_cons, simp end 153 | 154 | theorem le_of_mem_degree : Π (φ : nnf) (Γ : list nnf), φ ∈ Γ → count_modal φ ≤ modal_degree Γ 155 | | φ [] h := absurd h $ list.not_mem_nil _ 156 | | φ (hd::tl) h := 157 | begin 158 | cases h, 159 | {rw h, rw cons_degree, apply le_max_left}, 160 | {rw cons_degree, apply le_max_right_of_le, apply le_of_mem_degree, exact h} 161 | end 162 | 163 | theorem le_degree_of_subset : Π {l₁ l₂ : list nnf}, l₁ ⊆ l₂ → 164 | modal_degree l₁ ≤ modal_degree l₂ 165 | := 166 | begin 167 | intros, 168 | dsimp [modal_degree], 169 | apply maximum_subset, 170 | {exact zero_le}, 171 | {apply map_subset _ a} 172 | end 173 | 174 | @[simp] def seqt_size (s : seqt) : ℕ × ℕ := (modal_degree $ s.main ++ s.hdld, node_size s.main) 175 | 176 | instance : has_well_founded seqt := 177 | ⟨ prod.measure_lex seqt_size, prod.measure_lex_wf seqt_size ⟩ 178 | 179 | lemma size_erase_add_eq_size_sum {φ} : Π (Γ : list nnf) (h : φ ∈ Γ), node_size (Γ.erase φ) + sizeof φ = node_size Γ 180 | | [] h := absurd h $ not_mem_nil _ 181 | | (hd :: tl) h := 182 | begin 183 | by_cases eq : φ = hd, 184 | {dsimp [list.erase], rw if_pos, rw eq, apply add_comm, rw eq}, 185 | {dsimp [list.erase], rw if_neg, dsimp, rw [add_assoc], have : node_size (list.erase tl φ) + sizeof φ = node_size tl, 186 | {apply size_erase_add_eq_size_sum, cases h, contradiction, assumption}, 187 | rw this, intro h', rw h' at eq, contradiction} 188 | end 189 | 190 | theorem split_lt_and_size {φ ψ} (Γ : list nnf) (h : nnf.and φ ψ ∈ Γ) : 191 | node_size (φ :: ψ :: Γ.erase (nnf.and φ ψ)) < node_size Γ := 192 | begin 193 | dsimp [node_size], 194 | rw ←size_erase_add_eq_size_sum Γ, 195 | swap, exact h, 196 | rw [←add_assoc, add_comm], 197 | apply add_lt_add_left, dsimp [sizeof, has_sizeof.sizeof, nnf.sizeof], 198 | rw [add_assoc], 199 | rw [nat.one_add], 200 | apply nat.lt_succ_self 201 | end 202 | 203 | theorem split_le_and_degree {φ ψ} (Γ : list nnf) (h : nnf.and φ ψ ∈ Γ) : 204 | modal_degree (φ :: ψ :: Γ.erase (nnf.and φ ψ)) ≤ modal_degree Γ := 205 | begin 206 | rw cons_degree, rw cons_degree, 207 | {rw ←max_assoc, apply max_le, 208 | {apply max_le, 209 | repeat {apply le_trans, swap 3, exact count_modal (and φ ψ), simp, apply le_of_mem_degree _ _ h}}, 210 | {apply le_degree_of_subset, apply erase_subset} } 211 | end 212 | 213 | theorem split_lt_and_seqt {φ ψ} (Γ : seqt) (h : nnf.and φ ψ ∈ Γ.main) : 214 | prod.measure_lex seqt_size 215 | (and_child Γ h) Γ := 216 | begin 217 | have hmem : and φ ψ ∈ Γ.main ++ Γ.hdld, {apply mem_append_left _ h}, 218 | have := lt_or_eq_of_le (split_le_and_degree _ hmem), 219 | cases this, 220 | {left, 221 | simp only [seqt_size, and_child], rw erase_append_left at this, rw cons_append, rw cons_append, exact this, exact h 222 | }, 223 | {dsimp only [prod.measure_lex, inv_image, seqt_size], rw [erase_append_left _ h] at this, rw ←this, right, 224 | apply split_lt_and_size _ h} 225 | end 226 | 227 | theorem split_lt_or_size_left {φ ψ} (Γ : list nnf) (h : nnf.or φ ψ ∈ Γ) : 228 | node_size (φ :: Γ.erase (nnf.or φ ψ)) < node_size Γ := 229 | begin 230 | dsimp, 231 | rw ←size_erase_add_eq_size_sum Γ, 232 | swap, exact h, 233 | rw [add_comm], 234 | apply add_lt_add_left, dsimp [sizeof, has_sizeof.sizeof, nnf.sizeof], 235 | rw [add_comm, ←add_assoc], 236 | apply nat.lt_add_of_pos_left, 237 | apply nat.succ_pos 238 | end 239 | 240 | theorem split_lt_or_size_right {φ ψ} (Γ : list nnf) (h : nnf.or φ ψ ∈ Γ) : 241 | node_size (ψ :: Γ.erase (nnf.or φ ψ)) < node_size Γ := 242 | begin 243 | dsimp, 244 | rw ←size_erase_add_eq_size_sum Γ, 245 | swap, exact h, 246 | rw [add_comm], 247 | apply add_lt_add_left, dsimp [sizeof, has_sizeof.sizeof, nnf.sizeof], 248 | apply nat.lt_add_of_pos_left, 249 | rw nat.one_add, 250 | apply nat.succ_pos 251 | end 252 | 253 | theorem split_le_or_degree_left {φ ψ} (Γ : list nnf) (h : nnf.or φ ψ ∈ Γ) : 254 | modal_degree (φ :: Γ.erase (nnf.or φ ψ)) ≤ modal_degree Γ := 255 | begin 256 | rw cons_degree, 257 | {apply max_le, 258 | -- works even when exact count_modal (and φ ψ). How smart `apply` is? 259 | repeat {apply le_trans, swap 3, exact count_modal (or φ ψ), simp, apply le_of_mem_degree (nnf.or φ ψ) _ h}, 260 | {apply le_degree_of_subset, apply erase_subset}} 261 | end 262 | 263 | theorem split_le_or_degree_right {φ ψ} (Γ : list nnf) (h : nnf.or φ ψ ∈ Γ) : 264 | modal_degree (ψ :: Γ.erase (nnf.or φ ψ)) ≤ modal_degree Γ := 265 | begin 266 | rw cons_degree, 267 | {apply max_le, 268 | repeat {apply le_trans, swap 3, exact count_modal (or φ ψ), simp, apply le_of_mem_degree (nnf.or φ ψ) _ h}, 269 | {apply le_degree_of_subset, apply erase_subset}} 270 | end 271 | 272 | theorem split_lt_or_seqt_left {φ ψ} (Γ : seqt) (h : nnf.or φ ψ ∈ Γ.main) : 273 | prod.measure_lex seqt_size (or_child_left Γ h) Γ := 274 | begin 275 | have hmem : or φ ψ ∈ Γ.main ++ Γ.hdld, {apply mem_append_left _ h}, 276 | have := lt_or_eq_of_le (split_le_or_degree_left _ hmem), 277 | cases this, 278 | {left, simp only [seqt_size, or_child_left], rw erase_append_left at this, rw cons_append, exact this, exact h}, 279 | {dsimp only [prod.measure_lex, inv_image, seqt_size], rw [erase_append_left _ h] at this, rw ←this, right, 280 | apply split_lt_or_size_left _ h} 281 | end 282 | 283 | theorem split_lt_or_seqt_right {φ ψ} (Γ : seqt) (h : nnf.or φ ψ ∈ Γ.main) : 284 | prod.measure_lex seqt_size (or_child_right Γ h) Γ := 285 | begin 286 | have hmem : or φ ψ ∈ Γ.main ++ Γ.hdld, {apply mem_append_left _ h}, 287 | have := lt_or_eq_of_le (split_le_or_degree_right _ hmem), 288 | cases this, 289 | {left, simp only [seqt_size, or_child_right], rw erase_append_left at this, rw cons_append, exact this, exact h}, 290 | {dsimp only [prod.measure_lex, inv_image, seqt_size], rw [erase_append_left _ h] at this, rw ←this, right, 291 | apply split_lt_or_size_right _ h} 292 | end 293 | 294 | theorem copy_lt_size {φ} (Γ : list nnf) (h : nnf.box φ ∈ Γ) : 295 | node_size (φ :: Γ.erase (nnf.box φ)) < node_size Γ := 296 | begin 297 | dsimp [node_size], 298 | rw ←size_erase_add_eq_size_sum Γ, 299 | swap, exact h, 300 | rw [add_comm], simp [sizeof,has_sizeof.sizeof, nnf.sizeof] 301 | end 302 | 303 | theorem copy_eq_degree {φ} (Γ Λ: list nnf) (h : nnf.box φ ∈ Γ) : 304 | modal_degree (φ :: Γ.erase (nnf.box φ) ++ (box φ :: Λ)) = modal_degree (Γ ++ Λ) := 305 | begin 306 | rw cons_append, rw cons_degree, 307 | have : count_modal φ ≤ count_modal (box φ), {simp}, 308 | have : count_modal φ ≤ (modal_degree (list.erase Γ (box φ) ++ box φ :: Λ)), {apply le_trans, exact this, apply le_of_mem_degree, simp}, 309 | have heq := max_eq_right this, 310 | rw heq, 311 | apply nat.le_antisymm, 312 | {apply le_degree_of_subset, intros x hx, rw mem_append at hx, 313 | cases hx, 314 | {rw mem_append, left, apply mem_of_mem_erase hx}, 315 | {cases hx, rw mem_append, left, rw ←hx at h, exact h, rw mem_append, right, exact hx}}, 316 | {apply le_degree_of_subset, intros x hx, rw mem_append at hx, 317 | cases hx, 318 | {rw mem_append, by_cases heq : x = box φ, 319 | {right, simp [heq]}, 320 | {left, rw mem_erase_of_ne heq, exact hx}}, 321 | {rw mem_append, right, simp [hx]} } 322 | end 323 | 324 | theorem copy_lt_seqt {φ} (Γ : seqt) (h : nnf.box φ ∈ Γ.main) : 325 | prod.measure_lex seqt_size (box_child Γ h) Γ := 326 | begin 327 | have hmem : box φ ∈ Γ.main ++ Γ.hdld, {apply mem_append_left _ h}, 328 | dsimp only [prod.measure_lex, inv_image, seqt_size, box_child], 329 | rw [copy_eq_degree], right, apply copy_lt_size, 330 | repeat {assumption} 331 | end 332 | -------------------------------------------------------------------------------- /src/KT/vanilla.lean: -------------------------------------------------------------------------------- 1 | import .rules 2 | 3 | open psum nnf node 4 | 5 | set_option eqn_compiler.zeta true 6 | 7 | def tableau : Π Γ : seqt, node Γ 8 | | Γ := 9 | match get_contra_seqt Γ with 10 | | inl w := contra_rule_seqt w.2 11 | | inr no_contra := 12 | match get_and_seqt Γ with 13 | | inl p := 14 | let Δ := p.1.1 :: p.1.2 :: Γ.main.erase (and p.val.1 p.val.2) in 15 | let inst := and_instance_seqt.cons p.2 in 16 | have h : prod.measure_lex seqt_size ⟨_,_,_,_⟩ Γ, 17 | begin apply split_lt_and_seqt, exact p.2 end, 18 | let d_delta : node ⟨Δ, Γ.hdld,_,_⟩ := tableau (and_child Γ p.2) in 19 | and_rule_seqt inst d_delta 20 | | inr no_and := 21 | match get_or_seqt Γ with 22 | | inl p := 23 | let Γ₁ := p.1.1 :: Γ.main.erase (nnf.or p.val.1 p.val.2) in 24 | let Γ₂ := p.1.2 :: Γ.main.erase (nnf.or p.val.1 p.val.2) in 25 | let inst := or_instance_seqt.cons p.2 in 26 | have h₁ : prod.measure_lex seqt_size ⟨Γ₁, Γ.hdld,_,_⟩ Γ, 27 | begin apply split_lt_or_seqt_left, exact p.2 end, 28 | have h₂ : prod.measure_lex seqt_size ⟨Γ₂, Γ.hdld,_,_⟩ Γ, 29 | begin apply split_lt_or_seqt_right, exact p.2 end, 30 | let d_Γ₁ : node ⟨Γ₁, Γ.hdld,_,_⟩ := tableau (or_child_left Γ p.2) in 31 | match d_Γ₁ with 32 | | closed pr := or_rule_seqt inst (closed pr) (tableau (or_child_right Γ p.2)) 33 | | open_ w := open_rule_seqt inst w.2 34 | end 35 | | inr no_or := 36 | match get_box_seqt Γ with 37 | | inl p := 38 | let Γ₁ := p.1 :: Γ.main.erase (nnf.box p.1) in 39 | let inst := copy_instance_seqt.cons p.2 in 40 | have h : prod.measure_lex seqt_size ⟨Γ₁, box p.1 :: Γ.hdld,_,_⟩ Γ, 41 | begin apply copy_lt_seqt, exact p.2 end, 42 | let d_delta : node ⟨Γ₁, box p.1 :: Γ.hdld,_,_⟩ := tableau (box_child Γ p.2) in 43 | copy_rule_seqt inst d_delta 44 | | inr no_box := 45 | match get_dia_seqt Γ with 46 | | inl p := 47 | let ma : modal_applicable Γ := 48 | {satu := {no_and := no_and, no_or := no_or}, 49 | no_contra_main := no_contra, 50 | no_box_main := no_box, 51 | v := get_var Γ.main, 52 | hv := λ n, get_var_iff, 53 | φ := p.1, 54 | ex := p.2} in 55 | let l := @dia_rule_seqt (λ Δ, prod.measure_lex seqt_size Δ Γ) 56 | (λ x h, tableau x) (unmodal_seqt Γ) 57 | (unmodal_seqt_size Γ) in 58 | match l with 59 | | inl w := 60 | begin left, {exact unsat_of_unsat_unmodal ma w.1 w.2} end 61 | | inr w := 62 | begin right, split, apply sat_of_batch_sat, exact ma, exact w.2 end 63 | end 64 | | inr no_dia := 65 | let mc : model_constructible Γ := 66 | {satu := {no_and := no_and, no_or := no_or}, 67 | no_box_main := no_box, 68 | no_contra_main := no_contra, 69 | v := get_var Γ.main, 70 | hv := λ n, get_var_iff, 71 | no_dia := no_dia} in 72 | begin right, split, apply build_model_seqt, exact mc end 73 | end 74 | end 75 | end 76 | end 77 | end 78 | using_well_founded {rel_tac := λ _ _, `[exact ⟨_, prod.measure_lex_wf seqt_size⟩], dec_tac := `[assumption]} 79 | 80 | @[simp] def mk_seqt (Γ : list nnf) : seqt := 81 | { main := Γ, 82 | hdld := [], 83 | pmain := begin 84 | intros v l φ h hb hall, exfalso, 85 | apply list.not_mem_nil, exact hb 86 | end, 87 | phdld := box_only_nil } 88 | 89 | def is_sat (Γ : list nnf) : bool := 90 | match tableau (mk_seqt Γ) with 91 | | closed _ := ff 92 | | open_ _ := tt 93 | end 94 | 95 | theorem correctness (Γ : list nnf) : is_sat Γ = tt ↔ ∃ (st : Type) (k : KT st) s, sat k s Γ := 96 | begin 97 | cases h : is_sat Γ, 98 | constructor, 99 | {intro, contradiction}, 100 | {intro hsat, cases eq : tableau (mk_seqt Γ), 101 | rcases hsat with ⟨w, k, s, hsat⟩, 102 | apply false.elim, apply a, simp, exact hsat, 103 | {dsimp [is_sat] at h, dsimp at eq, rw eq at h, contradiction} }, 104 | {split, intro, dsimp [is_sat] at h, 105 | cases eq : tableau (mk_seqt Γ), 106 | { dsimp at eq, rw eq at h, contradiction }, 107 | { split, split, split, have := a_1.2, simp at this, exact this}, 108 | { simp } } 109 | end 110 | 111 | def test := [box (var 1), (neg 1)] 112 | 113 | #eval is_sat test 114 | -------------------------------------------------------------------------------- /src/S4/S4_defs.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2018-2019 Minchao Wu. All rights reserved. 3 | Released under MIT license as described in the file LICENSE. 4 | Author: Minchao Wu 5 | -/ 6 | 7 | import defs data.list.perm .data 8 | open nnf tactic 9 | 10 | meta def frame_restriction : tactic unit := 11 | do intro `a >> `[simp] 12 | 13 | structure S4 (states : Type) extends kripke states := 14 | (refl : reflexive rel . frame_restriction) 15 | (trans : transitive rel . frame_restriction) 16 | 17 | instance inhabited_S4 : inhabited (S4 ℕ) := 18 | ⟨ { val := λ a b, tt, rel := λ a b, tt } ⟩ 19 | 20 | @[simp] def force {states : Type} (k : S4 states) : states → nnf → Prop 21 | | s (var n) := k.val n s 22 | | s (neg n) := ¬ k.val n s 23 | | s (and φ ψ) := force s φ ∧ force s ψ 24 | | s (or φ ψ) := force s φ ∨ force s ψ 25 | | s (box φ) := ∀ s', k.rel s s' → force s' φ 26 | | s (dia φ) := ∃ s', k.rel s s' ∧ force s' φ 27 | 28 | def sat {st} (k : S4 st) (s) (Γ : list nnf) : Prop := 29 | ∀ φ ∈ Γ, force k s φ 30 | 31 | def unsatisfiable (Γ : list nnf) : Prop := 32 | ∀ (st) (k : S4 st) s, ¬ sat k s Γ 33 | 34 | theorem unsat_singleton {φ} : unsatisfiable [φ] → ∀ (st) (k : S4 st) s, ¬ force k s φ 35 | := 36 | begin 37 | intros h _ _ _ hf, 38 | apply h, intros ψ hψ, rw list.mem_singleton at hψ, rw hψ, exact hf 39 | end 40 | 41 | theorem sat_of_empty {st} (k : S4 st) (s) : sat k s [] := 42 | λ φ h, absurd h $ list.not_mem_nil _ 43 | 44 | theorem ne_empty_of_unsat {Γ} (h : unsatisfiable Γ): Γ ≠ [] := 45 | begin 46 | intro heq, rw heq at h, 47 | apply h, apply sat_of_empty, exact nat, 48 | apply inhabited_S4.1, exact 0 49 | end 50 | 51 | open tmodel 52 | 53 | @[simp] def minfo : Π m : tmodel, info 54 | | (cons i l ba) := i 55 | 56 | @[simp] def htk : Π m : tmodel, list nnf 57 | | (cons i l ba) := i.htk 58 | 59 | def hist : Π m : tmodel, list nnf 60 | | (cons i l ba) := i.id.h 61 | 62 | @[simp] def msig : Π m : tmodel, sig 63 | | (cons i l ba) := i.id.s 64 | 65 | @[simp] def manc : Π m : tmodel, list psig 66 | | (cons i l ba) := i.id.a 67 | 68 | def bhist : Π m : tmodel, list nnf 69 | | (cons i l ba) := i.id.b 70 | 71 | @[simp] def request : Π m : tmodel, list psig 72 | | (cons i l ba) := ba 73 | 74 | @[simp] def proper_request_box : Π m : tmodel, Prop 75 | | (cons i l ba) := ∀ rq : psig, rq ∈ ba → ∀ φ, (box φ ∈ i.htk ∨ box φ ∈ i.id.b) → box φ ∈ rq.b 76 | 77 | @[simp] def subset_request : Π m : tmodel, Prop 78 | | (cons i l ba) := ba ⊆ i.id.a 79 | 80 | @[simp] def tmodel_step_bhist : Π m : tmodel, Prop 81 | | m@(cons i l ba) := ∀ s ∈ l, ∀ φ, box φ ∈ i.id.b → box φ ∈ htk s 82 | 83 | @[simp] def tmodel_step_box : Π m : tmodel, Prop 84 | | m@(cons i l ba) := ∀ s ∈ l, ∀ φ, box φ ∈ i.htk → box φ ∈ htk s 85 | 86 | -- Can be strenghtened 87 | @[simp] def tmodel_dia : Π m : tmodel, Prop 88 | | m@(cons i l ba) := ∀ φ, dia φ ∈ i.htk → (∃ rq : psig, rq ∈ ba ∧ rq.d = φ) ∨ ∃ s ∈ l, φ ∈ htk s 89 | 90 | @[simp] def child : tmodel → tmodel → bool 91 | | s (cons i l ba) := s ∈ l 92 | 93 | inductive tc' {α : Type} (r : α → α → Prop) : α → α → Prop 94 | | base : ∀ a b, r a b → tc' a b 95 | | step : ∀ a b c, r a b → tc' b c → tc' a c 96 | 97 | theorem tc'.trans {α : Type} {r : α → α → Prop} {a b c : α} : 98 | tc' r a b → tc' r b c → tc' r a c := 99 | begin 100 | intros h₁ h₂, 101 | induction h₁, 102 | apply tc'.step, exact h₁_a_1, exact h₂, 103 | apply tc'.step, exact h₁_a_1, apply h₁_ih, exact h₂ 104 | end 105 | 106 | def desc : tmodel → tmodel → Prop := tc' (λ s m, child s m) 107 | 108 | theorem desc_not_nil : Π c i ba m, m = cons i [] ba → desc c m → false := 109 | begin 110 | intros c i ba m heq h, 111 | induction h, 112 | {rw heq at h_a_1, simp at h_a_1, exact h_a_1}, 113 | {apply h_ih, exact heq} 114 | end 115 | 116 | theorem desc_iff_eq_child_aux : Π i₁ i₂ s₁ s₂ l₁ l₂ m₁ m₂ m₃, 117 | m₁ = cons i₁ l₁ s₁ → m₂ = cons i₂ l₂ s₂ → l₁ = l₂ → 118 | (desc m₃ m₁ ↔ desc m₃ m₂) := 119 | begin 120 | intros i₁ i₂ s₁ s₂ l₁ l₂ m₁ m₂ m₃ heq₁ heq₂ heq, 121 | split, 122 | {intro hd, induction hd, 123 | {rw heq₁ at hd_a_1, simp at hd_a_1, 124 | rw heq at hd_a_1, 125 | apply tc'.base, rw heq₂, simp, exact hd_a_1}, 126 | {apply tc'.trans, apply tc'.base, exact hd_a_1, apply hd_ih, exact heq₁}}, 127 | {intro hd, induction hd, 128 | {rw heq₂ at hd_a_1, simp at hd_a_1, 129 | rw ←heq at hd_a_1, 130 | apply tc'.base, rw heq₁, simp, exact hd_a_1}, 131 | {apply tc'.trans, apply tc'.base, exact hd_a_1, apply hd_ih, exact heq₂}} 132 | end 133 | 134 | theorem eq_desc_of_eq_children {i₁ i₂ s₁ s₂ l c} : 135 | desc c (cons i₁ l s₁) = desc c (cons i₂ l s₂) := 136 | begin rw desc_iff_eq_child_aux, repeat {refl} end 137 | 138 | theorem desc_step : Π c i l ba, c ∈ l → desc c (cons i l ba) 139 | | c i [] ba h := absurd h $ list.not_mem_nil _ 140 | | c i (hd::tl) ba h := 141 | begin 142 | constructor, 143 | simp, cases h, 144 | left, exact h, right, exact h 145 | end 146 | 147 | theorem desc_ex : Π c i l ba, (∃ m ∈ l, desc c m) → desc c (cons i l ba) 148 | | c i [] ba h := begin rcases h with ⟨w, hmem, hw⟩, exact (absurd hmem $ list.not_mem_nil _) end 149 | | c i (hd::tl) ba h := 150 | begin 151 | rcases h with ⟨w, hmem, hw⟩, 152 | cases hw, 153 | {apply tc'.step, 154 | swap 3, {exact w}, 155 | {exact hw_a_1}, 156 | {apply tc'.base, simp, cases hmem, left, exact hmem, right, exact hmem}}, 157 | {apply tc'.step, exact hw_a_1, apply tc'.trans, exact hw_a_2, apply tc'.base, simp, exact hmem} 158 | end 159 | 160 | theorem ex_desc : Π c i l ba m, m = (cons i l ba) → desc c m → (c ∈ l ∨ ∃ m ∈ l, desc c m) := 161 | begin 162 | intros c i l ba m heq h, 163 | induction h, 164 | {left, rw heq at h_a_1, simp at h_a_1, exact h_a_1}, 165 | {cases h_ih heq, 166 | {right, split, split, exact h, apply tc'.base, exact h_a_1}, 167 | {rcases h with ⟨w, hmem, hw⟩, right, split, split, exact hmem, apply tc'.step, exact h_a_1, exact hw}} 168 | end 169 | 170 | theorem ex_desc' : Π c i l ba, desc c (cons i l ba) → (c ∈ l ∨ ∃ m ∈ l, desc c m) := 171 | begin intros c i l ba h, apply ex_desc, repeat {refl}, exact h end 172 | 173 | @[simp] def tmodel_anc : Π m : tmodel, Prop 174 | | m@(cons i l ba) := ∀ s rq, desc s m → rq ∈ request s → 175 | (rq ∈ manc m) ∨ 176 | (∃ d, desc d m ∧ some rq = msig d) 177 | 178 | structure ptmodel (m : tmodel) : Prop := 179 | (bhist : tmodel_step_bhist m) 180 | (sbox : tmodel_step_box m) 181 | (pdia : tmodel_dia m) 182 | (bdia : tmodel_anc m) 183 | (reqb : proper_request_box m) 184 | (sreq : subset_request m) 185 | 186 | def global_pt (m : tmodel) := ∀ s, desc s m → ptmodel s 187 | 188 | open subtype 189 | 190 | def model : Type := {m : tmodel // ptmodel m ∧ global_pt m} 191 | 192 | def rmodel : Type := {m : tmodel // ptmodel m} 193 | 194 | inductive reach_step : rmodel → rmodel → Prop 195 | | fwd_base (s : rmodel) (i l ba h) : s.1 ∈ l → reach_step ⟨(cons i l ba), h⟩ s 196 | | bwd_base (s : rmodel) (i l ba h) : (∃ rq ∈ ba, some rq = msig s.1) → reach_step ⟨(cons i l ba), h⟩ s 197 | 198 | theorem reach_step_box (s₁ s₂ φ) (h₁ : reach_step s₁ s₂) (h₂ : box φ ∈ htk s₁.1) : box φ ∈ htk s₂.1 := 199 | begin 200 | cases h₁, 201 | {cases s₂ with s₂ ps₂, 202 | cases s₂ with i₂ l₂ sg₂, 203 | simp, 204 | have := h₁_h.sbox, 205 | simp at this, simp at h₂, 206 | have hmem := this _ h₁_a _ h₂, 207 | simp at hmem, exact hmem }, 208 | {cases s₂ with s₂ ps₂, 209 | cases s₂ with i₂ l₂ sg₂, 210 | simp, 211 | rcases h₁_a with ⟨w,hmem,hw⟩, 212 | simp at hw, 213 | apply i₂.mhtk, 214 | have := i₂.id.ps₂, 215 | rw ←hw at this, 216 | have hneq : some w ≠ none, 217 | {intro heq, contradiction}, 218 | have hsub := this hneq, 219 | apply hsub, 220 | have := h₁_h.reqb, 221 | simp at this, simp at h₂, 222 | have hc := this w hmem φ (or.inl h₂), 223 | cases w, 224 | dsimp [bsig], exact hc} 225 | end 226 | 227 | inductive rtc {α : Type} (r : α → α → Prop) : α → α → Prop 228 | | refl : Π a, rtc a a 229 | | step : Π a b c, r a b → rtc b c → rtc a c 230 | 231 | theorem rtc.trans {α : Type} {r : α → α → Prop} {a b c : α} : 232 | rtc r a b → rtc r b c → rtc r a c := 233 | begin 234 | intros h₁ h₂, 235 | induction h₁, 236 | exact h₂, 237 | apply rtc.step, exact h₁_a_1, apply h₁_ih, exact h₂ 238 | end 239 | 240 | theorem rtc_step {α : Type} {r : α → α → Prop} {a b : α} (h : r a b) : 241 | rtc r a b := 242 | by apply rtc.step _ _ _ h; apply rtc.refl 243 | 244 | def reach (s₁ s₂ : rmodel) := rtc reach_step s₁ s₂ 245 | 246 | theorem refl_reach : Π s, reach s s := λ s, rtc.refl s 247 | 248 | theorem trans_reach : Π s₁ s₂ s₃, reach s₁ s₂ → reach s₂ s₃ → reach s₁ s₃ := λ s₁ s₂ s₃ h₁ h₂, rtc.trans h₁ h₂ 249 | 250 | @[simp] def builder (m : tmodel) : S4 {x : rmodel // x.1 = m ∨ desc x.1 m} := 251 | {val := λ v s, var v ∈ htk s.1.1, 252 | rel := λ s₁ s₂, reach s₁ s₂, 253 | refl := λ s, refl_reach s, 254 | trans := λ a b c, trans_reach a b c} 255 | 256 | open rtc 257 | 258 | theorem reach_box (s₁ s₂ φ) (h₁ : reach s₁ s₂) (h₂ : box φ ∈ htk s₁.1) : φ ∈ htk s₂.1 := 259 | begin 260 | induction h₁ with m m₁ m₂ m₃ h₁₂ h₂₃ ih, 261 | {cases m with tm hm, cases tm with i l sg, 262 | simp, 263 | apply i.hhtk.hbox, 264 | simp at h₂, exact h₂}, 265 | {apply ih, 266 | apply reach_step_box, 267 | exact h₁₂, exact h₂} 268 | end 269 | 270 | theorem reach_step_dia (s : rmodel) (rt : model) (φ) 271 | (h₁ : desc s.1 rt.1) 272 | (h₂ : manc rt.1 = []) (h₃ : dia φ ∈ htk s.1) : 273 | ∃ s', reach_step s s' ∧ φ ∈ htk s'.1 ∧ desc s'.1 rt.1 := 274 | begin 275 | cases s with s ps, 276 | cases s with i l sg, 277 | have := ps.pdia, 278 | simp at this, simp at h₃, 279 | have hc := this _ h₃, 280 | cases hc, 281 | {cases rt with rt prt, 282 | cases rt with irt lrt sgrt, 283 | rcases hc with ⟨w, hmem, hw⟩, 284 | have := prt.1.bdia, 285 | simp at this, simp at h₁, 286 | have hcaux := this _ w h₁, 287 | simp at hcaux, 288 | have hcc := hcaux hmem, 289 | simp at h₂, 290 | cases hcc, 291 | {rw h₂ at hcc, exfalso, apply list.not_mem_nil, exact hcc}, 292 | {rcases hcc with ⟨m, hml, hmr⟩, 293 | have pm := prt.2 m hml, 294 | split, split, 295 | swap 3, exact ⟨m, pm⟩, 296 | apply reach_step.bwd_base, 297 | split, split, exact hmem, simp, exact hmr, 298 | split, 299 | {cases m with im lm sgm, simp, 300 | apply im.mhtk, 301 | have := im.id.ps₁, 302 | simp at hmr, rw ←hmr at this, 303 | have hneq : some w ≠ none, {intro, contradiction}, 304 | have hmem := this hneq, 305 | cases w, dsimp [dsig] at hmem, 306 | rw ←hw, exact hmem_1 }, 307 | {exact hml} } }, 308 | {rcases hc with ⟨m, pml, pmr⟩, 309 | have hdm : desc m rt.1, 310 | {apply tc'.trans, apply tc'.base, 311 | swap 3, exact (⟨cons i l sg, ps⟩ : rmodel).val, 312 | simp, exact pml, exact h₁}, 313 | cases rt with rt prt, 314 | cases rt with irt lrt sgrt, 315 | have pm := prt.2 m hdm, 316 | split, split, swap 3, 317 | exact ⟨m, pm⟩, 318 | apply reach_step.fwd_base, 319 | exact pml, split, 320 | {exact pmr}, 321 | {exact hdm} } 322 | end 323 | 324 | theorem reach_dia (s : rmodel) (rt : model) (φ) 325 | (h₁ : desc s.1 rt.1) 326 | (h₂ : manc rt.1 = []) (h₃ : dia φ ∈ htk s.1) : 327 | ∃ s', reach s s' ∧ φ ∈ htk s'.1 ∧ desc s'.1 rt.1:= 328 | begin 329 | have := reach_step_dia s rt φ h₁ h₂ h₃, 330 | rcases this with ⟨w, hwl, hwr⟩, 331 | split, split, swap 3, exact w, 332 | apply rtc_step hwl, exact hwr 333 | end 334 | 335 | theorem reach_step_dia_root (s : rmodel) (rt : model) (φ) 336 | (h₁ : s.1 = rt.1) 337 | (h₂ : manc rt.1 = []) (h₃ : dia φ ∈ htk s.1) : 338 | ∃ s', reach_step s s' ∧ φ ∈ htk s'.1 ∧ desc s'.1 rt.1 := 339 | begin 340 | cases s with s ps, 341 | cases s with is ls sgs, 342 | have := ps.pdia, 343 | simp at this, simp at h₃, 344 | have hc := this _ h₃, 345 | cases hc, 346 | {have := ps.sreq, simp at this, 347 | rcases hc with ⟨w, hmw, hw⟩, 348 | have hmem := this hmw, 349 | cases rt with rt prt, 350 | rw ←h₁ at h₂, 351 | simp at h₂, rw h₂ at hmem, 352 | exfalso, apply list.not_mem_nil, exact hmem}, 353 | {cases rt with rt prt, 354 | rcases hc with ⟨w, hwl, hwr⟩, 355 | have ptw : ptmodel w, 356 | {apply prt.2, apply tc'.base, simp at h₁, rw ←h₁, simp, exact hwl}, 357 | split, split, swap 3, exact ⟨w, ptw⟩, 358 | apply reach_step.fwd_base, exact hwl, split, 359 | {exact hwr}, 360 | {apply tc'.base, simp, simp at h₁, rw ←h₁, simp, exact hwl} } 361 | end 362 | 363 | theorem reach_dia_root (s : rmodel) (rt : model) (φ) 364 | (h₁ : s.1 = rt.1) 365 | (h₂ : manc rt.1 = []) (h₃ : dia φ ∈ htk s.1) : 366 | ∃ s', reach s s' ∧ φ ∈ htk s'.1 ∧ desc s'.1 rt.1 := 367 | begin 368 | have := reach_step_dia_root s rt φ h₁ h₂ h₃, 369 | rcases this with ⟨w, hwl, hwr⟩, 370 | split, split, swap 3, exact w, 371 | apply rtc_step hwl, exact hwr 372 | end 373 | 374 | theorem good_model (m : model) (hrt : manc m.1 = []): 375 | Π (s : {x : rmodel // x.1 = m.1 ∨ desc x.1 m.1}) (φ : nnf), 376 | φ ∈ htk s.1.1 → force (builder m.1) s φ 377 | | s (var n) h := begin simp, exact h end 378 | | s (neg n) h := begin 379 | simp, intro hin, 380 | cases s with s ps, 381 | cases s with s pts, 382 | cases s with i l sg, 383 | have := i.hhtk.hno_contra, 384 | simp at hin, 385 | apply this hin, simp at h, exact h 386 | end 387 | | s (and φ ψ) h := begin 388 | split, 389 | {apply good_model, 390 | cases s with s ps, 391 | cases s with s pts, 392 | cases s with i l sg, 393 | have := i.hhtk.hand_left, 394 | simp, apply this, simp at h, exact h}, 395 | {apply good_model, 396 | cases s with s ps, 397 | cases s with s pts, 398 | cases s with i l sg, 399 | have := i.hhtk.hand_right, 400 | simp, apply this, simp at h, exact h} 401 | end 402 | | s (or φ ψ) h := begin 403 | cases s with s ps, 404 | cases s with s pts, 405 | cases s with i l sg, 406 | have := i.hhtk.hor, 407 | simp at h, 408 | have hc := this h, 409 | cases hc, 410 | {simp, left, apply good_model, simp, exact hc}, 411 | {simp, right, apply good_model, simp, exact hc} 412 | end 413 | | s (box φ) h := begin 414 | intros m hm, 415 | apply good_model, 416 | apply reach_box, 417 | exact hm, 418 | exact h 419 | end 420 | | s (dia φ) h := begin 421 | cases s with s ps, 422 | cases ps, 423 | {simp, simp at h, 424 | have := reach_dia_root _ _ _ ps hrt h, 425 | rcases this with ⟨s', hs'l, hs'm, hs'r⟩, 426 | split, split, 427 | exact hs'l, split, apply good_model, simp, 428 | exact hs'm, right, exact hs'r }, 429 | {simp, simp at h, 430 | have := reach_dia _ _ _ ps hrt h, 431 | rcases this with ⟨s', hs'l, hs'm, hs'r⟩, 432 | split, split, 433 | exact hs'l, split, apply good_model, simp, 434 | exact hs'm, right, exact hs'r} 435 | end 436 | -------------------------------------------------------------------------------- /src/S4/data.lean: -------------------------------------------------------------------------------- 1 | import defs data.list.perm 2 | open nnf 3 | 4 | namespace list 5 | universes u v w 6 | variables {α : Type u} 7 | 8 | theorem length_sub_lt_of_nodup_subperm [decidable_eq α] {l₁ l₂ : list α} {a : α} 9 | (h₁ : l₁ <+~ l₂) (h₂ : a ∈ l₂) (h₃ : a ∉ l₁) (h₄ : nodup l₁): 10 | length l₂ - length (a :: l₁) < length l₂ - length l₁ 11 | := 12 | begin 13 | rw nat.sub_lt_sub_left_iff, 14 | { simp [zero_lt_one] }, 15 | { apply subperm.length_le, 16 | apply cons_subperm_of_mem; assumption } 17 | end 18 | 19 | end list 20 | 21 | structure psig : Type := 22 | (d : nnf) 23 | (b : list nnf) 24 | 25 | instance : decidable_eq psig := by tactic.mk_dec_eq_instance 26 | 27 | def sig : Type := option psig 28 | 29 | instance : decidable_eq sig := by tactic.mk_dec_eq_instance 30 | 31 | def dsig : Π (s : sig) (h : s ≠ none), nnf 32 | | none h := by contradiction 33 | | (some ⟨d, b⟩) h := d 34 | 35 | def bsig : Π (s : sig) (h : s ≠ none), list nnf 36 | | none h := by contradiction 37 | | (some ⟨d, b⟩) h := b 38 | 39 | instance (Γ) : decidable_eq (no_literals Γ) := by tactic.mk_dec_eq_instance 40 | 41 | instance (Γ) : decidable_eq (saturated Γ) := by tactic.mk_dec_eq_instance 42 | 43 | instance (Γ) : decidable_eq (box_only Γ) := by tactic.mk_dec_eq_instance 44 | 45 | structure sseqt : Type := 46 | (goal : list nnf) 47 | (s : sig) -- sig := option psig 48 | (a : list psig) -- psig is the signature, which is of the form (d, b) 49 | (h b m: list nnf) 50 | (ndh : list.nodup h) -- nodup says there is no duplicate elements 51 | (ndb : list.nodup b) 52 | (sph : h <+~ closure goal) -- <+~ denotes sublist permutation 53 | (spb : b <+~ closure goal) 54 | (sbm : m ⊆ closure goal) 55 | (ha : ∀ φ ∈ h, (⟨φ, b⟩ : psig) ∈ a) 56 | (hb : box_only b) 57 | -- dsig takes a signature (d, b) and a proof h, and returns d 58 | (ps₁ : Π (h : s ≠ none), dsig s h ∈ m) 59 | -- bsig takes a signature (d, b) and a proof h, and returns b 60 | (ps₂ : Π (h : s ≠ none), bsig s h ⊆ m) 61 | 62 | instance : decidable_eq sseqt := by tactic.mk_dec_eq_instance 63 | 64 | class val_constructible (Γ : sseqt) := 65 | (satu : saturated Γ.m) 66 | (no_box_main : ∀ {φ}, box φ ∉ Γ.m) 67 | (no_contra_main : ∀ {n}, var n ∈ Γ.m → neg n ∉ Γ.m) 68 | 69 | class modal_applicable (Γ : sseqt) extends val_constructible Γ := 70 | (φ : nnf) 71 | (ex : dia φ ∈ Γ.m) 72 | 73 | class model_constructible (Γ : sseqt) extends val_constructible Γ := 74 | (no_dia : ∀ {φ}, nnf.dia φ ∉ Γ.m) 75 | 76 | def and_child {φ ψ} (Γ : sseqt) (h : nnf.and φ ψ ∈ Γ.m) : sseqt := 77 | { goal := Γ.goal, 78 | s := none, 79 | m := φ :: ψ :: Γ.m.erase (and φ ψ), 80 | sbm := begin 81 | intros x hx, cases hx, 82 | {rw hx, apply (mem_closure_and _ (Γ.sbm h)).1}, 83 | {cases hx, 84 | {rw hx, apply (mem_closure_and _ (Γ.sbm h)).2}, 85 | {apply Γ.sbm, apply list.erase_subset, exact hx}} 86 | end, 87 | ps₁ := by intro; contradiction, 88 | ps₂ := by intro; contradiction, 89 | .. Γ} 90 | 91 | inductive and_instance_seqt (Γ : sseqt) : sseqt → Type 92 | | cons : Π {φ ψ} (h : nnf.and φ ψ ∈ Γ.m), 93 | and_instance_seqt $ and_child Γ h 94 | 95 | def or_child_left {φ ψ} (Γ : sseqt) (h : nnf.or φ ψ ∈ Γ.m) : sseqt := 96 | { goal := Γ.goal, 97 | s := none, 98 | m := φ :: Γ.m.erase (or φ ψ), 99 | sbm := begin 100 | intros x hx, cases hx, 101 | {rw hx, apply (mem_closure_or _ (Γ.sbm h)).1}, 102 | {apply Γ.sbm, apply list.erase_subset, exact hx} 103 | end, 104 | ps₁ := by intro; contradiction, 105 | ps₂ := by intro; contradiction, 106 | .. Γ} 107 | 108 | def or_child_right {φ ψ} (Γ : sseqt) (h : nnf.or φ ψ ∈ Γ.m) : sseqt := 109 | { goal := Γ.goal, 110 | s := none, 111 | m := ψ :: Γ.m.erase (or φ ψ), 112 | sbm := begin 113 | intros x hx, cases hx, 114 | {rw hx, apply (mem_closure_or _ (Γ.sbm h)).2}, 115 | {apply Γ.sbm, apply list.erase_subset, exact hx} 116 | end, 117 | ps₁ := by intro; contradiction, 118 | ps₂ := by intro; contradiction, 119 | .. Γ} 120 | 121 | inductive or_instance_seqt (Γ : sseqt) : sseqt → sseqt → Type 122 | | cons : Π {φ ψ} (h : nnf.or φ ψ ∈ Γ.m), 123 | or_instance_seqt (or_child_left Γ h) (or_child_right Γ h) 124 | 125 | def box_child_new {φ} (Γ : sseqt) (h₁ : nnf.box φ ∈ Γ.m) (h₂ : nnf.box φ ∉ Γ.b) : sseqt := 126 | { goal := Γ.goal, 127 | s := none, 128 | h := [], 129 | b := box φ :: Γ.b, 130 | m := φ :: Γ.m.erase (box φ), 131 | ndh := by simp, 132 | ndb := begin rw list.nodup_cons, split, exact h₂, exact Γ.ndb end, 133 | sph := begin apply list.nil_subperm end, 134 | spb := begin 135 | apply list.cons_subperm_of_mem Γ.ndb h₂, 136 | apply Γ.sbm h₁, apply Γ.spb 137 | end, 138 | sbm := begin 139 | intros x hx, cases hx, 140 | {rw hx, apply mem_closure_box _ (Γ.sbm h₁)}, 141 | {apply Γ.sbm, apply list.erase_subset, exact hx} 142 | end, 143 | ha := λ φ h, absurd h $ list.not_mem_nil _, 144 | hb := cons_box_only Γ.hb, 145 | ps₁ := by intro; contradiction, 146 | ps₂ := by intro; contradiction, 147 | .. Γ} 148 | 149 | inductive box_new_instance_seqt (Γ : sseqt) : sseqt → Type 150 | | cons : Π {φ} (h₁ : nnf.box φ ∈ Γ.m) (h₂ : nnf.box φ ∉ Γ.b), 151 | box_new_instance_seqt $ box_child_new Γ h₁ h₂ 152 | 153 | def box_child {φ} (Γ : sseqt) (h₁ : nnf.box φ ∈ Γ.m) : sseqt := 154 | { goal := Γ.goal, 155 | s := none, 156 | a := Γ.a, 157 | h := Γ.h, 158 | b := Γ.b, 159 | m := φ :: Γ.m.erase (box φ), 160 | ndh := Γ.ndh, 161 | ndb := Γ.ndb, 162 | sph := Γ.sph, 163 | spb := Γ.spb, 164 | sbm := begin 165 | intros x hx, cases hx, 166 | {rw hx, apply mem_closure_box _ (Γ.sbm h₁)}, 167 | {apply Γ.sbm, apply list.erase_subset, exact hx} 168 | end, 169 | ha := Γ.ha, 170 | hb := Γ.hb, 171 | ps₁ := by intro; contradiction, 172 | ps₂ := by intro; contradiction} 173 | 174 | inductive box_dup_instance_seqt (Γ : sseqt) : sseqt → Type 175 | | cons : Π {φ} (h₁ : nnf.box φ ∈ Γ.m) (h₂ : nnf.box φ ∈ Γ.b), 176 | box_dup_instance_seqt $ box_child Γ h₁ 177 | 178 | 179 | theorem hintikka_vc {Γ} (h : val_constructible Γ) : hintikka Γ.m := 180 | {hno_contra := h.no_contra_main, 181 | hand_left := begin intros φ ψ h₁, exfalso, apply h.satu.no_and, exact h₁ end, 182 | hand_right := begin intros φ ψ h₁, exfalso, apply h.satu.no_and, exact h₁ end, 183 | hor := begin intros φ ψ h₁, exfalso, apply h.satu.no_or, exact h₁ end, 184 | hbox := begin intros φ h₁, exfalso, apply h.no_box_main, exact h₁ end} 185 | 186 | theorem hintikka_ma {Γ} (h : modal_applicable Γ) : hintikka Γ.m := 187 | hintikka_vc h.to_val_constructible 188 | 189 | theorem hintikka_mc {Γ} (h : model_constructible Γ) : hintikka Γ.m := 190 | hintikka_vc h.to_val_constructible 191 | 192 | structure info : Type := 193 | (id : sseqt) 194 | (htk : list nnf) 195 | (hhtk : hintikka htk) 196 | (mhtk : id.m ⊆ htk) 197 | 198 | instance : decidable_eq info := by tactic.mk_dec_eq_instance 199 | 200 | inductive tmodel 201 | | cons : info → list tmodel → list psig → tmodel 202 | 203 | instance : decidable_eq tmodel := by tactic.mk_dec_eq_instance 204 | -------------------------------------------------------------------------------- /src/S4/full_language.lean: -------------------------------------------------------------------------------- 1 | import .semantics 2 | 3 | inductive fml 4 | | var (n : nat) 5 | | neg (φ : fml) 6 | | and (φ ψ : fml) 7 | | or (φ ψ : fml) 8 | | impl (φ ψ : fml) 9 | | box (φ : fml) 10 | | dia (φ : fml) 11 | 12 | 13 | open fml 14 | 15 | def fml.to_string : fml → string 16 | | (var n) := "P" ++ n.repr 17 | | (neg ϕ) := "¬" ++ "(" ++ fml.to_string ϕ ++ ")" 18 | | (and φ ψ) := fml.to_string φ ++ "∧" ++ fml.to_string ψ 19 | | (or φ ψ) := fml.to_string φ ++ "∨" ++ fml.to_string ψ 20 | | (impl φ ψ) := fml.to_string φ ++ "→" ++ fml.to_string ψ 21 | | (box φ) := "□" ++ "(" ++ fml.to_string φ ++ ")" 22 | | (dia φ) := "◇" ++ "(" ++ fml.to_string φ ++ ")" 23 | 24 | instance fml_repr : has_repr fml := ⟨fml.to_string⟩ 25 | 26 | instance dec_eq_fml : decidable_eq fml := by tactic.mk_dec_eq_instance 27 | 28 | @[simp] def fml_size : fml → ℕ 29 | | (var n) := 1 30 | | (neg φ) := fml_size φ + 1 31 | | (and φ ψ) := fml_size φ + fml_size ψ + 1 32 | | (or φ ψ) := fml_size φ + fml_size ψ + 1 33 | | (impl φ ψ) := fml_size φ + fml_size ψ + 2 34 | | (box φ) := fml_size φ + 1 35 | | (dia φ) := fml_size φ + 1 36 | 37 | @[simp] def fml_force {states : Type} (k : S4 states) : states → fml → Prop 38 | | s (var n) := k.val n s 39 | | s (neg φ) := ¬ fml_force s φ 40 | | s (and φ ψ) := fml_force s φ ∧ fml_force s ψ 41 | | s (or φ ψ) := fml_force s φ ∨ fml_force s ψ 42 | | s (impl φ ψ) := ¬ fml_force s φ ∨ fml_force s ψ 43 | | s (box φ) := ∀ s', k.rel s s' → fml_force s' φ 44 | | s (dia φ) := ∃ s', k.rel s s' ∧ fml_force s' φ 45 | 46 | @[simp] def fml.to_nnf : fml → nnf 47 | | (var n) := nnf.var n 48 | | (neg (var n)) := nnf.neg n 49 | | (neg (neg φ)) := fml.to_nnf φ 50 | | (neg (and φ ψ)) := nnf.or (fml.to_nnf (neg φ)) (fml.to_nnf (neg ψ)) 51 | | (neg (or φ ψ)) := nnf.and (fml.to_nnf (neg φ)) (fml.to_nnf (neg ψ)) 52 | | (neg (impl φ ψ)) := nnf.and (fml.to_nnf φ) (fml.to_nnf (neg ψ)) 53 | | (neg (box φ)) := nnf.dia (fml.to_nnf (neg φ)) 54 | | (neg (dia φ)) := nnf.box (fml.to_nnf (neg φ)) 55 | | (and φ ψ) := nnf.and (fml.to_nnf φ) (fml.to_nnf ψ) 56 | | (or φ ψ) := nnf.or (fml.to_nnf φ) (fml.to_nnf ψ) 57 | | (impl φ ψ) := nnf.or (fml.to_nnf (neg φ)) (fml.to_nnf ψ) 58 | | (box φ) := nnf.box (fml.to_nnf φ) 59 | | (dia φ) := nnf.dia (fml.to_nnf φ) 60 | using_well_founded {rel_tac := λ _ _, `[exact ⟨_, measure_wf fml_size⟩]} 61 | 62 | @[simp] def trans_size_left {st} (k : S4 st) : (Σ' (s) (φ : fml), fml_force k s φ) → ℕ := 63 | λ h, fml_size h.snd.fst 64 | 65 | @[simp] def trans_size_right {st} (k : S4 st) : (Σ' (s : st) (φ : fml), force k s (fml.to_nnf φ)) → ℕ := 66 | λ h, fml_size h.snd.fst 67 | 68 | theorem trans_left {st} (k : S4 st) : Π s φ, 69 | fml_force k s φ → force k s (fml.to_nnf φ) 70 | | s (var n) h := h 71 | | s (neg (var n)) h := h 72 | | s (neg (neg φ)) h := by dsimp at h; rw classical.not_not at h; exact trans_left _ _ h 73 | | s (neg (and φ ψ)) h := begin 74 | have := trans_left s (neg φ), 75 | dsimp at h, 76 | rw classical.not_and_distrib at h, 77 | cases h with l r, 78 | {simp [trans_left _ (neg φ) l]}, 79 | {simp [trans_left _ (neg ψ) r]} 80 | end 81 | | s (neg (or φ ψ)) h := begin 82 | dsimp at h, 83 | rw not_or_distrib at h, 84 | simp [and.intro (trans_left _ (neg φ) h.1) (trans_left _ (neg ψ) h.2)] 85 | end 86 | | s (neg (impl φ ψ)) h := begin 87 | dsimp at h, 88 | rw not_or_distrib at h, 89 | have : force k s (fml.to_nnf φ), 90 | { apply trans_left, have := h.1, 91 | rw classical.not_not at this, 92 | exact this }, 93 | simp [and.intro this (trans_left _ (neg ψ) h.2)] 94 | end 95 | | s (neg (box φ)) h := begin 96 | dsimp at h, 97 | rw classical.not_forall at h, 98 | cases h with w hw, 99 | rw classical.not_imp at hw, 100 | simp, split, split, 101 | {exact hw.1}, {apply trans_left, exact hw.2} 102 | end 103 | | s (neg (dia φ)) h := begin 104 | dsimp at h, simp, 105 | intros s' hs', 106 | rw not_exists at h, 107 | have := h s', 108 | rw not_and at this, 109 | have hnf := this hs', 110 | apply trans_left, exact hnf 111 | end 112 | | s (and φ ψ) h := by dsimp at h; simp [and.intro (trans_left _ _ h.1) (trans_left _ _ h.2)] 113 | | s (or φ ψ) h := begin 114 | dsimp at h, cases h, 115 | {exact or.inl (trans_left _ _ h)}, 116 | {exact or.inr (trans_left _ _ h)} 117 | end 118 | | s (impl φ ψ) h := begin 119 | dsimp at h, simp, 120 | cases h, 121 | {left, apply trans_left, exact h}, 122 | {right, apply trans_left, exact h} 123 | end 124 | | s (box φ) h := begin 125 | dsimp at h, simp, 126 | intros s' hs', 127 | exact trans_left _ _ (h s' hs') 128 | end 129 | | s (dia φ) h := begin 130 | dsimp at h, simp, 131 | cases h with w hw, 132 | split, split, {exact hw.1}, {exact trans_left _ _ hw.2} 133 | end 134 | using_well_founded {rel_tac := λ _ _, `[exact ⟨_, measure_wf (trans_size_left k)⟩]} 135 | 136 | theorem trans_right {st} (k : S4 st) : Π s φ, 137 | force k s (fml.to_nnf φ) → fml_force k s φ 138 | | s (var n) h := h 139 | | s (neg (var n)) h := h 140 | | s (neg (neg φ)) h := by dsimp; simp at h; exact not_not_intro (trans_right _ _ h) 141 | | s (neg (and φ ψ)) h := begin 142 | dsimp, simp at h, 143 | apply not_and_of_not_or_not, 144 | cases h, 145 | {left, exact trans_right _ _ h}, 146 | {right, exact trans_right _ _ h} 147 | end 148 | | s (neg (or φ ψ)) h := begin 149 | dsimp, simp at h, 150 | apply not_or, 151 | {exact trans_right _ _ h.1}, {exact trans_right _ _ h.2} 152 | end 153 | | s (neg (impl φ ψ)) h := begin 154 | dsimp, simp at h, 155 | apply not_or, 156 | {exact not_not_intro (trans_right _ _ h.1)}, {exact trans_right _ _ h.2} 157 | end 158 | | s (neg (box φ)) h := begin 159 | dsimp, simp at h, 160 | apply not_forall_of_exists_not, 161 | cases h with w hw, 162 | split, 163 | {apply not_imp_of_and_not, split, exact hw.1, exact trans_right _ _ hw.2} 164 | end 165 | | s (neg (dia φ)) h := begin 166 | dsimp, simp at h, 167 | rw not_exists, 168 | intros s' hs', 169 | have := h s' hs'.1, 170 | have hn := trans_right _ _ this, 171 | have := hs'.2, 172 | contradiction 173 | end 174 | | s (and φ ψ) h := begin 175 | dsimp, simp at h, 176 | split, {exact trans_right _ _ h.1}, {exact trans_right _ _ h.2} 177 | end 178 | | s (or φ ψ) h := begin 179 | dsimp, simp at h, 180 | cases h, {left, exact trans_right _ _ h}, {right, exact trans_right _ _ h} 181 | end 182 | | s (impl φ ψ) h := begin 183 | dsimp, simp at h, cases h, 184 | {left, exact trans_right _ _ h}, {right, exact trans_right _ _ h} 185 | end 186 | | s (box φ) h := begin 187 | dsimp, simp at h, 188 | intros s' hs', 189 | exact trans_right _ _ (h s' hs') 190 | end 191 | | s (dia φ) h := begin 192 | dsimp, simp at h, 193 | cases h with w hw, 194 | split, split, {exact hw.1}, {exact trans_right _ _ hw.2} 195 | end 196 | using_well_founded {rel_tac := λ _ _, `[exact ⟨_, measure_wf (trans_size_right k)⟩]} 197 | 198 | theorem trans_iff {st} (k : S4 st) (s φ) : fml_force k s φ ↔ force k s (fml.to_nnf φ) := 199 | ⟨trans_left k s φ, trans_right k s φ⟩ 200 | 201 | def fml_sat {st} (k : S4 st) (s) (Γ : list fml) : Prop := 202 | ∀ φ ∈ Γ, fml_force k s φ 203 | 204 | theorem fml_sat_of_empty {st} (k : S4 st) (s) : fml_sat k s [] := 205 | λ φ h, absurd h $ list.not_mem_nil _ 206 | 207 | def fml_unsatisfiable (Γ : list fml) : Prop := 208 | ∀ (st) (k : S4 st) s, ¬ fml_sat k s Γ 209 | 210 | theorem trans_sat_iff {st} (k : S4 st) (s) : Π Γ, 211 | fml_sat k s Γ ↔ sat k s (list.map fml.to_nnf Γ) 212 | | [] := by simp [sat_of_empty, fml_sat_of_empty] 213 | | (hd::tl) := 214 | begin 215 | split, 216 | { intro h, dsimp, intros φ hφ, cases hφ, 217 | { rw hφ, rw ←trans_iff, apply h, simp }, 218 | { have : fml_sat k s tl, 219 | { intros ψ hψ, apply h, simp [hψ] }, 220 | rw trans_sat_iff at this, 221 | exact this _ hφ } }, 222 | { intros h φ hφ, cases hφ, 223 | { rw hφ, rw trans_iff, apply h, simp }, 224 | { have : sat k s (list.map fml.to_nnf tl), 225 | { intros ψ hψ, apply h, simp [hψ] }, 226 | rw ←trans_sat_iff at this, 227 | exact this _ hφ } } 228 | end 229 | 230 | theorem trans_unsat_iff (Γ : list fml) : fml_unsatisfiable Γ ↔ unsatisfiable (list.map fml.to_nnf Γ) := 231 | begin 232 | split, 233 | {intros h _ _ _ _, apply h, rw trans_sat_iff, assumption}, 234 | {intros h _ _ _ _, apply h, rw ←trans_sat_iff, assumption} 235 | end 236 | -------------------------------------------------------------------------------- /src/S4/ops.lean: -------------------------------------------------------------------------------- 1 | import .size 2 | open nnf list 3 | 4 | namespace list 5 | universes u v w 6 | 7 | variables {α : Type u} {β : Type v} {γ : Type w} 8 | 9 | theorem mapp {p : β → Prop} (f : α → β) : Π (l : list α) (h : ∀ x∈l, p (f x)) x, x ∈ list.map f l → p x 10 | | [] h x := by simp 11 | | (hd::tl) h x := 12 | begin 13 | intro hmem, cases hmem, 14 | {rw hmem, apply h, simp}, 15 | {apply mapp tl, intros a ha, apply h, simp [ha], exact hmem} 16 | end 17 | 18 | theorem pmapp {p : β → Prop} {q : α → Prop} (f : Π a, q a → β) : Π (l : list α) 19 | (h : ∀ y (hy : q y) (h₁ : y ∈ l), p (f y hy)) 20 | (hf : ∀ y ∈ l, q y) x, x ∈ list.pmap f l hf → p x 21 | | [] h hf x := by simp 22 | | (hd::tl) h hf x := 23 | begin 24 | intro hmem, 25 | cases hmem, 26 | {rw hmem, apply h, simp}, 27 | {apply pmapp tl, intros a ha hamem, 28 | apply h, simp [ha], right, exact hamem, exact hmem} 29 | end 30 | 31 | @[simp] def ne_empty_head : Π l : list α, l ≠ [] → α 32 | | [] h := by contradiction 33 | | (a :: l) h := a 34 | 35 | theorem mem_pmap_of_mem {p : α → Prop} (f : Π a, p a → β) {a : α} {l : list α} (h : a ∈ l) (hf : ∀ y ∈ l, p y) : 36 | f a (hf a h) ∈ pmap f l hf := 37 | begin 38 | induction l with b l' ih, 39 | {cases h}, 40 | {rcases h with rfl | h, 41 | {exact or.inl rfl}, 42 | {apply or.inr, 43 | have hpy : ∀ (y : α), y ∈ l' → p y, 44 | {intros y hy, apply hf, right, exact hy}, 45 | have := ih h hpy, exact this}} 46 | end 47 | 48 | end list 49 | 50 | @[simp] def unbox : list nnf → list nnf 51 | | [] := [] 52 | | ((box φ) :: l) := φ :: unbox l 53 | | (e :: l) := unbox l 54 | 55 | theorem unbox_iff : Π {Γ φ}, box φ ∈ Γ ↔ φ ∈ unbox Γ 56 | | [] φ := begin split, repeat {intro h, simpa using h} end 57 | | (hd::tl) φ := 58 | begin 59 | split, 60 | { intro h, cases h₁ : hd, 61 | case nnf.box : ψ 62 | { dsimp [unbox], cases h, 63 | {left, rw h₁ at h, injection h}, 64 | {right, exact (@unbox_iff tl φ).1 h} }, 65 | all_goals 66 | { dsimp [unbox], cases h, 67 | {rw h₁ at h, contradiction}, 68 | {exact (@unbox_iff tl φ).1 h} } }, 69 | { intro h, cases h₁ : hd, 70 | case nnf.box : ψ 71 | { rw h₁ at h, dsimp [unbox] at h, cases h, 72 | {simp [h]}, {right, exact (@unbox_iff tl φ).2 h} }, 73 | all_goals 74 | { rw h₁ at h, dsimp [unbox] at h, right, exact (@unbox_iff tl φ).2 h } } 75 | end 76 | 77 | theorem unbox_size_aux : Π {Γ}, node_size (unbox Γ) ≤ node_size Γ 78 | | [] := by simp 79 | | (hd::tl) := 80 | begin 81 | cases h : hd, 82 | case nnf.box : ψ 83 | { dsimp, apply add_le_add, 84 | { dsimp [sizeof, has_sizeof.sizeof, nnf.sizeof], 85 | rw add_comm, apply nat.le_succ }, 86 | { apply unbox_size_aux } }, 87 | all_goals 88 | { dsimp, apply le_add_of_nonneg_of_le, 89 | { dsimp [sizeof, has_sizeof.sizeof, nnf.sizeof], rw add_comm, apply nat.zero_le }, 90 | { apply unbox_size_aux } } 91 | end 92 | 93 | @[simp] def rebox : list nnf → list nnf 94 | | [] := [] 95 | | (hd::tl) := box hd :: rebox tl 96 | 97 | theorem rebox_unbox_of_mem : Π {Γ} (h : ∀ {φ}, φ ∈ unbox Γ → box φ ∈ Γ), rebox (unbox Γ) ⊆ Γ 98 | | [] h := by simp 99 | | (hd::tl) h := 100 | begin 101 | cases hψ : hd, 102 | case nnf.box : φ {dsimp, simp [cons_subset_cons], apply subset_cons_of_subset, apply rebox_unbox_of_mem, simp [unbox_iff]}, 103 | all_goals {dsimp, apply subset_cons_of_subset, apply rebox_unbox_of_mem, simp [unbox_iff]} 104 | end 105 | 106 | theorem unbox_rebox : Π {Γ}, unbox (rebox Γ) = Γ 107 | | [] := by simp 108 | | (hd::tl) := by simp [unbox_rebox] 109 | 110 | -- Just that I don't want to say ∃ φ s.t. ... 111 | def box_only_rebox : Π {Γ}, box_only (rebox Γ) 112 | | [] := {no_var := by simp, 113 | no_neg := by simp, 114 | no_and := by simp, 115 | no_or := by simp, 116 | no_dia := by simp} 117 | | (hd::tl) := 118 | begin 119 | cases h : hd, 120 | all_goals { 121 | exact { no_var := begin 122 | intros n h, cases h, contradiction, 123 | apply (@box_only_rebox tl).no_var, assumption 124 | end, 125 | no_neg := begin 126 | intros n h, cases h, contradiction, 127 | apply (@box_only_rebox tl).no_neg, assumption 128 | end, 129 | no_and := begin 130 | intros φ ψ h, cases h, contradiction, 131 | apply (@box_only_rebox tl).no_and, assumption 132 | end, 133 | no_or := begin 134 | intros φ ψ h, cases h, contradiction, 135 | apply (@box_only_rebox tl).no_or, assumption 136 | end, 137 | no_dia := begin 138 | intros φ h, cases h, contradiction, 139 | apply (@box_only_rebox tl).no_dia, assumption 140 | end} } 141 | end 142 | 143 | theorem rebox_iff : Π {φ Γ}, box φ ∈ rebox Γ ↔ φ ∈ Γ 144 | | φ [] := by simp 145 | | φ (hd::tl) := 146 | begin 147 | split, 148 | {intro h, cases h₁ : hd, 149 | all_goals { cases h, 150 | {left, rw ←h₁, injection h}, 151 | {right, have := (@rebox_iff φ tl).1, exact this h } }}, 152 | {intro h, cases h₁ : hd, 153 | all_goals { dsimp, cases h, 154 | {left, rw ←h₁, rw h}, 155 | {right, have := (@rebox_iff φ tl).2, exact this h } } } 156 | end 157 | 158 | @[simp] def undia : list nnf → list nnf 159 | | [] := [] 160 | | ((dia φ) :: l) := φ :: undia l 161 | | (e :: l) := undia l 162 | 163 | theorem undia_iff : Π {Γ φ}, dia φ ∈ Γ ↔ φ ∈ undia Γ 164 | | [] φ := begin split, repeat {intro h, simpa using h} end 165 | | (hd::tl) φ := 166 | begin 167 | split, 168 | { intro h, cases h₁ : hd, 169 | case nnf.dia : ψ 170 | { dsimp [undia], cases h, 171 | {left, rw h₁ at h, injection h}, 172 | {right, exact (@undia_iff tl φ).1 h} }, 173 | all_goals 174 | { dsimp [undia], cases h, 175 | {rw h₁ at h, contradiction}, 176 | {exact (@undia_iff tl φ).1 h} } }, 177 | { intro h, cases h₁ : hd, 178 | case nnf.dia : ψ 179 | { rw h₁ at h, dsimp [undia] at h, cases h, 180 | {simp [h]}, {right, exact (@undia_iff tl φ).2 h} }, 181 | all_goals 182 | { rw h₁ at h, dsimp [undia] at h, right, exact (@undia_iff tl φ).2 h } } 183 | end 184 | 185 | def get_contra : Π Γ : list nnf, 186 | psum {p : nat // var p ∈ Γ ∧ neg p ∈ Γ} 187 | (∀ n, var n ∈ Γ → neg n ∉ Γ) 188 | | [] := psum.inr $ λ _ h, absurd h $ not_mem_nil _ 189 | | (hd :: tl) := 190 | begin 191 | cases h : hd, 192 | case nnf.var : n 193 | {apply dite (neg n ∈ tl), 194 | {intro t, 195 | exact psum.inl ⟨n, ⟨mem_cons_self _ _, mem_cons_of_mem _ t⟩⟩}, 196 | {intro e, 197 | cases (get_contra tl), 198 | {left, constructor, constructor, 199 | apply mem_cons_of_mem, exact val.2.1, 200 | apply mem_cons_of_mem, exact val.2.2}, 201 | {right, 202 | intros m hm hin, 203 | by_cases eq : m=n, 204 | {apply e, cases hin, contradiction, rw ←eq, assumption}, 205 | {cases hm, apply eq, injection hm, apply val, exact hm, 206 | cases hin, contradiction, assumption} } } 207 | }, 208 | case nnf.neg : n 209 | { apply dite (var n ∈ tl), 210 | { intro t, 211 | exact psum.inl ⟨n, ⟨mem_cons_of_mem _ t, mem_cons_self _ _⟩⟩ }, 212 | { intro e, 213 | cases (get_contra tl), 214 | {left, constructor, constructor, 215 | apply mem_cons_of_mem, exact val.2.1, 216 | apply mem_cons_of_mem, exact val.2.2 }, 217 | { right, 218 | intros m hm hin, 219 | by_cases eq : m=n, 220 | { apply e, cases hm, contradiction, rw ←eq, assumption }, 221 | { cases hin, apply eq, injection hin, apply val, 222 | swap, exact hin, cases hm, contradiction, assumption } 223 | } 224 | } 225 | }, 226 | all_goals 227 | { 228 | cases (get_contra tl), 229 | { left, constructor, constructor, 230 | apply mem_cons_of_mem, exact val.2.1, 231 | apply mem_cons_of_mem, exact val.2.2 }, 232 | { right, 233 | intros m hm hin, 234 | {apply val, swap 3, exact m, 235 | cases hm, contradiction, assumption, 236 | cases hin, contradiction, assumption} } 237 | } 238 | end 239 | 240 | def get_contra_seqt : Π Γ : sseqt, 241 | psum {p : nat // var p ∈ Γ.m ∧ neg p ∈ Γ.m} 242 | (∀ n, var n ∈ Γ.m → neg n ∉ Γ.m) 243 | := λ Γ, get_contra Γ.m 244 | 245 | def get_and : Π Γ : list nnf, 246 | psum {p : nnf × nnf // and p.1 p.2 ∈ Γ} 247 | (∀ φ ψ, nnf.and φ ψ ∉ Γ) 248 | | [] := psum.inr $ λ _ _, not_mem_nil _ 249 | | (hd :: tl) := 250 | begin 251 | cases h : hd, 252 | case nnf.and : φ ψ { left, constructor,swap, 253 | constructor, exact φ, exact ψ, simp 254 | }, 255 | all_goals 256 | { cases (get_and tl), 257 | {left, 258 | constructor, 259 | apply mem_cons_of_mem, 260 | exact val.2}, 261 | {right, intros γ ψ h, 262 | cases h, contradiction, 263 | apply val, assumption } 264 | } 265 | end 266 | 267 | def get_and_seqt : Π Γ : sseqt, 268 | psum {p : nnf × nnf // and p.1 p.2 ∈ Γ.m} 269 | (∀ φ ψ, nnf.and φ ψ ∉ Γ.m) 270 | := λ Γ, get_and Γ.m 271 | 272 | def get_or : Π Γ : list nnf, 273 | psum {p : nnf × nnf // or p.1 p.2 ∈ Γ} 274 | (∀ φ ψ, nnf.or φ ψ ∉ Γ) 275 | | [] := psum.inr $ λ _ _, not_mem_nil _ 276 | | (hd :: tl) := 277 | begin 278 | cases h : hd, 279 | case nnf.or : φ ψ { left, constructor,swap, 280 | constructor, exact φ, exact ψ, simp }, 281 | all_goals 282 | { cases (get_or tl), 283 | {left, 284 | constructor, 285 | apply mem_cons_of_mem, 286 | exact val.2}, 287 | {right, intros γ ψ h, 288 | cases h, contradiction, 289 | apply val, assumption} 290 | } 291 | end 292 | 293 | def get_or_seqt : Π Γ : sseqt, 294 | psum {p : nnf × nnf // or p.1 p.2 ∈ Γ.m} 295 | (∀ φ ψ, nnf.or φ ψ ∉ Γ.m) 296 | := λ Γ, get_or Γ.m 297 | 298 | def get_dia : Π Γ : list nnf, 299 | psum {p : nnf // dia p ∈ Γ} 300 | (∀ φ, nnf.dia φ ∉ Γ) 301 | | [] := psum.inr $ λ _, not_mem_nil _ 302 | | (hd :: tl) := 303 | begin 304 | cases h : hd, 305 | case nnf.dia : φ { left, constructor, swap, exact φ, simp }, 306 | all_goals 307 | { cases (get_dia tl), 308 | {left, 309 | constructor, 310 | apply mem_cons_of_mem, 311 | exact val.2}, 312 | {right, intros γ h, 313 | cases h, contradiction, 314 | apply val, assumption } } 315 | end 316 | 317 | def get_dia_seqt : Π Γ : sseqt, 318 | psum {p : nnf // dia p ∈ Γ.m} 319 | (∀ φ, nnf.dia φ ∉ Γ.m) 320 | := λ Γ, get_dia Γ.m 321 | 322 | def get_box : Π Γ : list nnf, 323 | psum {l : nnf // box l ∈ Γ} 324 | (∀ φ, nnf.box φ ∉ Γ) 325 | | [] := psum.inr $ λ _, not_mem_nil _ 326 | | (hd :: tl) := 327 | begin 328 | cases h : hd, 329 | case nnf.box : φ { left, constructor, swap, exact φ, simp }, 330 | all_goals 331 | { cases (get_box tl), 332 | {left, 333 | constructor, 334 | apply mem_cons_of_mem, 335 | exact val.2}, 336 | {right, intros γ h, 337 | cases h, contradiction, 338 | apply val, assumption } } 339 | end 340 | 341 | def get_box_seqt : Π Γ : sseqt, 342 | psum {p : nnf // box p ∈ Γ.m} 343 | (∀ φ, nnf.box φ ∉ Γ.m) 344 | := λ Γ, get_box Γ.m 345 | 346 | @[simp] def get_var : list nnf → list ℕ 347 | | [] := [] 348 | | ((var n) :: l) := n :: get_var l 349 | | (e :: l) := get_var l 350 | 351 | theorem get_var_iff : Π {Γ n}, var n ∈ Γ ↔ n ∈ get_var Γ 352 | | [] φ := begin split, repeat {intro h, simpa using h} end 353 | | (hd::tl) φ := 354 | begin 355 | split, 356 | { intro h, cases h₁ : hd, 357 | case nnf.var : n 358 | { dsimp, cases h, 359 | {left, rw h₁ at h, injection h}, 360 | {right, exact (@get_var_iff tl φ).1 h} }, 361 | all_goals 362 | { dsimp, cases h, 363 | {rw h₁ at h, contradiction}, 364 | {exact (@get_var_iff tl φ).1 h} } }, 365 | { intro h, cases h₁ : hd, 366 | case nnf.var : n 367 | { rw h₁ at h, dsimp at h, cases h, 368 | {simp [h]}, {right, exact (@get_var_iff tl φ).2 h} }, 369 | all_goals 370 | { rw h₁ at h, dsimp [undia] at h, right, exact (@get_var_iff tl φ).2 h } } 371 | end 372 | -------------------------------------------------------------------------------- /src/S4/semantics.lean: -------------------------------------------------------------------------------- 1 | import .ops 2 | open subtype nnf tmodel list 3 | 4 | section 5 | variables (φ ψ : nnf) (Γ₁ Γ₂ Δ Λ: list nnf) {st : Type} 6 | variables (k : S4 st) (s : st) 7 | open list 8 | 9 | theorem sat_subset (h₁ : Γ₁ ⊆ Γ₂) (h₂ : sat k s Γ₂) : sat k s Γ₁ := 10 | λ x hx, h₂ _ (h₁ hx) 11 | 12 | theorem sat_sublist (h₁ : Γ₁ <+ Γ₂) (h₂ :sat k s Γ₂) : sat k s Γ₁ := 13 | sat_subset _ _ _ _ (sublist.subset h₁) h₂ 14 | 15 | theorem sat_append (h₁ : sat k s Γ₁) (h₂ : sat k s Γ₂) : sat k s (Γ₁ ++ Γ₂) := 16 | begin 17 | intros φ h, rw mem_append at h, cases h, 18 | apply h₁ _ h, apply h₂ _ h 19 | end 20 | 21 | theorem unsat_contra {Δ n} : var n ∈ Δ → neg n ∈ Δ → unsatisfiable Δ:= 22 | begin 23 | intros h₁ h₂, intros v hsat, intros s hsat, 24 | have := hsat _ h₁, have := hsat _ h₂, simpa 25 | end 26 | 27 | theorem unsat_contra_seqt {Δ : sseqt} {n} : var n ∈ Δ.m → neg n ∈ Δ.m → unsatisfiable (Δ.m ++ Δ.b):= 28 | begin 29 | intros h₁ h₂, intros st m, intros s hsat, 30 | have := unsat_contra h₁ h₂, 31 | have := this _ m s, 32 | apply this, 33 | apply sat_subset _ _ _ _ _ hsat, 34 | simp 35 | end 36 | 37 | theorem sat_of_and : force k s (and φ ψ) ↔ (force k s φ) ∧ (force k s ψ) := 38 | by split; {intro, simpa} 39 | 40 | theorem sat_of_sat_erase (h₁ : sat k s $ Δ.erase φ) (h₂ : force k s φ) : sat k s Δ := 41 | begin 42 | intro ψ, intro h, 43 | by_cases (ψ = φ), 44 | {rw h, assumption}, 45 | {have : ψ ∈ Δ.erase φ, 46 | rw mem_erase_of_ne, assumption, exact h, 47 | apply h₁, assumption} 48 | end 49 | 50 | theorem unsat_and_of_unsat_split 51 | (h₁ : and φ ψ ∈ Δ) 52 | (h₂ : unsatisfiable $ φ :: ψ :: Δ.erase (and φ ψ)) : 53 | unsatisfiable Δ := 54 | begin 55 | intro st, intros, intro h, 56 | apply h₂, swap 3, exact k, swap, exact s, 57 | intro e, intro he, 58 | cases he, 59 | {rw he, have := h _ h₁, rw sat_of_and at this, exact this.1}, 60 | {cases he, 61 | {rw he, have := h _ h₁, rw sat_of_and at this, exact this.2}, 62 | {have := h _ h₁, apply h, apply mem_of_mem_erase he} } 63 | end 64 | 65 | theorem unsat_and_of_unsat_split_seqt {Γ} 66 | (h₁ : and φ ψ ∈ Δ) 67 | (h₂ : unsatisfiable $ (φ :: ψ :: Δ.erase (and φ ψ)++Γ)) : 68 | unsatisfiable (Δ++Γ) := 69 | begin 70 | intro st, intros, intro h, 71 | apply h₂, swap 3, exact k, swap, exact s, 72 | intro e, intro he, 73 | cases he, 74 | {rw he, have := h _ (mem_append_left _ h₁), rw sat_of_and at this, exact this.1}, 75 | {cases he, 76 | {rw he, have := h _ (mem_append_left _ h₁), rw sat_of_and at this, exact this.2}, 77 | {have := h _ (mem_append_left _ h₁), apply h, apply mem_of_mem_erase, rw erase_append_left, exact he, exact h₁} } 78 | end 79 | 80 | theorem sat_and_of_sat_split 81 | (h₁ : and φ ψ ∈ Δ) 82 | (h₂ : sat k s $ φ :: ψ :: Δ.erase (and φ ψ)) : 83 | sat k s Δ := 84 | begin 85 | intro e, intro he, 86 | by_cases (e = and φ ψ), 87 | { rw h, dsimp, split, repeat {apply h₂, simp} }, 88 | { have : e ∈ Δ.erase (and φ ψ), 89 | { rw mem_erase_of_ne, repeat { assumption } }, 90 | apply h₂, simp [this] } 91 | end 92 | 93 | theorem sat_and_of_sat_split_seqt {Γ} 94 | (h₁ : and φ ψ ∈ Δ) 95 | (h₂ : sat k s $ (φ :: ψ :: Δ.erase (and φ ψ)++Γ)) : 96 | sat k s (Δ++Γ) := 97 | begin 98 | intro e, intro he, 99 | by_cases (e = and φ ψ), 100 | { rw h, dsimp, split, repeat {apply h₂, simp} }, 101 | { have : e ∈ Δ.erase (and φ ψ) ++ Γ, 102 | { rw ←erase_append_left, rw mem_erase_of_ne, repeat {assumption} }, 103 | apply h₂, simp [this] } 104 | end 105 | 106 | theorem sat_split_of_sat_and_seqt {Γ} 107 | (h₁ : and φ ψ ∈ Δ) 108 | (h₂ : sat k s (Δ++Γ)) : 109 | sat k s $ (φ :: ψ :: Δ.erase (and φ ψ)++Γ) := 110 | begin 111 | intros e he, rw mem_append at he, cases he, 112 | have : force k s (and φ ψ), {apply h₂, simp [h₁]}, rw sat_of_and at this, 113 | {cases he, 114 | {rw he, exact this.left}, 115 | {cases he, rw he, exact this.right, apply h₂, rw mem_append, left, apply mem_of_mem_erase he} 116 | }, 117 | {apply h₂, rw mem_append, right, exact he} 118 | end 119 | 120 | theorem unsat_or_of_unsat_split_seqt {Γ} 121 | (h : or φ ψ ∈ Δ) 122 | (h₁ : unsatisfiable $ (φ :: Δ.erase (nnf.or φ ψ)++Γ)) 123 | (h₂ : unsatisfiable $ (ψ :: Δ.erase (nnf.or φ ψ)++Γ)) : 124 | unsatisfiable $ (Δ++Γ) := 125 | begin 126 | intro, intros, intro hsat, 127 | have := hsat _ (mem_append_left _ h), 128 | dsimp at this, 129 | cases this, 130 | {apply h₁, swap 3, exact k, swap, exact s, intro e, intro he, 131 | cases he, rw he, exact this, apply hsat, 132 | apply mem_of_mem_erase, rw erase_append_left, exact he, exact h}, 133 | {apply h₂, swap 3, exact k, swap, exact s, intro e, intro he, 134 | cases he, rw he, exact this, apply hsat, apply mem_of_mem_erase, rw erase_append_left, exact he, exact h} 135 | end 136 | 137 | theorem sat_or_of_sat_split_left 138 | (h : or φ ψ ∈ Δ) 139 | (hl : sat k s $ φ :: Δ.erase (nnf.or φ ψ)) : 140 | sat k s Δ := 141 | begin 142 | intros e he, 143 | by_cases (e = or φ ψ), 144 | { rw h, dsimp, left, apply hl, simp}, 145 | {have : e ∈ Δ.erase (or φ ψ), 146 | { rw mem_erase_of_ne, repeat { assumption } }, 147 | apply hl, simp [this]} 148 | end 149 | 150 | theorem sat_or_of_sat_split_right 151 | (h : or φ ψ ∈ Δ) 152 | (hl : sat k s $ ψ :: Δ.erase (nnf.or φ ψ)) : 153 | sat k s Δ := 154 | begin 155 | intros e he, 156 | by_cases (e = or φ ψ), 157 | { rw h, dsimp, right, apply hl, simp}, 158 | { have : e ∈ Δ.erase (or φ ψ), 159 | { rw mem_erase_of_ne, repeat { assumption } }, 160 | apply hl, simp [this] } 161 | end 162 | 163 | /- S4-specific lemmas -/ 164 | 165 | theorem force_of_force_box (h : force k s $ box φ) : force k s φ 166 | := by apply h; apply k.refl 167 | 168 | theorem force_box_box_of_force_box : force k s (box φ) → force k s (box (box φ)) := 169 | by intros h s₁ rs₁ s₂ rs₂; apply h; apply k.trans rs₁ rs₂ 170 | 171 | theorem unsat_of_unsat_box_new 172 | (h₁ : box φ ∈ Δ) 173 | (h₂ : unsatisfiable $ (φ :: Δ.erase (box φ)) ++ box φ :: Λ) : 174 | unsatisfiable (Δ ++ Λ) := 175 | begin 176 | intros st k s h, 177 | apply h₂, swap 3, exact k, swap, exact s, 178 | intros e he, 179 | rw [mem_append] at he, 180 | cases he, 181 | {cases he, 182 | {rw he, apply force_of_force_box, apply h (box φ), simp [h₁]}, 183 | {apply h, rw mem_append, left, apply mem_of_mem_erase he }}, 184 | {cases he, 185 | {rw ←he at h₁, apply h, rw mem_append, left, exact h₁}, 186 | {apply h, rw mem_append, right, assumption}} 187 | end 188 | 189 | theorem sat_copy_of_sat_box_new 190 | (h₁ : box φ ∈ Δ) 191 | (h₂ : sat k s $ (φ :: Δ.erase (box φ)) ++ box φ :: Λ) : 192 | sat k s (Δ ++ Λ) := 193 | begin 194 | intros ψ hφ, 195 | rw mem_append at hφ, 196 | cases hφ, 197 | {by_cases heq : ψ = box φ, 198 | {rw heq, apply h₂ (box φ), simp}, 199 | {have := mem_erase_of_ne heq, rw ←this at hφ, apply h₂, simp, right, left, exact hφ}}, 200 | {apply h₂, simp, repeat {right}, exact hφ} 201 | end 202 | 203 | theorem unsat_of_unsat_box_dup 204 | (h₁ : box φ ∈ Δ) 205 | (h₃ : unsatisfiable $ (φ :: Δ.erase (box φ)) ++ Λ) : 206 | unsatisfiable (Δ ++ Λ) := 207 | begin 208 | intros st k s h, 209 | apply h₃, swap 3, exact k, swap, exact s, 210 | intros e he, 211 | cases he, 212 | {rw he, apply force_of_force_box, apply h (box φ), simp [h₁]}, 213 | {have := mem_append.1 he, cases this, 214 | {apply h, apply mem_append_left, apply mem_of_mem_erase this}, 215 | {apply h, apply mem_append_right, exact this}} 216 | end 217 | 218 | end 219 | 220 | def unmodal_seqt (Γ : sseqt) : list sseqt := 221 | @list.pmap _ _ (λ φ, φ ∉ Γ.h ∧ dia φ ∈ Γ.m) 222 | (λ d h, 223 | ({s := some {d := d, b := Γ.b}, 224 | a := {d := d, b := Γ.b} :: Γ.a, 225 | h := d :: Γ.h, 226 | m := d :: Γ.b, 227 | ndh := begin rw list.nodup_cons, split, exact h.1, exact Γ.ndh end, 228 | sph := begin apply list.cons_subperm_of_mem, apply Γ.ndh, exact h.1, apply mem_closure_dia, apply Γ.sbm, exact h.2, exact Γ.sph end, 229 | sbm := begin rw list.cons_subset, split, apply mem_closure_dia, apply Γ.sbm, exact h.2, apply list.subperm.subset, exact Γ.spb end, 230 | ha := begin intros φ h, cases h with l r, rw l, simp, right, apply Γ.ha, exact r end, 231 | ps₁ := begin simp [dsig] end, 232 | ps₂ := begin simp [bsig] end, 233 | .. Γ} : sseqt)) 234 | (filter_undia Γ.h Γ.m) 235 | (begin 236 | intros φ hmem,split, 237 | {apply mem_filter_dia_right, exact hmem}, 238 | {apply mem_filter_dia_right_aux, exact hmem} 239 | end) 240 | 241 | def unmodal_seqt_size (Γ : sseqt) : ∀ (i : sseqt), i ∈ unmodal_seqt Γ → (prod.measure_lex' sseqt_size i Γ) := 242 | list.pmapp _ _ 243 | begin 244 | intros φ h hmem, 245 | right, left, 246 | apply length_sub_lt_of_nodup_subperm, 247 | {apply Γ.sph}, 248 | {apply mem_closure_dia, apply Γ.sbm, exact h.2}, 249 | {exact h.1}, 250 | {exact Γ.ndh} 251 | end 252 | _ 253 | 254 | def sat_unmodal_of_sat {Γ : sseqt} : ∀ (i : sseqt), i ∈ unmodal_seqt Γ → 255 | (∀ {st : Type} (k : S4 st) s, 256 | sat k s (Γ.m ++ Γ.b) → ∃ s', sat k s' (i.m ++ i.b)) := 257 | list.pmapp _ _ 258 | begin 259 | intros φ hninh hmem st k s h, 260 | have hd : force k s (dia φ), 261 | { apply h, rw mem_append, left, exact hninh.2 }, 262 | have hb : ∀ φ ∈ Γ.b, force k s (box φ), 263 | { intros γ hγ, 264 | have := box_only_ex Γ.hb hγ, cases this with w hw, 265 | rw hw, apply force_box_box_of_force_box, 266 | have := h _ (mem_append_right _ hγ), rw hw at this, exact this}, 267 | rcases hd with ⟨w, hrw, hfw⟩, 268 | split, swap, exact w, 269 | intros ψ hψ, simp at hψ, cases hψ, 270 | {rw hψ, exact hfw}, 271 | {apply hb, exact hψ, exact hrw} 272 | end 273 | _ 274 | 275 | def unsat_of_unsat_unmodal {Γ : sseqt} (i : sseqt) : i ∈ unmodal_seqt Γ ∧ unsatisfiable (i.m ++ i.b) → unsatisfiable (Γ.m ++ Γ.b) := 276 | begin 277 | intro hex, intros st k s h, 278 | have := sat_unmodal_of_sat i hex.1 k s h, 279 | cases this with w hw, 280 | have := hex.2, 281 | exact this _ _ _ hw 282 | end 283 | 284 | def unmodal_mem_box (Γ : sseqt) : ∀ (i : sseqt), i ∈ unmodal_seqt Γ → (∀ φ, box φ ∈ Γ.b → box φ ∈ i.m) := 285 | list.pmapp _ _ begin intros φ h hmem ψ hψ, right, exact hψ end _ 286 | 287 | def mem_unmodal_seqt (Γ : sseqt) (φ) (h : φ ∉ Γ.h ∧ dia φ ∈ Γ.m) : ∃ (i : sseqt), i ∈ unmodal_seqt Γ ∧ φ ∈ i.m := 288 | begin 289 | split, swap, 290 | {exact 291 | ({s := some {d := φ, b := Γ.b}, 292 | a := {d := φ, b := Γ.b} :: Γ.a, 293 | h := φ :: Γ.h, 294 | m := φ :: Γ.b, 295 | ndh := begin rw list.nodup_cons, split, exact h.1, exact Γ.ndh end, 296 | sph := begin apply list.cons_subperm_of_mem, apply Γ.ndh, exact h.1, apply mem_closure_dia, apply Γ.sbm, exact h.2, exact Γ.sph end, 297 | sbm := begin rw list.cons_subset, split, apply mem_closure_dia, apply Γ.sbm, exact h.2, apply list.subperm.subset, exact Γ.spb end, 298 | ha := begin intros φ h, cases h with l r, rw l, simp, right, apply Γ.ha, exact r end, 299 | ps₁ := begin intro, simp [dsig] end, 300 | ps₂ := begin intro, simp [bsig] end, 301 | .. Γ} : sseqt)}, 302 | { dsimp [unmodal_seqt], split, 303 | {let mf := (λ (d : nnf) (h : d ∉ Γ.h ∧ dia d ∈ Γ.m), 304 | ({s := some {d := d, b := Γ.b}, 305 | a := {d := d, b := Γ.b} :: Γ.a, 306 | h := d :: Γ.h, 307 | m := d :: Γ.b, 308 | ndh := begin rw list.nodup_cons, split, exact h.1, exact Γ.ndh end, 309 | sph := begin apply list.cons_subperm_of_mem, apply Γ.ndh, exact h.1, apply mem_closure_dia, apply Γ.sbm, exact h.2, exact Γ.sph end, 310 | sbm := begin rw list.cons_subset, split, apply mem_closure_dia, apply Γ.sbm, exact h.2, apply list.subperm.subset, exact Γ.spb end, 311 | ha := begin intros φ h, cases h with l r, rw l, simp, right, apply Γ.ha, exact r end, 312 | ps₁ := begin simp [dsig] end, 313 | ps₂ := begin simp [bsig] end, 314 | .. Γ} :sseqt)), 315 | have hmem := mem_filter_undia_left _ _ _ h.2 h.1, 316 | have hf : ∀ (y : nnf), y ∈ filter_undia (Γ.h) (Γ.m) → y ∉ Γ.h ∧ dia y ∈ Γ.m, 317 | {intros h hy, split, {apply mem_filter_dia_right, exact hy}, {apply mem_filter_dia_right_aux, exact hy}}, 318 | exact mem_pmap_of_mem mf hmem hf}, 319 | { simp } } 320 | end 321 | 322 | theorem unmodal_sig (Γ : sseqt) : ∀ (i : sseqt), i ∈ unmodal_seqt Γ → (∀ a, a ∈ i.a → some a = i.s ∨ a ∈ Γ.a) := 323 | list.pmapp _ _ 324 | begin 325 | intros φ h hmem a ha, simp at ha, 326 | cases ha, 327 | {left, simp, exact ha}, 328 | {right, exact ha} 329 | end 330 | _ 331 | 332 | 333 | theorem unsat_of_closed_and {Γ Δ} (i : and_instance Γ Δ) (h : unsatisfiable Δ) : unsatisfiable Γ := 334 | by cases i; { apply unsat_and_of_unsat_split, repeat {assumption} } 335 | 336 | theorem unsat_of_closed_and_seqt {Γ Δ} (i : and_instance_seqt Γ Δ) (h : unsatisfiable (Δ.m++Δ.b)) : unsatisfiable (Γ.m++Γ.b) := 337 | by cases i; {apply unsat_and_of_unsat_split_seqt, repeat {assumption} } 338 | 339 | theorem unsat_of_closed_or_seqt {Γ₁ Γ₂ Δ : sseqt} (i : or_instance_seqt Δ Γ₁ Γ₂) 340 | (h₁ : unsatisfiable (Γ₁.m++Γ₁.b)) 341 | (h₂ : unsatisfiable (Γ₂.m++Γ₂.b)) : 342 | unsatisfiable (Δ.m++Δ.b) := 343 | by cases i; {apply unsat_or_of_unsat_split_seqt, repeat { assumption }} 344 | 345 | theorem unsat_of_closed_box_new {Γ Δ} (i : box_new_instance_seqt Γ Δ) (h : unsatisfiable $ (Δ.m++Δ.b)) : unsatisfiable (Γ.m++Γ.b) := 346 | by cases i; { apply unsat_of_unsat_box_new, repeat { assumption } } 347 | 348 | theorem unsat_of_closed_box_dup {Γ Δ} (i : box_dup_instance_seqt Γ Δ) (h : unsatisfiable $ (Δ.m++Δ.b)) : unsatisfiable (Γ.m++Γ.b) := 349 | by cases i; { apply unsat_of_unsat_box_dup, repeat { assumption } } 350 | -------------------------------------------------------------------------------- /src/S4/size.lean: -------------------------------------------------------------------------------- 1 | import .S4_defs .data 2 | open nnf prod list 3 | 4 | namespace prod 5 | 6 | universe u 7 | 8 | def measure_lex {α : Sort u} : (α → ℕ × ℕ) → α → α → Prop := 9 | inv_image $ lex (<) (<) 10 | 11 | def measure_lex' {α : Sort u} : (α → ℕ × ℕ × ℕ) → α → α → Prop := 12 | inv_image $ lex (<) (measure_lex (λ x, x)) 13 | 14 | def measure_lex_wf {α : Sort u} (f : α → ℕ × ℕ) : well_founded (measure_lex f) := 15 | inv_image.wf f (lex_wf nat.lt_wf nat.lt_wf) 16 | 17 | def measure_lex_wf' {α : Sort u} (f : α → ℕ × ℕ × ℕ) : well_founded (measure_lex' f) := 18 | inv_image.wf f (lex_wf nat.lt_wf (measure_lex_wf (λ x, x))) 19 | 20 | end prod 21 | 22 | def sseqt_size (s : sseqt) : ℕ × ℕ × ℕ := 23 | ((closure s.goal).length - s.b.length, 24 | (closure s.goal).length - s.h.length, 25 | node_size s.m) 26 | 27 | instance : has_well_founded sseqt := 28 | ⟨prod.measure_lex' sseqt_size, prod.measure_lex_wf' sseqt_size⟩ 29 | 30 | lemma size_erase_add_eq_size_sum {φ} : Π (Γ : list nnf) (h : φ ∈ Γ), node_size (Γ.erase φ) + sizeof φ = node_size Γ 31 | | [] h := absurd h $ not_mem_nil _ 32 | | (hd :: tl) h := 33 | begin 34 | by_cases eq : φ = hd, 35 | {dsimp [list.erase], rw if_pos, rw eq, apply add_comm, rw eq}, 36 | {dsimp [list.erase], rw if_neg, dsimp, rw [add_assoc], have : node_size (list.erase tl φ) + sizeof φ = node_size tl, 37 | {apply size_erase_add_eq_size_sum, cases h, contradiction, assumption}, 38 | rw this, intro h', rw h' at eq, contradiction} 39 | end 40 | 41 | theorem split_lt_and_size {φ ψ} (Γ : list nnf) (h : nnf.and φ ψ ∈ Γ) : 42 | node_size (φ :: ψ :: Γ.erase (nnf.and φ ψ)) < node_size Γ := 43 | begin 44 | dsimp [node_size], 45 | rw ←size_erase_add_eq_size_sum Γ, 46 | swap, exact h, 47 | rw [←add_assoc, add_comm], 48 | apply add_lt_add_left, dsimp [sizeof, has_sizeof.sizeof, nnf.sizeof], 49 | rw [add_assoc], 50 | rw [nat.one_add], 51 | apply nat.lt_succ_self 52 | end 53 | 54 | theorem split_lt_and_seqt {φ ψ} (Γ : sseqt) (h : nnf.and φ ψ ∈ Γ.m) : 55 | prod.measure_lex' sseqt_size 56 | (and_child Γ h) Γ := 57 | begin 58 | right, right, 59 | simp only [and_child], 60 | apply split_lt_and_size _ h 61 | end 62 | 63 | theorem split_lt_or_size_left {φ ψ} (Γ : list nnf) (h : nnf.or φ ψ ∈ Γ) : 64 | node_size (φ :: Γ.erase (nnf.or φ ψ)) < node_size Γ := 65 | begin 66 | dsimp, 67 | rw ←size_erase_add_eq_size_sum Γ, 68 | swap, exact h, 69 | rw [add_comm], 70 | apply add_lt_add_left, dsimp [sizeof, has_sizeof.sizeof, nnf.sizeof], 71 | rw [add_comm, ←add_assoc], 72 | apply nat.lt_add_of_pos_left, 73 | apply nat.succ_pos 74 | end 75 | 76 | theorem split_lt_or_size_right {φ ψ} (Γ : list nnf) (h : nnf.or φ ψ ∈ Γ) : 77 | node_size (ψ :: Γ.erase (nnf.or φ ψ)) < node_size Γ := 78 | begin 79 | dsimp, 80 | rw ←size_erase_add_eq_size_sum Γ, 81 | swap, exact h, 82 | rw [add_comm], 83 | apply add_lt_add_left, dsimp [sizeof, has_sizeof.sizeof, nnf.sizeof], 84 | apply nat.lt_add_of_pos_left, 85 | rw nat.one_add, 86 | apply nat.succ_pos 87 | end 88 | 89 | theorem split_lt_or_seqt_left {φ ψ} (Γ : sseqt) (h : nnf.or φ ψ ∈ Γ.m) : 90 | prod.measure_lex' sseqt_size (or_child_left Γ h) Γ := 91 | begin 92 | right, right, 93 | simp only [or_child_left], 94 | apply split_lt_or_size_left _ h 95 | end 96 | 97 | theorem split_lt_or_seqt_right {φ ψ} (Γ : sseqt) (h : nnf.or φ ψ ∈ Γ.m) : 98 | prod.measure_lex' sseqt_size (or_child_right Γ h) Γ := 99 | begin 100 | right, right, 101 | simp only [or_child_right], 102 | apply split_lt_or_size_right _ h 103 | end 104 | 105 | theorem copy_lt_size {φ} (Γ : list nnf) (h : nnf.box φ ∈ Γ) : 106 | node_size (φ :: Γ.erase (nnf.box φ)) < node_size Γ := 107 | begin 108 | dsimp [node_size], 109 | rw ←size_erase_add_eq_size_sum Γ, 110 | swap, exact h, 111 | rw [add_comm], simp [sizeof,has_sizeof.sizeof, nnf.sizeof] 112 | end 113 | 114 | theorem copy_lt_seqt {φ} (Γ : sseqt) (h : nnf.box φ ∈ Γ.m) : 115 | prod.measure_lex' sseqt_size (box_child Γ h) Γ := 116 | begin 117 | right, right, 118 | simp only [box_child], 119 | apply copy_lt_size _ h 120 | end 121 | 122 | theorem box_new_lt_seqt {φ} (Γ : sseqt) 123 | (h₁ : nnf.box φ ∈ Γ.m) (h₂ : nnf.box φ ∉ Γ.b) : 124 | prod.measure_lex' sseqt_size (box_child_new Γ h₁ h₂) Γ := 125 | begin 126 | left, 127 | apply length_sub_lt_of_nodup_subperm, 128 | {dsimp [box_child_new], exact Γ.spb}, 129 | {dsimp [box_child_new], apply Γ.sbm, exact h₁}, 130 | {exact h₂}, 131 | {exact Γ.ndb} 132 | end 133 | -------------------------------------------------------------------------------- /src/S4/tableau.lean: -------------------------------------------------------------------------------- 1 | import .full_language .vanilla 2 | 3 | def fml_is_sat (Γ : list fml) : bool := 4 | let l : list nnf := list.map fml.to_nnf Γ in 5 | match tableau (mk_sseqt l) with 6 | | node.closed _ := ff 7 | | node.open_ _ := tt 8 | end 9 | 10 | theorem classical_correctness (Γ : list fml) : 11 | fml_is_sat Γ = tt ↔ ∃ (st : Type) (k : S4 st) s, fml_sat k s Γ := 12 | begin 13 | cases h : fml_is_sat Γ, 14 | constructor, 15 | { intro, contradiction }, 16 | { intro hsat, cases eq : tableau (mk_sseqt (list.map fml.to_nnf Γ)), 17 | rcases hsat with ⟨w, k, s, hsat⟩, 18 | apply false.elim, apply a, 19 | rw trans_sat_iff at hsat, 20 | swap 3, exact k, swap, exact s, 21 | dsimp [mk_sseqt], rw list.append_nil, 22 | exact hsat, 23 | { dsimp [fml_is_sat] at h, dsimp [mk_sseqt] at eq, rw eq at h, contradiction } }, 24 | { split, intro, dsimp [fml_is_sat] at h, 25 | cases eq : tableau (mk_sseqt (list.map fml.to_nnf Γ)), 26 | { dsimp [mk_sseqt] at eq, rw eq at h, contradiction }, 27 | { have he := model_existence a_1.1, 28 | have h12:= a_1.2, 29 | cases a_1.val with tm ptm, 30 | cases tm with itm ltm sgtm, 31 | simp at h12, dsimp [manc] at he, 32 | have hman : manc (tmodel.cons itm ltm sgtm) = [], {simp, rw h12}, 33 | have hsub : itm.id.m ⊆ htk (tmodel.cons itm ltm sgtm), {simp, exact itm.mhtk}, 34 | have := he hman _ hsub, 35 | rcases this with ⟨st, k, w, hw⟩, 36 | split, split, split, 37 | rw h12 at hw, simp at hw, 38 | rw ←trans_sat_iff at hw, 39 | exact hw}, 40 | { simp } } 41 | end 42 | 43 | open fml 44 | 45 | def fml.exm : fml := or (var 1) (neg (var 1)) 46 | 47 | def K : fml := impl (box (impl (var 1) (var 2))) (impl (box $ var 1) (box $ var 2)) 48 | -------------------------------------------------------------------------------- /src/S4/test.lean: -------------------------------------------------------------------------------- 1 | import .tableau 2 | open fml 3 | 4 | def fml.iff (φ ψ : fml) := and (impl φ ψ) (impl ψ φ) 5 | 6 | def test (ϕ : fml) := fml_is_sat [neg ϕ] 7 | 8 | local infix ` ∨ `:1000 := fml.or 9 | local infix ` ∧ `:1000 := fml.and 10 | local infix ` -> `:1000 := fml.impl 11 | local infix ` <-> `:1000 := fml.iff 12 | local prefix `~` := fml.neg 13 | 14 | precedence `p`:max 15 | local prefix ` p ` := fml.var 16 | 17 | set_option profiler true 18 | 19 | -- #eval test () 20 | -------------------------------------------------------------------------------- /src/S4/vanilla.lean: -------------------------------------------------------------------------------- 1 | import .rules 2 | 3 | open psum nnf node 4 | 5 | set_option eqn_compiler.zeta true 6 | 7 | def tableau : Π Γ : sseqt, node Γ 8 | | Γ := 9 | match get_contra_seqt Γ with 10 | | inl w := contra_rule_seqt w.2 11 | | inr no_contra := 12 | match get_and_seqt Γ with 13 | | inl w := 14 | let inst := and_instance_seqt.cons w.2 in 15 | have h : prod.measure_lex' sseqt_size (and_child Γ w.2) Γ, 16 | begin apply split_lt_and_seqt end, 17 | let d_delta := tableau (and_child Γ w.2) in 18 | and_rule_seqt inst d_delta 19 | | inr no_and := 20 | match get_or_seqt Γ with 21 | | inl w := 22 | let inst := or_instance_seqt.cons w.2 in 23 | have h₁ : prod.measure_lex' sseqt_size (or_child_left Γ w.2) Γ, 24 | begin apply split_lt_or_seqt_left end, 25 | have h₂ : prod.measure_lex' sseqt_size (or_child_right Γ w.2) Γ, 26 | begin apply split_lt_or_seqt_right end, 27 | let Γ₁ := tableau (or_child_left Γ w.2) in 28 | match Γ₁ with 29 | | closed p := or_rule_seqt inst (closed p) (tableau (or_child_right Γ w.2)) 30 | | open_ w := open_rule_seqt inst w.2 31 | end 32 | | inr no_or := 33 | match get_box_seqt Γ with 34 | | inl w := -- term mode here helps termination 35 | if hb : box w.1 ∈ Γ.b then 36 | let inst := box_dup_instance_seqt.cons w.2 hb in 37 | have h : prod.measure_lex' sseqt_size (box_child Γ w.2) Γ, 38 | from copy_lt_seqt _ _, 39 | let d_delta := tableau (box_child Γ w.2) in 40 | box_rule_seqt inst d_delta 41 | else 42 | let inst := box_new_instance_seqt.cons w.2 hb in 43 | have h : prod.measure_lex' sseqt_size (box_child_new Γ w.2 hb) Γ, 44 | from box_new_lt_seqt _ _ _, 45 | let d_delta := tableau (box_child_new Γ w.2 hb) in 46 | box_new_rule_seqt inst d_delta 47 | | inr no_box := 48 | match get_dia_seqt Γ with 49 | | inl w := 50 | let ma : modal_applicable Γ := 51 | {satu := {no_and := no_and, no_or := no_or}, 52 | no_contra_main := no_contra, 53 | no_box_main := no_box, 54 | φ := w.1, 55 | ex := w.2} in 56 | let l := @dia_rule_seqt (λ Δ, prod.measure_lex' sseqt_size Δ Γ) 57 | (λ x h, tableau x) (unmodal_seqt Γ) 58 | (unmodal_seqt_size Γ) in 59 | match l with 60 | | inl w := begin left, exact unsat_of_unsat_unmodal w.1 w.2 end 61 | | inr w := 62 | begin 63 | right, split, swap, 64 | {let lm := models_to_tmodels w.1, 65 | let sgm := dia_rule_loop Γ.h Γ.b Γ.m, 66 | let mΓ : sseqt := Γ, 67 | let mhtk := Γ.m, 68 | have mhhtk : hintikka Γ.m, {apply hintikka_ma ma}, 69 | have mmhtk : Γ.m ⊆ Γ.m, {simp}, 70 | split, swap, 71 | {let minfo : info := ⟨mΓ, mhtk, mhhtk, mmhtk⟩, 72 | exact tmodel.cons minfo lm sgm }, 73 | {split, 74 | {split, 75 | {simp, intros s hs, have := list.mem_map.1 hs, 76 | rcases this with ⟨i, hmem, hi⟩, intros φ hφ, rw ←hi, 77 | apply mem_be_box, exact w.2, exact hmem, exact hφ}, 78 | {simp, intros s hs φ hφ, exfalso, apply no_box, exact hφ}, 79 | {simp, intros φ hφ, 80 | by_cases hc : φ ∈ Γ.h, 81 | {left, have := mem_loop_left _ Γ.b _ _ hc hφ, split, split, exact this, simp}, 82 | {right, have := mem_be_dia _ _ w.2 _ hφ hc, 83 | rcases this with ⟨i, hmem, hi⟩, 84 | have := list.mem_map_of_mem (λ x : model, x.val) hmem, 85 | split, split, exact this, exact hi} }, 86 | {simp, intros s rq hdesc hmem, 87 | have := ex_desc' _ _ lm _ hdesc, 88 | cases this with hl hr, 89 | {have hc := list.mem_map.1 hl, 90 | rcases hc with ⟨ms, pmsl, pmsr⟩, 91 | have hci := be_ex _ _ w.2 _ pmsl, 92 | rcases hci with ⟨iw, imem, pi⟩, 93 | have ps := pt_of_m_to_tm _ _ hl, 94 | have hsub := ps.1.sreq, 95 | cases s with is ls sgs, 96 | simp at hmem, simp at hsub, 97 | have hin := hsub hmem, 98 | rw pmsr at pi, simp at pi, rw pi at hin, 99 | have := unmodal_sig _ _ imem _ hin, 100 | cases this, 101 | {right, split, swap, exact tmodel.cons is ls sgs, split, exact hdesc, simp, rw this, rw pi}, 102 | {left, exact this}}, 103 | {rcases hr with ⟨c, memc, pdc⟩, 104 | have pc := pt_of_m_to_tm _ _ memc, 105 | have := pc.1.bdia, 106 | cases c with ic lc sgc, simp at this, 107 | have hc := this _ _ pdc hmem, 108 | cases hc with hl hr, 109 | {have hcc := list.mem_map.1 memc, 110 | rcases hcc with ⟨ms, pmsl, pmsr⟩, 111 | have hci := be_ex _ _ w.2 _ pmsl, 112 | rcases hci with ⟨iw, imem, pi⟩, 113 | have ps := pt_of_m_to_tm _ _ memc, 114 | rw pmsr at pi, simp at pi, rw pi at hl, 115 | have hccc := unmodal_sig _ _ imem _ hl, 116 | cases hccc, 117 | {right, split, swap, exact tmodel.cons ic lc sgc, split, apply tc'.base, simp, exact memc, simp, rw hccc, rw pi}, 118 | {left, exact hccc} }, 119 | {right, 120 | rcases hr with ⟨d, ddesc, hd⟩, 121 | split, swap, exact d, split, 122 | {apply desc_ex, split, split, exact memc, exact ddesc}, 123 | {exact hd}}}}, 124 | {simp, intros rq hrq φ hφ, cases hφ, 125 | {exfalso, apply ma.no_box_main, exact hφ}, 126 | {have := mem_loop_box _ _ _ _ hrq, rw this, exact hφ}}, 127 | {simp, intros s hs, have := mem_loop_dia _ _ _ _ hs, 128 | rcases this with ⟨w, hmem, hw⟩, 129 | have := Γ.ha _ hmem, 130 | have sb := mem_loop_box _ _ _ _ hs, 131 | cases s, rw ←hw at this, rw ←sb at this, simp at this, exact this}}, 132 | {intros d hd, have := ex_desc' _ _ _ _ hd, 133 | cases this, 134 | {have pd := pt_of_m_to_tm _ _ this, exact pd.1}, 135 | {rcases this with ⟨m, memm, mdesc⟩, 136 | have pd := pt_of_m_to_tm _ _ memm, 137 | apply pd.2, exact mdesc}}}}, 138 | {simp} 139 | end 140 | end 141 | | inr no_dia := 142 | let mc : model_constructible Γ := 143 | {satu := {no_and := no_and, no_or := no_or}, 144 | no_box_main := no_box, 145 | no_contra_main := no_contra, 146 | no_dia := no_dia} in build_model mc 147 | end 148 | end 149 | end 150 | end 151 | end 152 | using_well_founded {rel_tac := λ _ _, `[exact ⟨_, prod.measure_lex_wf' sseqt_size⟩], dec_tac := `[assumption]} 153 | 154 | @[simp] def mk_sseqt (Γ : list nnf) : sseqt := 155 | { goal := Γ, 156 | s := none, 157 | a := [], 158 | h := [], 159 | b := [], 160 | m := Γ, 161 | ndh := list.nodup_nil, 162 | ndb := list.nodup_nil, 163 | sph := list.nil_subperm, 164 | spb := list.nil_subperm, 165 | sbm := mem_closure_self _, 166 | ha := λ x hx, absurd hx $ list.not_mem_nil _, 167 | hb := box_only_nil, 168 | ps₁ := λ h, by contradiction, 169 | ps₂ := λ h, by contradiction } 170 | 171 | def is_sat (Γ : list nnf) : bool := 172 | match tableau (mk_sseqt Γ) with 173 | | closed _ := ff 174 | | open_ _ := tt 175 | end 176 | 177 | theorem model_existence (m : model) (hrt : manc m.1 = []) 178 | (Γ : list nnf) (h : Γ ⊆ htk m.1) : ∃ (st : Type) (k : S4 st) s, sat k s Γ := 179 | begin 180 | split, swap, 181 | exact {x : rmodel // x.1 = m.1 ∨ desc x.1 m.1}, 182 | split, swap, 183 | exact builder m.1, 184 | split, swap, 185 | split, 186 | left, swap, exact ⟨m.1, m.2.1⟩, 187 | simp, 188 | intros φ hφ, apply good_model, 189 | exact hrt, apply h, exact hφ 190 | end 191 | 192 | theorem correctness (Γ : list nnf) : is_sat Γ = tt ↔ ∃ (st : Type) (k : S4 st) s, sat k s Γ := 193 | begin 194 | cases h : is_sat Γ, 195 | constructor, 196 | {intro, contradiction}, 197 | {intro hsat, cases eq : tableau (mk_sseqt Γ), 198 | rcases hsat with ⟨w, k, s, hsat⟩, 199 | apply false.elim, apply a, simp, exact hsat, 200 | {dsimp [is_sat] at h, dsimp at eq, rw eq at h, contradiction}}, 201 | {split, intro, dsimp [is_sat] at h, 202 | cases eq : tableau (mk_sseqt Γ), 203 | {dsimp at eq, rw eq at h, contradiction}, 204 | {apply model_existence, swap 3, exact a_1.1, 205 | have := a_1.2, simp at this, 206 | cases a_1.val with tm ptm, 207 | cases tm with itm ltm sgtm, 208 | simp, simp at this, rw this, 209 | have := a_1.2, simp at this, 210 | cases a_1.val with tm ptm, 211 | cases tm with itm ltm sgtm, 212 | have hsub := itm.mhtk, simp, simp at this, 213 | rw this at hsub, dsimp at hsub, exact hsub}, 214 | {simp}} 215 | end 216 | 217 | -- negation of K 218 | def φ : nnf := 219 | and (box $ or (neg 1) (var 2)) (and (box $ var 1) (dia $ neg 2)) 220 | 221 | #eval is_sat [φ] -- ff 222 | 223 | -- negation of S4 224 | def ψ : nnf := and (box (var 1)) (dia (dia (neg 1))) 225 | 226 | #eval is_sat [ψ] -- ff 227 | 228 | def γ : nnf := and (and (dia (var 1)) (dia (var 2))) (box (or (neg 1) (neg 2))) 229 | 230 | #eval is_sat [γ] -- tt 231 | -------------------------------------------------------------------------------- /src/apps/examples.lean: -------------------------------------------------------------------------------- 1 | -- import .topo_translation ..K.jump 2 | 3 | -- open nnf 4 | 5 | -- def boom : nnf := dia (and (var 1) (neg 1)) 6 | 7 | -- -- #eval is_sat [boom] 8 | 9 | -- -- #check tableau [boom] 10 | 11 | -- def topo_fact_of (φ : nnf) : psum (∀ (α) (tm : topo_model α) s, ¬ topo_force tm s φ) unit := 12 | -- match tableau [φ] with 13 | -- | node.closed _ h _ := psum.inl $ not_topo_force_of_unsat h 14 | -- | node.open_ _ := psum.inr () 15 | -- end 16 | -------------------------------------------------------------------------------- /src/apps/topo_translation.lean: -------------------------------------------------------------------------------- 1 | import .topological_semantics ..K.semantics 2 | open nnf 3 | 4 | -- local attribute [instance] classical.prop_decidable 5 | 6 | @[simp] def topo_to_kripke {α : Type} (tm : topo_model α) : kripke α := 7 | { rel := λ s t, s ∈ @_root_.closure _ tm.to_topological_space {t}, 8 | val := λ n s, tm.v n s } 9 | 10 | theorem trans_force_left {α : Type} {tm : topo_model α} : Π {s} {φ : nnf}, (topo_force tm s φ) → force (topo_to_kripke tm) s φ 11 | | s (var n) h := by dsimp at h; simp [h] 12 | | s (neg n) h := by dsimp at h; simp [h] 13 | | s (and φ ψ) h := begin cases h with l r, split, apply trans_force_left l, apply trans_force_left r end 14 | | s (or φ ψ) h := begin cases h, left, apply trans_force_left h, right, apply trans_force_left h end 15 | | s (box φ) h := 16 | begin 17 | rcases h with ⟨w, hw, hmem⟩, 18 | intros s' hs', 19 | apply trans_force_left, 20 | apply hw.2, 21 | rw ←set.inter_singleton_nonempty, 22 | have := (@mem_closure_iff _ tm.to_topological_space _ _).1, 23 | swap 3, exact {s'}, apply this, simpa using hs', 24 | exact hw.1, exact hmem 25 | end 26 | | s (dia φ) h := 27 | begin 28 | let ts := tm.to_topological_space, 29 | let o := ⋂₀ {x : set α | s ∈ x ∧ @is_open _ ts x}, 30 | have openo: @is_open _ ts o, 31 | { apply tm.is_alex, intros, exact H.2 }, 32 | have hmem : s ∈ o, 33 | { rw set.mem_sInter, intros, exact H.1 }, 34 | have ex := (@mem_closure_iff _ tm.to_topological_space _ _).1 h o openo hmem, 35 | cases ex with w hw, dsimp, 36 | split, split, swap 3, 37 | { exact w }, 38 | rw (@mem_closure_iff _ tm.to_topological_space _ _), 39 | { --apply to_bool_true, 40 | intros o' hopen' hmem', 41 | have : o ⊆ o', 42 | { intro x, intro hx, rw set.mem_sInter at hx, apply hx, split, repeat {assumption} }, 43 | rw set.inter_singleton_nonempty, 44 | apply this hw.1 }, 45 | { exact trans_force_left hw.2 } 46 | end 47 | 48 | theorem sat_of_topo_sat {Γ : list nnf} 49 | {α : Type} (tm : topo_model α) (s) (h : topo_sat tm s Γ) : 50 | sat (topo_to_kripke tm) s Γ := λ φ hφ, trans_force_left $ h φ hφ 51 | 52 | theorem unsat_topo_of_unsat {Γ : list nnf} : unsatisfiable Γ → topo_unsatisfiable Γ := 53 | λ h, (λ α tm s hsat, @h _ (topo_to_kripke tm) s (sat_of_topo_sat _ _ hsat)) 54 | 55 | def not_topo_force_of_unsat {φ} : unsatisfiable [φ] → ∀ (α) (tm : topo_model α) s, ¬ topo_force tm s φ := 56 | λ h, topo_unsat_singleton $ unsat_topo_of_unsat h 57 | 58 | -------------------------------------------------------------------------------- /src/apps/topological_semantics.lean: -------------------------------------------------------------------------------- 1 | import defs topology.basic 2 | open nnf 3 | 4 | structure topo_model (α : Type) extends topological_space α := 5 | (v : ℕ → set α) 6 | (is_alex : ∀s, (∀t∈s, is_open t) → is_open (⋂₀ s)) 7 | 8 | @[simp] def topo_force {α : Type} (tm : topo_model α) : α → nnf → Prop 9 | | s (var n) := tm.v n s 10 | | s (neg n) := ¬ tm.v n s 11 | | s (and φ ψ) := topo_force s φ ∧ topo_force s ψ 12 | | s (or φ ψ) := topo_force s φ ∨ topo_force s ψ 13 | | s (box φ) := @interior _ tm.to_topological_space (λ a, topo_force a φ) s 14 | | s (dia φ) := @_root_.closure _ tm.to_topological_space (λ a, topo_force a φ) s 15 | 16 | def topo_sat {α : Type} (tm : topo_model α) (s) (Γ : list nnf) : Prop := ∀ φ ∈ Γ, topo_force tm s φ 17 | 18 | def topo_unsatisfiable (Γ : list nnf) : Prop := 19 | ∀ (α) (tm : topo_model α) s, ¬ topo_sat tm s Γ 20 | 21 | def topo_unsat_singleton {φ} : topo_unsatisfiable [φ] → 22 | ∀ (α) (tm : topo_model α) s, ¬ topo_force tm s φ := 23 | begin 24 | intro h, intros, intro hf, 25 | apply h, intros ψ hψ, rw list.mem_singleton at hψ, rw hψ, exact hf 26 | end 27 | -------------------------------------------------------------------------------- /src/defs.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2018-2019 Minchao Wu. All rights reserved. 3 | Released under MIT license as described in the file LICENSE. 4 | Author: Minchao Wu 5 | -/ 6 | import data.list 7 | 8 | open nat tactic subtype 9 | 10 | inductive nnf : Type 11 | | var (n : nat) 12 | | neg (n : nat) 13 | | and (φ ψ : nnf) 14 | | or (φ ψ : nnf) 15 | | box (φ : nnf) 16 | | dia (φ : nnf) 17 | 18 | open nnf 19 | instance : inhabited nnf := ⟨nnf.var 0⟩ 20 | 21 | class no_literals (Γ : list nnf) := 22 | (no_var : ∀ {n}, var n ∉ Γ) 23 | (no_neg : ∀ {n}, neg n ∉ Γ) 24 | 25 | class saturated (Γ : list nnf) := 26 | (no_and : ∀ {φ ψ}, nnf.and φ ψ ∉ Γ) 27 | (no_or : ∀ {φ ψ}, nnf.or φ ψ ∉ Γ) 28 | 29 | class box_only (Γ : list nnf) extends no_literals Γ, saturated Γ := 30 | (no_dia : ∀ {φ}, nnf.dia φ ∉ Γ) 31 | 32 | def box_only_nil : box_only [] := 33 | {no_var := λ n hn, absurd hn $ list.not_mem_nil _, 34 | no_neg := λ n hn, absurd hn $ list.not_mem_nil _, 35 | no_and := λ φ ψ hn, absurd hn $ list.not_mem_nil _, 36 | no_or := λ φ ψ hn, absurd hn $ list.not_mem_nil _, 37 | no_dia := λ n hn, absurd hn $ list.not_mem_nil _} 38 | 39 | def cons_box_only {φ Γ} (h : box_only Γ) : box_only (box φ :: Γ) := 40 | {no_var := begin intros n hn, cases hn, contradiction, apply h.no_var, exact hn end, 41 | no_neg := begin intros n hn, cases hn, contradiction, apply h.no_neg, exact hn end, 42 | no_and := begin intros ψ₁ ψ₂ hψ, cases hψ, contradiction, apply h.no_and, exact hψ end, 43 | no_or := begin intros ψ₁ ψ₂ hψ, cases hψ, contradiction, apply h.no_or, exact hψ end, 44 | no_dia := begin intros n hn, cases hn, contradiction, apply h.no_dia, exact hn end} 45 | 46 | def box_only_ex : Π {φ Γ} (h₁ : box_only Γ) (h₂ : φ ∈ Γ), ∃ ψ, φ = box ψ 47 | | (var n) Γ h₁ h₂ := begin exfalso, apply h₁.no_var, exact h₂ end 48 | | (neg n) Γ h₁ h₂ := begin exfalso, apply h₁.no_neg, exact h₂ end 49 | | (and φ ψ) Γ h₁ h₂ := begin exfalso, apply h₁.no_and, exact h₂ end 50 | | (or φ ψ) Γ h₁ h₂ := begin exfalso, apply h₁.no_or, exact h₂ end 51 | | (box φ) Γ h₁ h₂ := ⟨φ, rfl⟩ 52 | | (dia φ) Γ h₁ h₂ := begin exfalso, apply h₁.no_dia, exact h₂ end 53 | 54 | def nnf.to_string : nnf → string 55 | | (var n) := "P" ++ n.repr 56 | | (neg n) := "¬P" ++ n.repr 57 | | (and φ ψ) := nnf.to_string φ ++ "∧" ++ nnf.to_string ψ 58 | | (or φ ψ) := nnf.to_string φ ++ "∨" ++ nnf.to_string ψ 59 | | (box φ) := "□" ++ "(" ++ nnf.to_string φ ++ ")" 60 | | (dia φ) := "◇" ++ "(" ++ nnf.to_string φ ++ ")" 61 | 62 | instance nnf_repr : has_repr nnf := ⟨nnf.to_string⟩ 63 | 64 | instance dec_eq_nnf : decidable_eq nnf := by mk_dec_eq_instance 65 | 66 | structure kripke (states : Type) := 67 | (val : ℕ → states → Prop) 68 | (rel : states → states → Prop) 69 | 70 | instance inhabited_kripke : inhabited (kripke ℕ) := 71 | ⟨{ val := λ a b, tt, rel := λ a b, tt }⟩ 72 | 73 | open nnf 74 | 75 | /- Do not use map -/ 76 | @[simp] def node_size : list nnf → ℕ 77 | | [] := 0 78 | | (hd :: tl) := sizeof hd + node_size tl 79 | 80 | inductive close_instance (Γ : list nnf) 81 | | cons : Π {n}, var n ∈ Γ → neg n ∈ Γ → close_instance 82 | 83 | inductive and_instance (Γ : list nnf) : list nnf → Type 84 | | cons : Π {φ ψ}, nnf.and φ ψ ∈ Γ → 85 | and_instance (φ :: ψ :: Γ.erase (nnf.and φ ψ)) 86 | 87 | inductive or_instance (Γ : list nnf) : list nnf → list nnf → Type 88 | | cons : Π {φ ψ}, nnf.or φ ψ ∈ Γ → 89 | or_instance (φ :: Γ.erase (nnf.or φ ψ)) 90 | (ψ :: Γ.erase (nnf.or φ ψ)) 91 | 92 | def left_prcp : Π {Γ₁ Γ₂ Δ : list nnf} (i : or_instance Δ Γ₁ Γ₂), nnf 93 | | Γ₁ Γ₂ Δ (@or_instance.cons _ φ ψ _) := φ 94 | 95 | @[simp] def sub_nnf : nnf → list nnf 96 | | (var n) := [var n] 97 | | (neg n) := [neg n] 98 | | (and φ ψ) := and φ ψ :: (sub_nnf φ ++ sub_nnf ψ) 99 | | (or φ ψ) := or φ ψ :: (sub_nnf φ ++ sub_nnf ψ) 100 | | (box φ) := box φ :: sub_nnf φ 101 | | (dia φ) := dia φ :: sub_nnf φ 102 | 103 | theorem mem_sub_nnf_self : Π φ : nnf, φ ∈ sub_nnf φ := 104 | begin intro φ, cases φ, all_goals {simp} end 105 | 106 | theorem mem_sub_nnf_and {φ ψ : nnf} : Π γ : nnf, and φ ψ ∈ sub_nnf γ → φ ∈ sub_nnf γ ∧ ψ ∈ sub_nnf γ 107 | | (var n) h := begin simp at h, contradiction end 108 | | (neg n) h := begin simp at h, contradiction end 109 | | (and φ₁ ψ₁) h := 110 | begin 111 | simp only [sub_nnf] at h, 112 | cases h, 113 | {injection h with eq₁ eq₂, 114 | rw eq₁, rw eq₂, 115 | simp only [sub_nnf], 116 | split, 117 | {simp, right, left, apply mem_sub_nnf_self}, 118 | {simp, right, right, apply mem_sub_nnf_self}}, 119 | {have := (@list.mem_append _ _ _ _).1 h, 120 | cases this, 121 | {split, 122 | {simp, right, left, apply (mem_sub_nnf_and _ _).1, exact this}, 123 | {simp, right, left, apply (mem_sub_nnf_and _ _).2, exact this}}, 124 | {split, 125 | {simp, right, right, apply (mem_sub_nnf_and _ _).1, exact this}, 126 | {simp, right, right, apply (mem_sub_nnf_and _ _).2, exact this}}} 127 | end 128 | | (or φ₁ ψ₁) h := 129 | begin 130 | simp at h, 131 | cases h, 132 | {split, 133 | {simp, right, left, apply (mem_sub_nnf_and _ _).1, exact h}, 134 | {simp, right, left, apply (mem_sub_nnf_and _ _).2, exact h}}, 135 | {split, 136 | {simp, right, right, apply (mem_sub_nnf_and _ _).1, exact h}, 137 | {simp, right, right, apply (mem_sub_nnf_and _ _).2, exact h}} 138 | end 139 | | (box φ₁) h := 140 | begin 141 | simp at h, 142 | {split, 143 | {simp, right, apply (mem_sub_nnf_and _ _).1, exact h}, 144 | {simp, right, apply (mem_sub_nnf_and _ _).2, exact h}} 145 | end 146 | | (dia φ₁) h := 147 | begin 148 | simp at h, 149 | {split, 150 | {simp, right, apply (mem_sub_nnf_and _ _).1, exact h}, 151 | {simp, right, apply (mem_sub_nnf_and _ _).2, exact h}} 152 | end 153 | 154 | theorem mem_sub_nnf_or {φ ψ : nnf} : Π γ : nnf, or φ ψ ∈ sub_nnf γ → φ ∈ sub_nnf γ ∧ ψ ∈ sub_nnf γ 155 | | (var n) h := begin simp at h, contradiction end 156 | | (neg n) h := begin simp at h, contradiction end 157 | | (and φ₁ ψ₁) h := 158 | begin 159 | simp at h, 160 | cases h, 161 | {split, 162 | {simp, right, left, apply (mem_sub_nnf_or _ _).1, exact h}, 163 | {simp, right, left, apply (mem_sub_nnf_or _ _).2, exact h}}, 164 | {split, 165 | {simp, right, right, apply (mem_sub_nnf_or _ _).1, exact h}, 166 | {simp, right, right, apply (mem_sub_nnf_or _ _).2, exact h}} 167 | end 168 | | (or φ₁ ψ₁) h := 169 | begin 170 | simp only [sub_nnf] at h, 171 | cases h, 172 | {injection h with eq₁ eq₂, 173 | rw eq₁, rw eq₂, 174 | simp only [sub_nnf], 175 | split, 176 | {simp, right, left, apply mem_sub_nnf_self}, 177 | {simp, right, right, apply mem_sub_nnf_self}}, 178 | {have := (@list.mem_append _ _ _ _).1 h, 179 | cases this, 180 | {split, 181 | {simp, right, left, apply (mem_sub_nnf_or _ _).1, exact this}, 182 | {simp, right, left, apply (mem_sub_nnf_or _ _).2, exact this}}, 183 | {split, 184 | {simp, right, right, apply (mem_sub_nnf_or _ _).1, exact this}, 185 | {simp, right, right, apply (mem_sub_nnf_or _ _).2, exact this}}} 186 | end 187 | | (box φ₁) h := 188 | begin 189 | simp at h, 190 | {split, 191 | {simp, right, apply (mem_sub_nnf_or _ _).1, exact h}, 192 | {simp, right, apply (mem_sub_nnf_or _ _).2, exact h}} 193 | end 194 | | (dia φ₁) h := 195 | begin 196 | simp at h, 197 | {split, 198 | {simp, right, apply (mem_sub_nnf_or _ _).1, exact h}, 199 | {simp, right, apply (mem_sub_nnf_or _ _).2, exact h}} 200 | end 201 | 202 | theorem mem_sub_nnf_box {φ : nnf} : Π γ : nnf, box φ ∈ sub_nnf γ → φ ∈ sub_nnf γ 203 | | (var n) h := begin simp at h, contradiction end 204 | | (neg n) h := begin simp at h, contradiction end 205 | | (and φ₁ ψ₁) h := 206 | begin 207 | simp at h, 208 | cases h, 209 | {simp, right, left, apply mem_sub_nnf_box, exact h}, 210 | {simp, right, right, apply mem_sub_nnf_box, exact h} 211 | end 212 | | (or φ₁ ψ₁) h := 213 | begin 214 | simp at h, 215 | cases h, 216 | {simp, right, left, apply mem_sub_nnf_box, exact h}, 217 | {simp, right, right, apply mem_sub_nnf_box, exact h} 218 | end 219 | | (box φ₁) h := 220 | begin 221 | simp at h, 222 | cases h, 223 | {rw h, simp, right, apply mem_sub_nnf_self}, 224 | {simp, right, apply mem_sub_nnf_box, exact h} 225 | end 226 | | (dia φ₁) h := 227 | begin 228 | simp at h, 229 | simp, right, apply mem_sub_nnf_box, exact h 230 | end 231 | 232 | theorem mem_sub_nnf_dia {φ : nnf} : Π γ : nnf, dia φ ∈ sub_nnf γ → φ ∈ sub_nnf γ 233 | | (var n) h := begin simp at h, contradiction end 234 | | (neg n) h := begin simp at h, contradiction end 235 | | (and φ₁ ψ₁) h := 236 | begin 237 | simp at h, 238 | cases h, 239 | {simp, right, left, apply mem_sub_nnf_dia, exact h}, 240 | {simp, right, right, apply mem_sub_nnf_dia, exact h} 241 | end 242 | | (or φ₁ ψ₁) h := 243 | begin 244 | simp at h, 245 | cases h, 246 | {simp, right, left, apply mem_sub_nnf_dia, exact h}, 247 | {simp, right, right, apply mem_sub_nnf_dia, exact h} 248 | end 249 | | (box φ₁) h := 250 | begin 251 | simp at h, 252 | simp, right, apply mem_sub_nnf_dia, exact h 253 | end 254 | | (dia φ₁) h := 255 | begin 256 | simp at h, 257 | cases h, 258 | {rw h, simp, right, apply mem_sub_nnf_self}, 259 | {simp, right, apply mem_sub_nnf_dia, exact h} 260 | end 261 | 262 | @[simp] def nnf.closure : list nnf → list nnf 263 | | [] := [] 264 | | (hd::tl) := sub_nnf hd ++ nnf.closure tl 265 | 266 | theorem mem_closure_self : Π Γ : list nnf, Γ ⊆ nnf.closure Γ 267 | | [] := λ x h, absurd h $ list.not_mem_nil _ 268 | | (hd::tl) := 269 | begin 270 | intros x hx, cases hx, 271 | {rw hx, simp, left, apply mem_sub_nnf_self}, 272 | {simp, right, apply mem_closure_self, exact hx} 273 | end 274 | 275 | theorem mem_closure_and {φ ψ : nnf} : Π Γ : list nnf, and φ ψ ∈ nnf.closure Γ → φ ∈ nnf.closure Γ ∧ ψ ∈ nnf.closure Γ 276 | | [] h := absurd h $ list.not_mem_nil _ 277 | | (hd::tl) h := 278 | begin 279 | simp at h, 280 | cases h, 281 | {split, 282 | {simp, left, apply (mem_sub_nnf_and hd _).1, swap, exact h}, 283 | {simp, left, apply (mem_sub_nnf_and hd _).2, swap, exact h}}, 284 | {split, 285 | {simp, right, apply (mem_closure_and _ _).1, exact h}, 286 | {simp, right, apply (mem_closure_and _ _).2, exact h} } 287 | end 288 | 289 | theorem mem_closure_or {φ ψ : nnf} : Π Γ : list nnf, or φ ψ ∈ nnf.closure Γ → φ ∈ nnf.closure Γ ∧ ψ ∈ nnf.closure Γ 290 | | [] h := absurd h $ list.not_mem_nil _ 291 | | (hd::tl) h := 292 | begin 293 | simp at h, 294 | cases h, 295 | {split, 296 | {simp, left, apply (mem_sub_nnf_or hd _).1, swap, exact h}, 297 | {simp, left, apply (mem_sub_nnf_or hd _).2, swap, exact h}}, 298 | {split, 299 | {simp, right, apply (mem_closure_or _ _).1, exact h}, 300 | {simp, right, apply (mem_closure_or _ _).2, exact h} } 301 | end 302 | 303 | theorem mem_closure_box {φ : nnf} : Π Γ : list nnf, box φ ∈ nnf.closure Γ → φ ∈ nnf.closure Γ 304 | | [] h := absurd h $ list.not_mem_nil _ 305 | | (hd::tl) h := 306 | begin 307 | simp at h, 308 | cases h, 309 | {simp, left, apply mem_sub_nnf_box _ h}, 310 | {simp, right, apply mem_closure_box _ h} 311 | end 312 | 313 | theorem mem_closure_dia {φ : nnf} : Π Γ : list nnf, dia φ ∈ nnf.closure Γ → φ ∈ nnf.closure Γ 314 | | [] h := absurd h $ list.not_mem_nil _ 315 | | (hd::tl) h := 316 | begin 317 | simp at h, 318 | cases h, 319 | {simp, left, apply mem_sub_nnf_dia _ h}, 320 | {simp, right, apply mem_closure_dia _ h} 321 | end 322 | 323 | @[simp] def filter_undia (l : list nnf) : list nnf → list nnf 324 | | [] := [] 325 | | (dia φ::tl) := if φ ∈ l then filter_undia tl else φ :: filter_undia tl 326 | | (e::tl) := filter_undia tl 327 | 328 | theorem mem_filter_undia_left : Π l Γ φ (h₁ : dia φ ∈ Γ) (h₂ : φ ∉ l), 329 | φ ∈ filter_undia l Γ 330 | | l [] φ h₁ h₂ := absurd h₁ $ list.not_mem_nil _ 331 | | l (hd :: tl) φ h₁ h₂ := 332 | begin 333 | cases h : hd, 334 | case nnf.dia : ψ 335 | {rw h at h₁, simp, 336 | by_cases hc : ψ ∈ l, 337 | {rw if_pos, apply mem_filter_undia_left, cases h₁, {injection h₁ with h₁', rw ←h₁' at hc, contradiction}, {exact h₁}, repeat {assumption}}, 338 | {rw if_neg, cases h₁, {injection h₁ with h₁', rw h₁', simp}, {right, apply mem_filter_undia_left, exact h₁, exact h₂}, exact hc}}, 339 | all_goals 340 | {simp, rw h at h₁, cases h₁, contradiction, apply mem_filter_undia_left, repeat {assumption}} 341 | end 342 | 343 | theorem mem_filter_dia_right_aux : Π l Γ φ, φ ∈ filter_undia l Γ → dia φ ∈ Γ 344 | | l [] φ h := absurd h $ list.not_mem_nil _ 345 | | l (hd :: tl) φ h := 346 | begin 347 | cases h₁ : hd, 348 | case nnf.dia : ψ 349 | {rw h₁ at h, dsimp at h, 350 | by_cases hc : ψ ∈ l, 351 | {rw if_pos at h, have := mem_filter_dia_right_aux l tl φ h, right, exact this, exact hc}, 352 | {rw if_neg at h, cases h, {rw h, simp}, {have := mem_filter_dia_right_aux l tl φ h, right, exact this}, exact hc}}, 353 | all_goals 354 | {rw h₁ at h, simp at h, right, apply mem_filter_dia_right_aux, exact h} 355 | end 356 | 357 | theorem mem_filter_dia_right : Π l Γ φ (h₂ : φ ∈ filter_undia l Γ), φ ∉ l 358 | | l [] φ h := absurd h $ list.not_mem_nil _ 359 | | l (hd :: tl) φ h := 360 | begin 361 | cases h₁ : hd, 362 | case nnf.dia : ψ 363 | {rw h₁ at h, dsimp at h, 364 | by_cases hc : ψ ∈ l, 365 | {rw if_pos at h, have := mem_filter_dia_right l tl φ h, exact this, exact hc}, 366 | {rw if_neg at h, cases h, {rw h, exact hc}, {have := mem_filter_dia_right l tl φ h, exact this}, exact hc}}, 367 | all_goals 368 | {rw h₁ at h, simp at h, apply mem_filter_dia_right, exact h} 369 | end 370 | 371 | structure hintikka (Γ : list nnf) : Prop := 372 | (hno_contra : ∀ {n}, var n ∈ Γ → neg n ∉ Γ) 373 | (hand_left : ∀ {φ ψ}, nnf.and φ ψ ∈ Γ → φ ∈ Γ) 374 | (hand_right : ∀ {φ ψ}, nnf.and φ ψ ∈ Γ → ψ ∈ Γ) 375 | (hor : ∀ {φ ψ}, nnf.or φ ψ ∈ Γ → φ ∈ Γ ∨ ψ ∈ Γ) 376 | (hbox : ∀ {φ}, nnf.box φ ∈ Γ → φ ∈ Γ) 377 | --------------------------------------------------------------------------------