doc-src/ZF/FOL_examples.thy
author wenzelm
Mon, 09 Mar 2009 22:43:25 +0100
changeset 30399 a4772a650b4e
parent 16417 9bc16273c2d4
child 48517 0f8c8ac6cc0e
permissions -rw-r--r--
tuned;

header{*Examples of Classical Reasoning*}

theory FOL_examples imports FOL begin

lemma "EX y. ALL x. P(y)-->P(x)"
  --{* @{subgoals[display,indent=0,margin=65]} *}
apply (rule exCI)
  --{* @{subgoals[display,indent=0,margin=65]} *}
apply (rule allI)
  --{* @{subgoals[display,indent=0,margin=65]} *}
apply (rule impI)
  --{* @{subgoals[display,indent=0,margin=65]} *}
apply (erule allE)
  --{* @{subgoals[display,indent=0,margin=65]} *}
txt{*see below for @{text allI} combined with @{text swap}*}
apply (erule allI [THEN [2] swap])
  --{* @{subgoals[display,indent=0,margin=65]} *}
apply (rule impI)
  --{* @{subgoals[display,indent=0,margin=65]} *}
apply (erule notE)
  --{* @{subgoals[display,indent=0,margin=65]} *}
apply assumption
done

text {*
@{thm[display] allI [THEN [2] swap]}
*}

lemma "EX y. ALL x. P(y)-->P(x)"
by blast

end