| author | wenzelm | 
| Sat, 24 Jul 2010 21:40:48 +0200 | |
| changeset 37951 | 4e2aaf080572 | 
| parent 35416 | d8d7d1b785af | 
| child 41777 | 1f7cbe39d425 | 
| permissions | -rw-r--r-- | 
| 14151 | 1 | (* Title: FOL/ex/If.ML | 
| 2 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | |
| 3 | Copyright 1991 University of Cambridge | |
| 14957 | 4 | *) | 
| 14151 | 5 | |
| 14957 | 6 | header {* First-Order Logic: the 'if' example *}
 | 
| 14151 | 7 | |
| 16417 | 8 | theory If imports FOL begin | 
| 14151 | 9 | |
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
31974diff
changeset | 10 | definition "if" :: "[o,o,o]=>o" where | 
| 19796 | 11 | "if(P,Q,R) == P&Q | ~P&R" | 
| 14151 | 12 | |
| 13 | lemma ifI: | |
| 14 | "[| P ==> Q; ~P ==> R |] ==> if(P,Q,R)" | |
| 15 | apply (simp add: if_def, blast) | |
| 16 | done | |
| 17 | ||
| 18 | lemma ifE: | |
| 19 | "[| if(P,Q,R); [| P; Q |] ==> S; [| ~P; R |] ==> S |] ==> S" | |
| 20 | apply (simp add: if_def, blast) | |
| 21 | done | |
| 22 | ||
| 23 | lemma if_commute: "if(P, if(Q,A,B), if(Q,C,D)) <-> if(Q, if(P,A,C), if(P,B,D))" | |
| 24 | apply (rule iffI) | |
| 25 | apply (erule ifE) | |
| 26 | apply (erule ifE) | |
| 27 | apply (rule ifI) | |
| 28 | apply (rule ifI) | |
| 29 | oops | |
| 30 | ||
| 31 | text{*Trying again from the beginning in order to use @{text blast}*}
 | |
| 32 | declare ifI [intro!] | |
| 33 | declare ifE [elim!] | |
| 34 | ||
| 35 | lemma if_commute: "if(P, if(Q,A,B), if(Q,C,D)) <-> if(Q, if(P,A,C), if(P,B,D))" | |
| 36 | by blast | |
| 37 | ||
| 38 | ||
| 39 | lemma "if(if(P,Q,R), A, B) <-> if(P, if(Q,A,B), if(R,A,B))" | |
| 40 | by blast | |
| 41 | ||
| 42 | text{*Trying again from the beginning in order to prove from the definitions*}
 | |
| 43 | lemma "if(if(P,Q,R), A, B) <-> if(P, if(Q,A,B), if(R,A,B))" | |
| 44 | by (simp add: if_def, blast) | |
| 45 | ||
| 46 | ||
| 47 | text{*An invalid formula.  High-level rules permit a simpler diagnosis*}
 | |
| 48 | lemma "if(if(P,Q,R), A, B) <-> if(P, if(Q,A,B), if(R,B,A))" | |
| 49 | apply auto | |
| 14957 | 50 |   -- {*The next step will fail unless subgoals remain*}
 | 
| 14151 | 51 | apply (tactic all_tac) | 
| 52 | oops | |
| 53 | ||
| 54 | text{*Trying again from the beginning in order to prove from the definitions*}
 | |
| 55 | lemma "if(if(P,Q,R), A, B) <-> if(P, if(Q,A,B), if(R,B,A))" | |
| 56 | apply (simp add: if_def, auto) | |
| 14957 | 57 |   -- {*The next step will fail unless subgoals remain*}
 | 
| 14151 | 58 | apply (tactic all_tac) | 
| 59 | oops | |
| 60 | ||
| 0 | 61 | end |