src/Doc/Logics_ZF/FOL_examples.thy
author wenzelm
Wed, 21 Mar 2018 17:55:17 +0100
changeset 67912 a7731d581bbc
parent 67443 3abf6a722518
child 69505 cc2d676d5395
permissions -rw-r--r--
clarified result;

section\<open>Examples of Classical Reasoning\<close>

theory FOL_examples imports FOL begin

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

text \<open>
@{thm[display] allI [THEN [2] swap]}
\<close>

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

end