| author | wenzelm | 
| Mon, 28 Sep 2020 22:22:56 +0200 | |
| changeset 72323 | e36f94e2eb6b | 
| parent 69505 | cc2d676d5395 | 
| permissions | -rw-r--r-- | 
| 59720 | 1 | (* Title: Doc/Logics_ZF/If.thy | 
| 14205 | 2 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | 
| 3 | Copyright 1991 University of Cambridge | |
| 4 | ||
| 5 | First-Order Logic: the 'if' example. | |
| 6 | *) | |
| 7 | ||
| 66453 
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
 wenzelm parents: 
59720diff
changeset | 8 | theory If imports FOL begin | 
| 14205 | 9 | |
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
19792diff
changeset | 10 | definition "if" :: "[o,o,o]=>o" where | 
| 19792 | 11 | "if(P,Q,R) == P&Q | ~P&R" | 
| 14205 | 12 | |
| 13 | lemma ifI: | |
| 14 | "[| P ==> Q; ~P ==> R |] ==> if(P,Q,R)" | |
| 67443 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 wenzelm parents: 
67406diff
changeset | 15 |   \<comment> \<open>@{subgoals[display,indent=0,margin=65]}\<close>
 | 
| 14205 | 16 | apply (simp add: if_def) | 
| 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>
 | 
| 14205 | 18 | apply blast | 
| 19 | done | |
| 20 | ||
| 21 | lemma ifE: | |
| 22 | "[| if(P,Q,R); [| P; Q |] ==> S; [| ~P; R |] ==> S |] ==> S" | |
| 67443 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 wenzelm parents: 
67406diff
changeset | 23 |   \<comment> \<open>@{subgoals[display,indent=0,margin=65]}\<close>
 | 
| 14205 | 24 | apply (simp add: if_def) | 
| 67443 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 wenzelm parents: 
67406diff
changeset | 25 |   \<comment> \<open>@{subgoals[display,indent=0,margin=65]}\<close>
 | 
| 14205 | 26 | apply blast | 
| 27 | done | |
| 28 | ||
| 29 | lemma if_commute: "if(P, if(Q,A,B), if(Q,C,D)) <-> if(Q, if(P,A,C), if(P,B,D))" | |
| 67443 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 wenzelm parents: 
67406diff
changeset | 30 |   \<comment> \<open>@{subgoals[display,indent=0,margin=65]}\<close>
 | 
| 14205 | 31 | apply (rule iffI) | 
| 67443 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 wenzelm parents: 
67406diff
changeset | 32 |   \<comment> \<open>@{subgoals[display,indent=0,margin=65]}\<close>
 | 
| 14205 | 33 | apply (erule ifE) | 
| 67443 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 wenzelm parents: 
67406diff
changeset | 34 |   \<comment> \<open>@{subgoals[display,indent=0,margin=65]}\<close>
 | 
| 14205 | 35 | apply (erule ifE) | 
| 67443 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 wenzelm parents: 
67406diff
changeset | 36 |   \<comment> \<open>@{subgoals[display,indent=0,margin=65]}\<close>
 | 
| 14205 | 37 | apply (rule ifI) | 
| 67443 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 wenzelm parents: 
67406diff
changeset | 38 |   \<comment> \<open>@{subgoals[display,indent=0,margin=65]}\<close>
 | 
| 14205 | 39 | apply (rule ifI) | 
| 67443 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 wenzelm parents: 
67406diff
changeset | 40 |   \<comment> \<open>@{subgoals[display,indent=0,margin=65]}\<close>
 | 
| 14205 | 41 | oops | 
| 42 | ||
| 69505 | 43 | text\<open>Trying again from the beginning in order to use \<open>blast\<close>\<close> | 
| 14205 | 44 | declare ifI [intro!] | 
| 45 | declare ifE [elim!] | |
| 46 | ||
| 47 | lemma if_commute: "if(P, if(Q,A,B), if(Q,C,D)) <-> if(Q, if(P,A,C), if(P,B,D))" | |
| 48 | by blast | |
| 49 | ||
| 50 | ||
| 51 | lemma "if(if(P,Q,R), A, B) <-> if(P, if(Q,A,B), if(R,A,B))" | |
| 67443 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 wenzelm parents: 
67406diff
changeset | 52 |   \<comment> \<open>@{subgoals[display,indent=0,margin=65]}\<close>
 | 
| 14205 | 53 | by blast | 
| 54 | ||
| 67406 | 55 | text\<open>Trying again from the beginning in order to prove from the definitions\<close> | 
| 14205 | 56 | lemma "if(if(P,Q,R), A, B) <-> if(P, if(Q,A,B), if(R,A,B))" | 
| 67443 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 wenzelm parents: 
67406diff
changeset | 57 |   \<comment> \<open>@{subgoals[display,indent=0,margin=65]}\<close>
 | 
| 14205 | 58 | apply (simp add: if_def) | 
| 67443 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 wenzelm parents: 
67406diff
changeset | 59 |   \<comment> \<open>@{subgoals[display,indent=0,margin=65]}\<close>
 | 
| 14205 | 60 | apply blast | 
| 61 | done | |
| 62 | ||
| 63 | ||
| 67406 | 64 | text\<open>An invalid formula. High-level rules permit a simpler diagnosis\<close> | 
| 14205 | 65 | lemma "if(if(P,Q,R), A, B) <-> if(P, if(Q,A,B), if(R,B,A))" | 
| 67443 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 wenzelm parents: 
67406diff
changeset | 66 |   \<comment> \<open>@{subgoals[display,indent=0,margin=65]}\<close>
 | 
| 14205 | 67 | apply auto | 
| 67443 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 wenzelm parents: 
67406diff
changeset | 68 |   \<comment> \<open>@{subgoals[display,indent=0,margin=65]}\<close>
 | 
| 14205 | 69 | (*The next step will fail unless subgoals remain*) | 
| 70 | apply (tactic all_tac) | |
| 71 | oops | |
| 72 | ||
| 67406 | 73 | text\<open>Trying again from the beginning in order to prove from the definitions\<close> | 
| 14205 | 74 | lemma "if(if(P,Q,R), A, B) <-> if(P, if(Q,A,B), if(R,B,A))" | 
| 67443 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 wenzelm parents: 
67406diff
changeset | 75 |   \<comment> \<open>@{subgoals[display,indent=0,margin=65]}\<close>
 | 
| 14205 | 76 | apply (simp add: if_def) | 
| 67443 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 wenzelm parents: 
67406diff
changeset | 77 |   \<comment> \<open>@{subgoals[display,indent=0,margin=65]}\<close>
 | 
| 14205 | 78 | apply (auto) | 
| 67443 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 wenzelm parents: 
67406diff
changeset | 79 |   \<comment> \<open>@{subgoals[display,indent=0,margin=65]}\<close>
 | 
| 14205 | 80 | (*The next step will fail unless subgoals remain*) | 
| 81 | apply (tactic all_tac) | |
| 82 | oops | |
| 83 | ||
| 84 | ||
| 85 | end |