This repository was archived by the owner on Mar 8, 2025. It is now read-only.
File tree Expand file tree Collapse file tree 2 files changed +13
-4
lines changed
Expand file tree Collapse file tree 2 files changed +13
-4
lines changed Original file line number Diff line number Diff line change @@ -18,7 +18,8 @@ lemma re_iff_sigma1 {P : ℕ → Prop} : RePred P ↔ 𝚺₁-Predicate P := by
1818
1919variable (T : Theory ℒₒᵣ) [𝐑₀ ≼ T] [Sigma1Sound T] [T.Delta1Definable]
2020
21- theorem incomplete : ¬System.Complete T := by
21+ /-- Gödel's First Incompleteness Theorem-/
22+ theorem goedel_first_incompleteness : ¬System.Complete T := by
2223 let D : ℕ → Prop := fun n : ℕ ↦ ∃ p : SyntacticSemiformula ℒₒᵣ 1 , n = ⌜p⌝ ∧ T ⊢! ∼p/[⌜p⌝]
2324 have D_re : RePred D := by
2425 have : 𝚺₁-Predicate fun p : ℕ ↦
Original file line number Diff line number Diff line change 138138
139139section
140140
141+ variable (T)
142+
141143variable [T.Delta1Definable]
142144
143145open LO.System LO.System.FiniteContext
@@ -179,11 +181,17 @@ lemma consistent_iff_goedel : T ⊢! ↑𝗖𝗼𝗻 ⭤ ↑𝗚 := by
179181 simpa [provable₀_iff] using contra₁'! (deduct'! this)
180182
181183/-- Gödel's Second Incompleteness Theorem-/
182- theorem consistent_unprovable [System.Consistent T] : T ⊬ ↑𝗖𝗼𝗻 := fun h ↦
183- goedel_unprovable <| and_left! consistent_iff_goedel ⨀ h
184+ theorem goedel_second_incompleteness [System.Consistent T] : T ⊬ ↑𝗖𝗼𝗻 := fun h ↦
185+ goedel_unprovable T <| and_left! ( consistent_iff_goedel T) ⨀ h
184186
185187theorem inconsistent_unprovable [ℕ ⊧ₘ* T] : T ⊬ ∼↑𝗖𝗼𝗻 := fun h ↦
186- not_goedel_unprovable <| contra₀'! (and_right! (consistent_iff_goedel (T := T))) ⨀ h
188+ not_goedel_unprovable T <| contra₀'! (and_right! (consistent_iff_goedel T)) ⨀ h
189+
190+ theorem inconsistent_undecidable [ℕ ⊧ₘ* T] : System.Undecidable T ↑𝗖𝗼𝗻 := by
191+ haveI : Consistent T := Sound.consistent_of_satisfiable ⟨_, (inferInstance : ℕ ⊧ₘ* T)⟩
192+ constructor
193+ · exact goedel_second_incompleteness T
194+ · exact inconsistent_unprovable T
187195
188196end
189197
You can’t perform that action at this time.
0 commit comments