Theory FOL

Up to index of Isabelle/FOL

theory FOL
imports IFOL
uses $ISABELLE_HOME/src/Provers/classical.ML $ISABELLE_HOME/src/Provers/blast.ML $ISABELLE_HOME/src/Provers/clasimp.ML $ISABELLE_HOME/src/Tools/induct.ML (cladata.ML) (blastdata.ML) (simpdata.ML)
begin

(*  Title:      FOL/FOL.thy
    ID:         $Id: FOL.thy,v 1.45 2007/10/04 12:42:47 wenzelm Exp $
    Author:     Lawrence C Paulson and Markus Wenzel
*)

header {* Classical first-order logic *}

theory FOL
imports IFOL
uses
  "~~/src/Provers/classical.ML"
  "~~/src/Provers/blast.ML"
  "~~/src/Provers/clasimp.ML"
  "~~/src/Tools/induct.ML"
  ("cladata.ML")
  ("blastdata.ML")
  ("simpdata.ML")
begin


subsection {* The classical axiom *}

axioms
  classical: "(~P ==> P) ==> P"


subsection {* Lemmas and proof tools *}

lemma ccontr: "(¬ P ==> False) ==> P"
  by (erule FalseE [THEN classical])

(*** Classical introduction rules for | and EX ***)

lemma disjCI: "(~Q ==> P) ==> P|Q"
  apply (rule classical)
  apply (assumption | erule meta_mp | rule disjI1 notI)+
  apply (erule notE disjI2)+
  done

(*introduction rule involving only EX*)
lemma ex_classical:
  assumes r: "~(EX x. P(x)) ==> P(a)"
  shows "EX x. P(x)"
  apply (rule classical)
  apply (rule exI, erule r)
  done

(*version of above, simplifying ~EX to ALL~ *)
lemma exCI:
  assumes r: "ALL x. ~P(x) ==> P(a)"
  shows "EX x. P(x)"
  apply (rule ex_classical)
  apply (rule notI [THEN allI, THEN r])
  apply (erule notE)
  apply (erule exI)
  done

lemma excluded_middle: "~P | P"
  apply (rule disjCI)
  apply assumption
  done

(*For disjunctive case analysis*)
ML {*
  fun excluded_middle_tac sP =
    res_inst_tac [("Q",sP)] (@{thm excluded_middle} RS @{thm disjE})
*}

lemma case_split_thm:
  assumes r1: "P ==> Q"
    and r2: "~P ==> Q"
  shows Q
  apply (rule excluded_middle [THEN disjE])
  apply (erule r2)
  apply (erule r1)
  done

lemmas case_split = case_split_thm [case_names True False]

(*HOL's more natural case analysis tactic*)
ML {*
  fun case_tac a = res_inst_tac [("P",a)] @{thm case_split_thm}
*}


(*** Special elimination rules *)


(*Classical implies (-->) elimination. *)
lemma impCE:
  assumes major: "P-->Q"
    and r1: "~P ==> R"
    and r2: "Q ==> R"
  shows R
  apply (rule excluded_middle [THEN disjE])
   apply (erule r1)
  apply (rule r2)
  apply (erule major [THEN mp])
  done

(*This version of --> elimination works on Q before P.  It works best for
  those cases in which P holds "almost everywhere".  Can't install as
  default: would break old proofs.*)
lemma impCE':
  assumes major: "P-->Q"
    and r1: "Q ==> R"
    and r2: "~P ==> R"
  shows R
  apply (rule excluded_middle [THEN disjE])
   apply (erule r2)
  apply (rule r1)
  apply (erule major [THEN mp])
  done

(*Double negation law*)
lemma notnotD: "~~P ==> P"
  apply (rule classical)
  apply (erule notE)
  apply assumption
  done

lemma contrapos2:  "[| Q; ~ P ==> ~ Q |] ==> P"
  apply (rule classical)
  apply (drule (1) meta_mp)
  apply (erule (1) notE)
  done

(*** Tactics for implication and contradiction ***)

(*Classical <-> elimination.  Proof substitutes P=Q in 
    ~P ==> ~Q    and    P ==> Q  *)
lemma iffCE:
  assumes major: "P<->Q"
    and r1: "[| P; Q |] ==> R"
    and r2: "[| ~P; ~Q |] ==> R"
  shows R
  apply (rule major [unfolded iff_def, THEN conjE])
  apply (elim impCE)
     apply (erule (1) r2)
    apply (erule (1) notE)+
  apply (erule (1) r1)
  done


(*Better for fast_tac: needs no quantifier duplication!*)
lemma alt_ex1E:
  assumes major: "EX! x. P(x)"
    and r: "!!x. [| P(x);  ALL y y'. P(y) & P(y') --> y=y' |] ==> R"
  shows R
  using major
proof (rule ex1E)
  fix x
  assume * : "∀y. P(y) --> y = x"
  assume "P(x)"
  then show R
  proof (rule r)
    { fix y y'
      assume "P(y)" and "P(y')"
      with * have "x = y" and "x = y'" by - (tactic "IntPr.fast_tac 1")+
      then have "y = y'" by (rule subst)
    } note r' = this
    show "∀y y'. P(y) ∧ P(y') --> y = y'" by (intro strip, elim conjE) (rule r')
  qed
qed

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

(* Elimination of True from asumptions: *)
lemma True_implies_equals: "(True ==> PROP P) == PROP P"
proof
  assume "True ==> PROP P"
  from this and TrueI show "PROP P" .
next
  assume "PROP P"
  then show "PROP P" .
qed

lemma uncurry: "P --> Q --> R ==> P & Q --> R"
  by blast

lemma iff_allI: "(!!x. P(x) <-> Q(x)) ==> (ALL x. P(x)) <-> (ALL x. Q(x))"
  by blast

lemma iff_exI: "(!!x. P(x) <-> Q(x)) ==> (EX x. P(x)) <-> (EX x. Q(x))"
  by blast

lemma all_comm: "(ALL x y. P(x,y)) <-> (ALL y x. P(x,y))" by blast

lemma ex_comm: "(EX x y. P(x,y)) <-> (EX y x. P(x,y))" by blast

use "simpdata.ML"
setup simpsetup
setup "Simplifier.method_setup Splitter.split_modifiers"
setup Splitter.setup
setup Clasimp.setup
setup EqSubst.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 where "induct_forall(P) == ∀x. P(x)"
  induct_implies where "induct_implies(A, B) == A --> B"
  induct_equal where "induct_equal(x, y) == x = y"
  induct_conj where "induct_conj(A, B) == A ∧ B"

lemma induct_forall_eq: "(!!x. P(x)) == Trueprop(induct_forall(λx. P(x)))"
  unfolding atomize_all induct_forall_def .

lemma induct_implies_eq: "(A ==> B) == Trueprop(induct_implies(A, B))"
  unfolding atomize_imp induct_implies_def .

lemma induct_equal_eq: "(x == y) == Trueprop(induct_equal(x, y))"
  unfolding atomize_eq induct_equal_def .

lemma induct_conj_eq:
  includes meta_conjunction_syntax
  shows "(A && B) == Trueprop(induct_conj(A, B))"
  unfolding atomize_conj induct_conj_def .

lemmas induct_atomize = induct_forall_eq induct_implies_eq induct_equal_eq induct_conj_eq
lemmas induct_rulify [symmetric, standard] = induct_atomize
lemmas induct_rulify_fallback =
  induct_forall_def induct_implies_def induct_equal_def induct_conj_def

hide const induct_forall induct_implies induct_equal induct_conj


text {* Method setup. *}

ML {*
  structure Induct = InductFun
  (
    val cases_default = @{thm case_split}
    val atomize = @{thms induct_atomize}
    val rulify = @{thms induct_rulify}
    val rulify_fallback = @{thms induct_rulify_fallback}
  );
*}

setup Induct.setup
declare case_split [cases type: o]

end

The classical axiom

Lemmas and proof tools

lemma ccontr:

  P ==> False) ==> P

lemma disjCI:

  Q ==> P) ==> PQ

lemma ex_classical:

  (¬ (∃x. P(x)) ==> P(a)) ==> ∃x. P(x)

lemma exCI:

  (∀x. ¬ P(x) ==> P(a)) ==> ∃x. P(x)

lemma excluded_middle:

  ¬ PP

lemma case_split_thm:

  [| P ==> Q; ¬ P ==> Q |] ==> Q

lemma case_split:

  [| P ==> Q; ¬ P ==> Q |] ==> Q

lemma impCE:

  [| P --> Q; ¬ P ==> R; Q ==> R |] ==> R

lemma impCE':

  [| P --> Q; Q ==> R; ¬ P ==> R |] ==> R

lemma notnotD:

  ¬ ¬ P ==> P

lemma contrapos2:

  [| Q; ¬ P ==> ¬ Q |] ==> P

lemma iffCE:

  [| P <-> Q; [| P; Q |] ==> R; [| ¬ P; ¬ Q |] ==> R |] ==> R

lemma alt_ex1E:

  [| ∃!x. P(x); !!x. [| P(x); ∀y y'. P(y) ∧ P(y') --> y = y' |] ==> R |] ==> R

Classical Reasoner

theorem imp_elim:

  [| Pa --> Q; ¬ P ==> Pa; Q ==> P |] ==> P

theorem swap:

  [| ¬ Pa; ¬ P ==> Pa |] ==> P

lemma ex1_functional:

  [| ∃!z. P(a, z); P(a, b); P(a, c) |] ==> b = c

lemma True_implies_equals:

  (True ==> PROP P) == PROP P

lemma uncurry:

  P --> Q --> R ==> PQ --> R

lemma iff_allI:

  (!!x. P(x) <-> Q(x)) ==> (∀x. P(x)) <-> (∀x. Q(x))

lemma iff_exI:

  (!!x. P(x) <-> Q(x)) ==> (∃x. P(x)) <-> (∃x. Q(x))

lemma all_comm:

  (∀x y. P(x, y)) <-> (∀y x. P(x, y))

lemma ex_comm:

  (∃x y. P(x, y)) <-> (∃y x. P(x, y))

theorem conj_simps:

  P ∧ True <-> P
  True ∧ P <-> P
  P ∧ False <-> False
  False ∧ P <-> False
  PP <-> P
  PPQ <-> PQ
  P ∧ ¬ P <-> False
  ¬ PP <-> False
  (PQ) ∧ R <-> PQR

theorem disj_simps:

  P ∨ True <-> True
  True ∨ P <-> True
  P ∨ False <-> P
  False ∨ P <-> P
  PP <-> P
  PPQ <-> PQ
  (PQ) ∨ R <-> PQR

theorem not_simps:

  ¬ (PQ) <-> ¬ P ∧ ¬ Q
  ¬ False <-> True
  ¬ True <-> False

theorem imp_simps:

  (P --> False) <-> ¬ P
  (P --> True) <-> True
  (False --> P) <-> True
  (True --> P) <-> P
  (P --> P) <-> True
  (P --> ¬ P) <-> ¬ P

theorem iff_simps:

  (True <-> P) <-> P
  (P <-> True) <-> P
  (P <-> P) <-> True
  (False <-> P) <-> ¬ P
  (P <-> False) <-> ¬ P

theorem 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 = tP(x)) <-> P(t)
  (∃x. t = xP(x)) <-> P(t)

theorem distrib_simps:

  P ∧ (QR) <-> PQPR
  (QR) ∧ P <-> QPRP
  (PQ --> 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

theorem int_ex_simps:

  (∃x. P(x) ∧ Q) <-> (∃x. P(x)) ∧ Q
  (∃x. PQ(x)) <-> P ∧ (∃x. Q(x))
  (∃x. P(x) ∨ Q) <-> (∃x. P(x)) ∨ Q
  (∃x. PQ(x)) <-> P ∨ (∃x. Q(x))

theorem cla_ex_simps:

  (∃x. P(x) --> Q) <-> (∀x. P(x)) --> Q
  (∃x. P --> Q(x)) <-> P --> (∃x. Q(x))

theorem ex_simps:

  (∃x. P(x) ∧ Q) <-> (∃x. P(x)) ∧ Q
  (∃x. PQ(x)) <-> P ∧ (∃x. Q(x))
  (∃x. P(x) ∨ Q) <-> (∃x. P(x)) ∨ Q
  (∃x. PQ(x)) <-> P ∨ (∃x. Q(x))
  (∃x. P(x) --> Q) <-> (∀x. P(x)) --> Q
  (∃x. P --> Q(x)) <-> P --> (∃x. Q(x))

theorem int_all_simps:

  (∀x. P(x) ∧ Q) <-> (∀x. P(x)) ∧ Q
  (∀x. PQ(x)) <-> P ∧ (∀x. Q(x))
  (∀x. P(x) --> Q) <-> (∃x. P(x)) --> Q
  (∀x. P --> Q(x)) <-> P --> (∀x. Q(x))

theorem cla_all_simps:

  (∀x. P(x) ∨ Q) <-> (∀x. P(x)) ∨ Q
  (∀x. PQ(x)) <-> P ∨ (∀x. Q(x))

theorem all_simps:

  (∀x. P(x) ∧ Q) <-> (∀x. P(x)) ∧ Q
  (∀x. PQ(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. PQ(x)) <-> P ∨ (∀x. Q(x))

theorem conj_commute:

  PQ <-> QP

theorem conj_left_commute:

  PQR <-> QPR

theorem conj_comms:

  PQ <-> QP
  PQR <-> QPR

theorem disj_commute:

  PQ <-> QP

theorem disj_left_commute:

  PQR <-> QPR

theorem disj_comms:

  PQ <-> QP
  PQR <-> QPR

theorem conj_disj_distribL:

  P ∧ (QR) <-> PQPR

theorem conj_disj_distribR:

  (PQ) ∧ R <-> PRQR

theorem disj_conj_distribL:

  PQR <-> (PQ) ∧ (PR)

theorem disj_conj_distribR:

  PQR <-> (PR) ∧ (QR)

theorem imp_conj_distrib:

  (P --> QR) <-> (P --> Q) ∧ (P --> R)

theorem imp_conj:

  (PQ --> R) <-> P --> Q --> R

theorem imp_disj:

  (PQ --> R) <-> (P --> R) ∧ (Q --> R)

theorem imp_disj1:

  (P --> Q) ∨ R <-> P --> QR

theorem imp_disj2:

  Q ∨ (P --> R) <-> P --> QR

theorem de_Morgan_disj:

  ¬ (PQ) <-> ¬ P ∧ ¬ Q

theorem de_Morgan_conj:

  ¬ (PQ) <-> ¬ 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_simps:

  (!!x. PROP V) == PROP V
  (True ==> PROP P) == PROP P

theorem IFOL_simps:

  a = a <-> True
  P ∧ True <-> P
  True ∧ P <-> P
  P ∧ False <-> False
  False ∧ P <-> False
  PP <-> P
  PPQ <-> PQ
  P ∧ ¬ P <-> False
  ¬ PP <-> False
  (PQ) ∧ R <-> PQR
  P ∨ True <-> True
  True ∨ P <-> True
  P ∨ False <-> P
  False ∨ P <-> P
  PP <-> P
  PPQ <-> PQ
  (PQ) ∨ R <-> PQR
  ¬ (PQ) <-> ¬ 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 = tP(x)) <-> P(t)
  (∃x. t = xP(x)) <-> P(t)

theorem notFalseI:

  ¬ False

theorem triv_rls:

  True
  a = a
  x == x
  P <-> P
  ¬ False

theorem cla_simps:

  ¬ (PQ) <-> ¬ P ∨ ¬ Q
  ¬ (PQ) <-> ¬ P ∧ ¬ Q
  (P --> Q) ∨ R <-> P --> QR
  Q ∨ (P --> R) <-> P --> QR
  ¬ (P --> Q) <-> P ∧ ¬ Q
  ¬ (∀x. P(x)) <-> (∃x. ¬ P(x))
  ¬ (∃x. P(x)) <-> (∀x. ¬ P(x))
  (P --> Q) ∧ (¬ P --> Q) <-> Q
  ¬ (PQ) <-> ¬ P ∨ ¬ Q
  P ∨ ¬ P
  ¬ PP
  ¬ ¬ P <-> P
  P --> P) <-> P
  P <-> ¬ Q) <-> P <-> Q

Other simple lemmas

lemma

  ((P --> R) <-> Q --> R) <-> (P <-> Q) ∨ R

lemma

  ((P --> Q) <-> P --> R) <-> P --> Q <-> R

lemma not_disj_iff_imp:

  ¬ PQ <-> P --> Q

lemma conj_mono:

  [| P1.0 --> Q1.0; P2.0 --> Q2.0 |] ==> P1.0P2.0 --> Q1.0Q2.0

lemma disj_mono:

  [| P1.0 --> Q1.0; P2.0 --> Q2.0 |] ==> P1.0P2.0 --> Q1.0Q2.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))

Proof by cases and induction

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_conj_eq:

  (A && B) == induct_conj(A, B)

lemma induct_atomize:

  (!!x. P(x)) == induct_forall(P)
  (A ==> B) == induct_implies(A, B)
  (x == y) == induct_equal(x, y)
  (A && B) == induct_conj(A, B)

lemma induct_rulify:

  induct_forall(P) == (!!x. P(x))
  induct_implies(A, B) == (A ==> B)
  induct_equal(x, y) == x == y
  induct_conj(A, B) == A && B

lemma induct_rulify_fallback:

  induct_forall(P) == ∀x. P(x)
  induct_implies(A, B) == A --> B
  induct_equal(x, y) == x = y
  induct_conj(A, B) == AB