Theory Propositional
section ‹Classical sequent calculus: examples with propositional connectives›
theory Propositional
imports "../LK"
begin
text "absorptive laws of ∧ and ∨"
lemma "⊢ P ∧ P ⟷ P"
by fast_prop
lemma "⊢ P ∨ P ⟷ P"
by fast_prop
text "commutative laws of ∧ and ∨"
lemma "⊢ P ∧ Q ⟷ Q ∧ P"
by fast_prop
lemma "⊢ P ∨ Q ⟷ Q ∨ P"
by fast_prop
text "associative laws of ∧ and ∨"
lemma "⊢ (P ∧ Q) ∧ R ⟷ P ∧ (Q ∧ R)"
by fast_prop
lemma "⊢ (P ∨ Q) ∨ R ⟷ P ∨ (Q ∨ R)"
by fast_prop
text "distributive laws of ∧ and ∨"
lemma "⊢ (P ∧ Q) ∨ R ⟷ (P ∨ R) ∧ (Q ∨ R)"
by fast_prop
lemma "⊢ (P ∨ Q) ∧ R ⟷ (P ∧ R) ∨ (Q ∧ R)"
by fast_prop
text "Laws involving implication"
lemma "⊢ (P ∨ Q ⟶ R) ⟷ (P ⟶ R) ∧ (Q ⟶ R)"
by fast_prop
lemma "⊢ (P ∧ Q ⟶ R) ⟷ (P ⟶ (Q ⟶ R))"
by fast_prop
lemma "⊢ (P ⟶ Q ∧ R) ⟷ (P ⟶ Q) ∧ (P ⟶ R)"
by fast_prop
text "Classical theorems"
lemma "⊢ P ∨ Q ⟶ P ∨ ¬ P ∧ Q"
by fast_prop
lemma "⊢ (P ⟶ Q) ∧ (¬ P ⟶ R) ⟶ (P ∧ Q ∨ R)"
by fast_prop
lemma "⊢ P ∧ Q ∨ ¬ P ∧ R ⟷ (P ⟶ Q) ∧ (¬ P ⟶ R)"
by fast_prop
lemma "⊢ (P ⟶ Q) ∨ (P ⟶ R) ⟷ (P ⟶ Q ∨ R)"
by fast_prop
lemma "⊢ (P ⟷ Q) ⟷ (Q ⟷ P)"
by fast_prop
lemma "⊢ ¬ (P ⟷ ¬ P)"
by fast_prop
lemma "⊢ (P ⟶ Q) ⟷ (¬ Q ⟶ ¬ P)"
by fast_prop
lemma "⊢ ¬ ¬ P ⟷ P"
by fast_prop
lemma "⊢ ¬ (P ⟶ Q) ⟶ (Q ⟶ P)"
by fast_prop
lemma "⊢ (¬ P ⟶ Q) ⟷ (¬ Q ⟶ P)"
by fast_prop
lemma "⊢ ((P ∨ Q) ⟶ (P ∨ R)) ⟶ (P ∨ (Q ⟶ R))"
by fast_prop
lemma "⊢ P ∨ ¬ P"
by fast_prop
lemma "⊢ P ∨ ¬ ¬ ¬ P"
by fast_prop
lemma "⊢ ((P ⟶ Q) ⟶ P) ⟶ P"
by fast_prop
lemma "⊢ ((P ∨ Q) ∧ (¬ P ∨ Q) ∧ (P ∨ ¬ Q)) ⟶ ¬ (¬ P ∨ ¬ Q)"
by fast_prop
lemma "Q ⟶ R, R ⟶ P ∧ Q, P ⟶ (Q ∨ R) ⊢ P ⟷ Q"
by fast_prop
lemma "⊢ P ⟷ P"
by fast_prop
lemma "⊢ ((P ⟷ Q) ⟷ R) ⟷ (P ⟷ (Q ⟷ R))"
by fast_prop
lemma "⊢ P ∨ (Q ∧ R) ⟷ (P ∨ Q) ∧ (P ∨ R)"
by fast_prop
lemma "⊢ (P ⟷ Q) ⟷ ((Q ∨ ¬ P) ∧ (¬ Q ∨ P))"
by fast_prop
lemma "⊢ (P ⟶ Q) ⟷ (¬ P ∨ Q)"
by fast_prop
lemma "⊢ (P ⟶ Q) ∨ (Q ⟶ P)"
by fast_prop
lemma "⊢ ((P ∧ (Q ⟶ R)) ⟶ S) ⟷ ((¬ P ∨ Q ∨ S) ∧ (¬ P ∨ ¬ R ∨ S))"
by fast_prop
end