doc-src/ZF/FOL_examples.thy
author kleing
Mon, 29 Dec 2003 06:49:26 +0100
changeset 14333 14f29eb097a3
parent 14152 12f6f18e7afc
child 16417 9bc16273c2d4
permissions -rw-r--r--
\<^bsub> .. \<^esub>

header{*Examples of Classical Reasoning*}

theory FOL_examples = FOL:

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