src/Doc/Logics_ZF/FOL_examples.thy
author wenzelm
Tue, 03 Dec 2019 16:40:04 +0100
changeset 71222 2bc39c80a95d
parent 69505 cc2d676d5395
permissions -rw-r--r--
clarified export of consts: recursion is accessible via spec_rules;

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 \<open>allI\<close> combined with \<open>swap\<close>\<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