Up to index of Isabelle/FOL
theory FOL(* Title: FOL/FOL.thy ID: $Id: FOL.thy,v 1.32 2005/06/17 14:15:08 haftmann Exp $ Author: Lawrence C Paulson and Markus Wenzel *) header {* Classical first-order logic *} theory FOL imports IFOL uses ("FOL_lemmas1.ML") ("cladata.ML") ("blastdata.ML") ("simpdata.ML") ("eqrule_FOL_data.ML") ("~~/src/Provers/eqsubst.ML") begin subsection {* The classical axiom *} axioms classical: "(~P ==> P) ==> P" subsection {* Lemmas and proof tools *} use "FOL_lemmas1.ML" theorems case_split = case_split_thm [case_names True False, cases type: o] use "cladata.ML" setup Cla.setup setup cla_setup setup case_setup use "blastdata.ML" setup Blast.setup lemma ex1_functional: "[| EX! z. P(a,z); P(a,b); P(a,c) |] ==> b = c" by blast ML {* val ex1_functional = thm "ex1_functional"; *} use "simpdata.ML" setup simpsetup setup "Simplifier.method_setup Splitter.split_modifiers" setup Splitter.setup setup Clasimp.setup subsection {* Lucas Dixon's eqstep tactic *} use "~~/src/Provers/eqsubst.ML"; use "eqrule_FOL_data.ML"; setup EQSubstTac.setup subsection {* Other simple lemmas *} lemma [simp]: "((P-->R) <-> (Q-->R)) <-> ((P<->Q) | R)" by blast lemma [simp]: "((P-->Q) <-> (P-->R)) <-> (P --> (Q<->R))" by blast lemma not_disj_iff_imp: "~P | Q <-> (P-->Q)" by blast (** Monotonicity of implications **) lemma conj_mono: "[| P1-->Q1; P2-->Q2 |] ==> (P1&P2) --> (Q1&Q2)" by fast (*or (IntPr.fast_tac 1)*) lemma disj_mono: "[| P1-->Q1; P2-->Q2 |] ==> (P1|P2) --> (Q1|Q2)" by fast (*or (IntPr.fast_tac 1)*) lemma imp_mono: "[| Q1-->P1; P2-->Q2 |] ==> (P1-->P2)-->(Q1-->Q2)" by fast (*or (IntPr.fast_tac 1)*) lemma imp_refl: "P-->P" by (rule impI, assumption) (*The quantifier monotonicity rules are also intuitionistically valid*) lemma ex_mono: "(!!x. P(x) --> Q(x)) ==> (EX x. P(x)) --> (EX x. Q(x))" by blast lemma all_mono: "(!!x. P(x) --> Q(x)) ==> (ALL x. P(x)) --> (ALL x. Q(x))" by blast subsection {* Proof by cases and induction *} text {* Proper handling of non-atomic rule statements. *} constdefs induct_forall :: "('a => o) => o" "induct_forall(P) == ∀x. P(x)" induct_implies :: "o => o => o" "induct_implies(A, B) == A --> B" induct_equal :: "'a => 'a => o" "induct_equal(x, y) == x = y" lemma induct_forall_eq: "(!!x. P(x)) == Trueprop(induct_forall(λx. P(x)))" by (simp only: atomize_all induct_forall_def) lemma induct_implies_eq: "(A ==> B) == Trueprop(induct_implies(A, B))" by (simp only: atomize_imp induct_implies_def) lemma induct_equal_eq: "(x == y) == Trueprop(induct_equal(x, y))" by (simp only: atomize_eq induct_equal_def) lemma induct_impliesI: "(A ==> B) ==> induct_implies(A, B)" by (simp add: induct_implies_def) lemmas induct_atomize = atomize_conj induct_forall_eq induct_implies_eq induct_equal_eq lemmas induct_rulify1 [symmetric, standard] = induct_forall_eq induct_implies_eq induct_equal_eq lemmas induct_rulify2 = induct_forall_def induct_implies_def induct_equal_def lemma all_conj_eq: "(ALL x. P(x)) & (ALL y. Q(y)) == (ALL x y. P(x) & Q(y))" by simp hide const induct_forall induct_implies induct_equal text {* Method setup. *} ML {* structure InductMethod = InductMethodFun (struct val dest_concls = FOLogic.dest_concls; val cases_default = thm "case_split"; val local_impI = thm "induct_impliesI"; val conjI = thm "conjI"; val atomize = thms "induct_atomize"; val rulify1 = thms "induct_rulify1"; val rulify2 = thms "induct_rulify2"; val localize = [Thm.symmetric (thm "induct_implies_def"), Thm.symmetric (thm "atomize_all"), thm "all_conj_eq"]; end); *} setup InductMethod.setup end
theorem ccontr:
(¬ P ==> False) ==> P
theorem disjCI:
(¬ Q ==> P) ==> P ∨ Q
theorem ex_classical:
(¬ (∃x. P(x)) ==> P(a)) ==> ∃x. P(x)
theorem exCI:
(∀x. ¬ P(x) ==> P(a)) ==> ∃x. P(x)
theorem excluded_middle:
¬ P ∨ P
theorem case_split_thm:
[| P ==> Q; ¬ P ==> Q |] ==> Q
theorem impCE:
[| P --> Q; ¬ P ==> R; Q ==> R |] ==> R
theorem impCE':
[| P --> Q; Q ==> R; ¬ P ==> R |] ==> R
theorem notnotD:
¬ ¬ P ==> P
theorem contrapos2:
[| Q; ¬ P ==> ¬ Q |] ==> P
theorem iffCE:
[| P <-> Q; [| P; Q |] ==> R; [| ¬ P; ¬ Q |] ==> R |] ==> R
theorems case_split:
[| P ==> Q; ¬ P ==> Q |] ==> Q
theorems case_split:
[| P ==> Q; ¬ P ==> Q |] ==> Q
theorem imp_elim:
[| P --> Q; ¬ Z_ ==> P; Q ==> Z_ |] ==> Z_
theorem swap:
[| ¬ Pa; ¬ P ==> Pa |] ==> P
theorem alt_ex1E:
[| ∃!x. P(x); !!x. [| P(x); ∀y y'. P(y) ∧ P(y') --> y = y' |] ==> R |] ==> R
lemma ex1_functional:
[| ∃!z. P(a, z); P(a, b); P(a, c) |] ==> b = c
theorem True_implies_equals:
(True ==> PROP P) == PROP P
theorems conj_simps:
P ∧ True <-> P
True ∧ P <-> P
P ∧ False <-> False
False ∧ P <-> False
P ∧ P <-> P
P ∧ P ∧ Q <-> P ∧ Q
P ∧ ¬ P <-> False
¬ P ∧ P <-> False
(P ∧ Q) ∧ R <-> P ∧ Q ∧ R
theorems disj_simps:
P ∨ True <-> True
True ∨ P <-> True
P ∨ False <-> P
False ∨ P <-> P
P ∨ P <-> P
P ∨ P ∨ Q <-> P ∨ Q
(P ∨ Q) ∨ R <-> P ∨ Q ∨ R
theorems not_simps:
¬ (P ∨ Q) <-> ¬ P ∧ ¬ Q
¬ False <-> True
¬ True <-> False
theorems imp_simps:
(P --> False) <-> ¬ P
(P --> True) <-> True
(False --> P) <-> True
(True --> P) <-> P
(P --> P) <-> True
(P --> ¬ P) <-> ¬ P
theorems iff_simps:
(True <-> P) <-> P
(P <-> True) <-> P
(P <-> P) <-> True
(False <-> P) <-> ¬ P
(P <-> False) <-> ¬ P
theorems quant_simps:
(∀x. P) <-> P
(∀x. x = t --> P(x)) <-> P(t)
(∀x. t = x --> P(x)) <-> P(t)
(∃x. P) <-> P
∃x. x = t
∃x. t = x
(∃x. x = t ∧ P(x)) <-> P(t)
(∃x. t = x ∧ P(x)) <-> P(t)
theorems distrib_simps:
P ∧ (Q ∨ R) <-> P ∧ Q ∨ P ∧ R
(Q ∨ R) ∧ P <-> Q ∧ P ∨ R ∧ P
(P ∨ Q --> R) <-> (P --> R) ∧ (Q --> R)
theorem P_iff_F:
¬ P ==> P <-> False
theorem iff_reflection_F:
¬ P ==> P == False
theorem P_iff_T:
P ==> P <-> True
theorem iff_reflection_T:
P ==> P == True
theorem cases_simp:
(P --> Q) ∧ (¬ P --> Q) <-> Q
theorems int_ex_simps:
(∃x. P(x) ∧ Q) <-> (∃x. P(x)) ∧ Q
(∃x. P ∧ Q(x)) <-> P ∧ (∃x. Q(x))
(∃x. P(x) ∨ Q) <-> (∃x. P(x)) ∨ Q
(∃x. P ∨ Q(x)) <-> P ∨ (∃x. Q(x))
theorems cla_ex_simps:
(∃x. P(x) --> Q) <-> (∀x. P(x)) --> Q
(∃x. P --> Q(x)) <-> P --> (∃x. Q(x))
theorems ex_simps:
(∃x. P(x) ∧ Q) <-> (∃x. P(x)) ∧ Q
(∃x. P ∧ Q(x)) <-> P ∧ (∃x. Q(x))
(∃x. P(x) ∨ Q) <-> (∃x. P(x)) ∨ Q
(∃x. P ∨ Q(x)) <-> P ∨ (∃x. Q(x))
(∃x. P(x) --> Q) <-> (∀x. P(x)) --> Q
(∃x. P --> Q(x)) <-> P --> (∃x. Q(x))
theorems int_all_simps:
(∀x. P(x) ∧ Q) <-> (∀x. P(x)) ∧ Q
(∀x. P ∧ Q(x)) <-> P ∧ (∀x. Q(x))
(∀x. P(x) --> Q) <-> (∃x. P(x)) --> Q
(∀x. P --> Q(x)) <-> P --> (∀x. Q(x))
theorems cla_all_simps:
(∀x. P(x) ∨ Q) <-> (∀x. P(x)) ∨ Q
(∀x. P ∨ Q(x)) <-> P ∨ (∀x. Q(x))
theorems all_simps:
(∀x. P(x) ∧ Q) <-> (∀x. P(x)) ∧ Q
(∀x. P ∧ Q(x)) <-> P ∧ (∀x. Q(x))
(∀x. P(x) --> Q) <-> (∃x. P(x)) --> Q
(∀x. P --> Q(x)) <-> P --> (∀x. Q(x))
(∀x. P(x) ∨ Q) <-> (∀x. P(x)) ∨ Q
(∀x. P ∨ Q(x)) <-> P ∨ (∀x. Q(x))
theorem conj_commute:
P ∧ Q <-> Q ∧ P
theorem conj_left_commute:
P ∧ Q ∧ R <-> Q ∧ P ∧ R
theorems conj_comms:
P ∧ Q <-> Q ∧ P
P ∧ Q ∧ R <-> Q ∧ P ∧ R
theorem disj_commute:
P ∨ Q <-> Q ∨ P
theorem disj_left_commute:
P ∨ Q ∨ R <-> Q ∨ P ∨ R
theorems disj_comms:
P ∨ Q <-> Q ∨ P
P ∨ Q ∨ R <-> Q ∨ P ∨ R
theorem conj_disj_distribL:
P ∧ (Q ∨ R) <-> P ∧ Q ∨ P ∧ R
theorem conj_disj_distribR:
(P ∨ Q) ∧ R <-> P ∧ R ∨ Q ∧ R
theorem disj_conj_distribL:
P ∨ Q ∧ R <-> (P ∨ Q) ∧ (P ∨ R)
theorem disj_conj_distribR:
P ∧ Q ∨ R <-> (P ∨ R) ∧ (Q ∨ R)
theorem imp_conj_distrib:
(P --> Q ∧ R) <-> (P --> Q) ∧ (P --> R)
theorem imp_conj:
(P ∧ Q --> R) <-> P --> Q --> R
theorem imp_disj:
(P ∨ Q --> R) <-> (P --> R) ∧ (Q --> R)
theorem imp_disj1:
(P --> Q) ∨ R <-> P --> Q ∨ R
theorem imp_disj2:
Q ∨ (P --> R) <-> P --> Q ∨ R
theorem de_Morgan_disj:
¬ (P ∨ Q) <-> ¬ P ∧ ¬ Q
theorem de_Morgan_conj:
¬ (P ∧ Q) <-> ¬ P ∨ ¬ Q
theorem not_imp:
¬ (P --> Q) <-> P ∧ ¬ Q
theorem not_iff:
¬ (P <-> Q) <-> P <-> ¬ Q
theorem not_all:
¬ (∀x. P(x)) <-> (∃x. ¬ P(x))
theorem imp_all:
((∀x. P(x)) --> Q) <-> (∃x. P(x) --> Q)
theorem not_ex:
¬ (∃x. P(x)) <-> (∀x. ¬ P(x))
theorem imp_ex:
((∃x. P(x)) --> Q) <-> (∀x. P(x) --> Q)
theorem ex_disj_distrib:
(∃x. P(x) ∨ Q(x)) <-> (∃x. P(x)) ∨ (∃x. Q(x))
theorem all_conj_distrib:
(∀x. P(x) ∧ Q(x)) <-> (∀x. P(x)) ∧ (∀x. Q(x))
theorem meta_eq_to_iff:
x == y ==> x <-> y
theorems meta_simps:
(!!x. PROP V) == PROP V
(True ==> PROP P) == PROP P
theorems IFOL_simps:
a = a <-> True
P ∧ True <-> P
True ∧ P <-> P
P ∧ False <-> False
False ∧ P <-> False
P ∧ P <-> P
P ∧ P ∧ Q <-> P ∧ Q
P ∧ ¬ P <-> False
¬ P ∧ P <-> False
(P ∧ Q) ∧ R <-> P ∧ Q ∧ R
P ∨ True <-> True
True ∨ P <-> True
P ∨ False <-> P
False ∨ P <-> P
P ∨ P <-> P
P ∨ P ∨ Q <-> P ∨ Q
(P ∨ Q) ∨ R <-> P ∨ Q ∨ R
¬ (P ∨ Q) <-> ¬ P ∧ ¬ Q
¬ False <-> True
¬ True <-> False
(P --> False) <-> ¬ P
(P --> True) <-> True
(False --> P) <-> True
(True --> P) <-> P
(P --> P) <-> True
(P --> ¬ P) <-> ¬ P
(True <-> P) <-> P
(P <-> True) <-> P
(P <-> P) <-> True
(False <-> P) <-> ¬ P
(P <-> False) <-> ¬ P
(∀x. P) <-> P
(∀x. x = t --> P(x)) <-> P(t)
(∀x. t = x --> P(x)) <-> P(t)
(∃x. P) <-> P
∃x. x = t
∃x. t = x
(∃x. x = t ∧ P(x)) <-> P(t)
(∃x. t = x ∧ P(x)) <-> P(t)
theorem notFalseI:
¬ False
theorems triv_rls:
True
a = a
x == x
P <-> P
¬ False
theorems cla_simps:
¬ (P ∧ Q) <-> ¬ P ∨ ¬ Q
¬ (P ∨ Q) <-> ¬ P ∧ ¬ Q
(P --> Q) ∨ R <-> P --> Q ∨ R
Q ∨ (P --> R) <-> P --> Q ∨ R
¬ (P --> Q) <-> P ∧ ¬ Q
¬ (∀x. P(x)) <-> (∃x. ¬ P(x))
¬ (∃x. P(x)) <-> (∀x. ¬ P(x))
(P --> Q) ∧ (¬ P --> Q) <-> Q
¬ (P ∧ Q) <-> ¬ P ∨ ¬ Q
P ∨ ¬ P
¬ P ∨ P
¬ ¬ P <-> P
(¬ P --> P) <-> P
(¬ P <-> ¬ Q) <-> P <-> Q
lemma
((P --> R) <-> Q --> R) <-> (P <-> Q) ∨ R
lemma
((P --> Q) <-> P --> R) <-> P --> Q <-> R
lemma not_disj_iff_imp:
¬ P ∨ Q <-> P --> Q
lemma conj_mono:
[| P1.0 --> Q1.0; P2.0 --> Q2.0 |] ==> P1.0 ∧ P2.0 --> Q1.0 ∧ Q2.0
lemma disj_mono:
[| P1.0 --> Q1.0; P2.0 --> Q2.0 |] ==> P1.0 ∨ P2.0 --> Q1.0 ∨ Q2.0
lemma imp_mono:
[| Q1.0 --> P1.0; P2.0 --> Q2.0 |] ==> (P1.0 --> P2.0) --> Q1.0 --> Q2.0
lemma imp_refl:
P --> P
lemma ex_mono:
(!!x. P(x) --> Q(x)) ==> (∃x. P(x)) --> (∃x. Q(x))
lemma all_mono:
(!!x. P(x) --> Q(x)) ==> (∀x. P(x)) --> (∀x. Q(x))
lemma induct_forall_eq:
(!!x. P(x)) == induct_forall(P)
lemma induct_implies_eq:
(A ==> B) == induct_implies(A, B)
lemma induct_equal_eq:
(x == y) == induct_equal(x, y)
lemma induct_impliesI:
(A ==> B) ==> induct_implies(A, B)
lemmas induct_atomize:
(A && B) == A ∧ B
(!!x. P(x)) == induct_forall(P)
(A ==> B) == induct_implies(A, B)
(x == y) == induct_equal(x, y)
lemmas induct_atomize:
(A && B) == A ∧ B
(!!x. P(x)) == induct_forall(P)
(A ==> B) == induct_implies(A, B)
(x == y) == induct_equal(x, y)
lemmas induct_rulify1:
induct_forall(P) == (!!x. P(x))
induct_implies(A, B) == (A ==> B)
induct_equal(x, y) == x == y
lemmas induct_rulify1:
induct_forall(P) == (!!x. P(x))
induct_implies(A, B) == (A ==> B)
induct_equal(x, y) == x == y
lemmas induct_rulify2:
induct_forall(P) == ∀x. P(x)
induct_implies(A, B) == A --> B
induct_equal(x, y) == x = y
lemmas induct_rulify2:
induct_forall(P) == ∀x. P(x)
induct_implies(A, B) == A --> B
induct_equal(x, y) == x = y
lemma all_conj_eq:
(∀x. P(x)) ∧ (∀y. Q(y)) == ∀x y. P(x) ∧ Q(y)