| author | wenzelm | 
| Sun, 13 Jan 2019 18:40:26 +0100 | |
| changeset 69645 | e4e5bc6ac214 | 
| parent 69505 | cc2d676d5395 | 
| permissions | -rw-r--r-- | 
| 67406 | 1 | section\<open>Examples of Classical Reasoning\<close> | 
| 14152 | 2 | |
| 66453 
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
 wenzelm parents: 
58889diff
changeset | 3 | theory FOL_examples imports FOL begin | 
| 14152 | 4 | |
| 5 | lemma "EX y. ALL x. P(y)-->P(x)" | |
| 67443 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 wenzelm parents: 
67406diff
changeset | 6 |   \<comment> \<open>@{subgoals[display,indent=0,margin=65]}\<close>
 | 
| 14152 | 7 | apply (rule exCI) | 
| 67443 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 wenzelm parents: 
67406diff
changeset | 8 |   \<comment> \<open>@{subgoals[display,indent=0,margin=65]}\<close>
 | 
| 14152 | 9 | apply (rule allI) | 
| 67443 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 wenzelm parents: 
67406diff
changeset | 10 |   \<comment> \<open>@{subgoals[display,indent=0,margin=65]}\<close>
 | 
| 14152 | 11 | apply (rule impI) | 
| 67443 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 wenzelm parents: 
67406diff
changeset | 12 |   \<comment> \<open>@{subgoals[display,indent=0,margin=65]}\<close>
 | 
| 14152 | 13 | apply (erule allE) | 
| 67443 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 wenzelm parents: 
67406diff
changeset | 14 |   \<comment> \<open>@{subgoals[display,indent=0,margin=65]}\<close>
 | 
| 69505 | 15 | txt\<open>see below for \<open>allI\<close> combined with \<open>swap\<close>\<close> | 
| 14152 | 16 | apply (erule allI [THEN [2] swap]) | 
| 67443 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 wenzelm parents: 
67406diff
changeset | 17 |   \<comment> \<open>@{subgoals[display,indent=0,margin=65]}\<close>
 | 
| 14152 | 18 | apply (rule impI) | 
| 67443 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 wenzelm parents: 
67406diff
changeset | 19 |   \<comment> \<open>@{subgoals[display,indent=0,margin=65]}\<close>
 | 
| 14152 | 20 | apply (erule notE) | 
| 67443 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 wenzelm parents: 
67406diff
changeset | 21 |   \<comment> \<open>@{subgoals[display,indent=0,margin=65]}\<close>
 | 
| 14152 | 22 | apply assumption | 
| 23 | done | |
| 24 | ||
| 67406 | 25 | text \<open> | 
| 14152 | 26 | @{thm[display] allI [THEN [2] swap]}
 | 
| 67406 | 27 | \<close> | 
| 14152 | 28 | |
| 29 | lemma "EX y. ALL x. P(y)-->P(x)" | |
| 30 | by blast | |
| 31 | ||
| 32 | end | |
| 33 | ||
| 34 |