Theory Quantifiers_Cla

(*  Title:      FOL/ex/Quantifiers_Cla.thy
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
    Copyright   1991  University of Cambridge
*)

section ‹First-Order Logic: quantifier examples (classical version)›

theory Quantifiers_Cla
imports FOL
begin

lemma (x y. P(x,y))  (y x. P(x,y))
  by fast

lemma (x y. P(x,y))  (y x. P(x,y))
  by fast


text ‹Converse is false.›
lemma (x. P(x))  (x. Q(x))  (x. P(x)  Q(x))
  by fast

lemma (x. P  Q(x))  (P  (x. Q(x)))
  by fast


lemma (x. P(x)  Q)  ((x. P(x))  Q)
  by fast


text ‹Some harder ones.›

lemma (x. P(x)  Q(x))  (x. P(x))  (x. Q(x))
  by fast

― ‹Converse is false.›
lemma (x. P(x)  Q(x))  (x. P(x))  (x. Q(x))
  by fast


text ‹Basic test of quantifier reasoning.›

― ‹TRUE›
lemma (y. x. Q(x,y))  (x. y. Q(x,y))
  by fast

lemma (x. Q(x))  (x. Q(x))
  by fast


text ‹The following should fail, as they are false!›

lemma (x. y. Q(x,y))  (y. x. Q(x,y))
  apply fast?
  oops

lemma (x. Q(x))  (x. Q(x))
  apply fast?
  oops

schematic_goal P(?a)  (x. P(x))
  apply fast?
  oops

schematic_goal (P(?a)  (x. Q(x)))  (x. P(x)  Q(x))
  apply fast?
  oops


text ‹Back to things that are provable \dots›

lemma (x. P(x)  Q(x))  (x. P(x))  (x. Q(x))
  by fast

text ‹An example of why exI› should be delayed as long as possible.›
lemma (P  (x. Q(x)))  P  (x. Q(x))
  by fast

schematic_goal (x. P(x)  Q(f(x)))  (x. Q(x)  R(g(x)))  P(d)  R(?a)
  by fast

lemma (x. Q(x))  (x. Q(x))
  by fast


text ‹Some slow ones›

text ‹Principia Mathematica *11.53›
lemma (x y. P(x)  Q(y))  ((x. P(x))  (y. Q(y)))
  by fast

(*Principia Mathematica *11.55  *)
lemma (x y. P(x)  Q(x,y))  (x. P(x)  (y. Q(x,y)))
  by fast

(*Principia Mathematica *11.61  *)
lemma (y. x. P(x)  Q(x,y))  (x. P(x)  (y. Q(x,y)))
  by fast

end