| author | wenzelm | 
| Thu, 13 Jul 2000 23:09:25 +0200 | |
| changeset 9311 | ab5b24cbaa16 | 
| parent 9245 | 428385c4bc50 | 
| child 10834 | a7897aebbffc | 
| permissions | -rw-r--r-- | 
| 2640 | 1 | (* Title: HOLCF/Tr.ML | 
| 2 | ID: $Id$ | |
| 3 | Author: Franz Regensburger | |
| 4 | Copyright 1993 Technische Universitaet Muenchen | |
| 5 | ||
| 9245 | 6 | Introduce infix if_then_else_fi and boolean connectives andalso, orelse | 
| 2640 | 7 | *) | 
| 8 | ||
| 9 | (* ------------------------------------------------------------------------ *) | |
| 10 | (* Exhaustion and Elimination for type one *) | |
| 11 | (* ------------------------------------------------------------------------ *) | |
| 9245 | 12 | |
| 13 | Goalw [FF_def,TT_def] "t=UU | t = TT | t = FF"; | |
| 14 | by (lift.induct_tac "t" 1); | |
| 15 | by (fast_tac HOL_cs 1); | |
| 16 | by (fast_tac (HOL_cs addss simpset()) 1); | |
| 17 | qed "Exh_tr"; | |
| 2640 | 18 | |
| 9169 | 19 | val prems = Goal "[| p=UU ==> Q; p = TT ==>Q; p = FF ==>Q|] ==>Q"; | 
| 20 | by (rtac (Exh_tr RS disjE) 1); | |
| 21 | by (eresolve_tac prems 1); | |
| 22 | by (etac disjE 1); | |
| 23 | by (eresolve_tac prems 1); | |
| 24 | by (eresolve_tac prems 1); | |
| 25 | qed "trE"; | |
| 2640 | 26 | |
| 27 | (* ------------------------------------------------------------------------ *) | |
| 28 | (* tactic for tr-thms with case split *) | |
| 29 | (* ------------------------------------------------------------------------ *) | |
| 30 | ||
| 7654 | 31 | bind_thms ("tr_defs", [andalso_def,orelse_def,neg_def,ifte_def,TT_def,FF_def]);
 | 
| 2640 | 32 | |
| 33 | fun prover t = prove_goal thy t | |
| 34 | (fn prems => | |
| 35 | [ | |
| 36 |         (res_inst_tac [("p","y")] trE 1),
 | |
| 4098 | 37 | (REPEAT(asm_simp_tac (simpset() addsimps | 
| 2640 | 38 | [o_def,flift1_def,flift2_def,inst_lift_po]@tr_defs) 1)) | 
| 39 | ]); | |
| 40 | ||
| 41 | (* ------------------------------------------------------------------------ *) | |
| 42 | (* distinctness for type tr *) | |
| 43 | (* ------------------------------------------------------------------------ *) | |
| 44 | ||
| 7654 | 45 | bind_thms ("dist_less_tr", map prover [
 | 
| 2640 | 46 | "~TT << UU", | 
| 47 | "~FF << UU", | |
| 48 | "~TT << FF", | |
| 49 | "~FF << TT" | |
| 7654 | 50 | ]); | 
| 2640 | 51 | |
| 52 | val dist_eq_tr = map prover ["TT~=UU","FF~=UU","TT~=FF"]; | |
| 7654 | 53 | bind_thms ("dist_eq_tr", dist_eq_tr @ (map (fn thm => (thm RS not_sym)) dist_eq_tr));
 | 
| 2640 | 54 | |
| 55 | (* ------------------------------------------------------------------------ *) | |
| 56 | (* lemmas about andalso, orelse, neg and if *) | |
| 57 | (* ------------------------------------------------------------------------ *) | |
| 58 | ||
| 7654 | 59 | bind_thms ("andalso_thms", map prover [
 | 
| 2640 | 60 | "(TT andalso y) = y", | 
| 61 | "(FF andalso y) = FF", | |
| 62 | "(UU andalso y) = UU", | |
| 63 | "(y andalso TT) = y", | |
| 64 | "(y andalso y) = y" | |
| 7654 | 65 | ]); | 
| 2640 | 66 | |
| 7654 | 67 | bind_thms ("orelse_thms", map prover [
 | 
| 2640 | 68 | "(TT orelse y) = TT", | 
| 69 | "(FF orelse y) = y", | |
| 70 | "(UU orelse y) = UU", | |
| 71 | "(y orelse FF) = y", | |
| 7654 | 72 | "(y orelse y) = y"]); | 
| 2640 | 73 | |
| 7654 | 74 | bind_thms ("neg_thms", map prover [
 | 
| 2640 | 75 | "neg`TT = FF", | 
| 76 | "neg`FF = TT", | |
| 77 | "neg`UU = UU" | |
| 7654 | 78 | ]); | 
| 2640 | 79 | |
| 7654 | 80 | bind_thms ("ifte_thms", map prover [
 | 
| 2640 | 81 | "If UU then e1 else e2 fi = UU", | 
| 82 | "If FF then e1 else e2 fi = e2", | |
| 7654 | 83 | "If TT then e1 else e2 fi = e1"]); | 
| 2640 | 84 | |
| 85 | Addsimps (dist_less_tr @ dist_eq_tr @ andalso_thms @ | |
| 86 | orelse_thms @ neg_thms @ ifte_thms); | |
| 87 | ||
| 3038 | 88 | (* ------------------------------------------------------------------- *) | 
| 89 | (* split-tac for If via If2 because the constant has to be a constant *) | |
| 90 | (* ------------------------------------------------------------------- *) | |
| 91 | ||
| 5068 | 92 | Goalw [If2_def] | 
| 3038 | 93 | "P (If2 Q x y ) = ((Q=UU --> P UU) & (Q=TT --> P x) & (Q=FF --> P y))"; | 
| 94 | by (res_inst_tac [("p","Q")] trE 1);
 | |
| 95 | by (REPEAT (Asm_full_simp_tac 1)); | |
| 4833 | 96 | qed"split_If2"; | 
| 2640 | 97 | |
| 3038 | 98 | val split_If_tac = | 
| 4833 | 99 | simp_tac (HOL_basic_ss addsimps [symmetric If2_def]) THEN' (split_tac [split_If2]); | 
| 2640 | 100 | |
| 3038 | 101 | |
| 2640 | 102 | |
| 3038 | 103 | (* ----------------------------------------------------------------- *) | 
| 104 | section"Rewriting of HOLCF operations to HOL functions"; | |
| 105 | (* ----------------------------------------------------------------- *) | |
| 2640 | 106 | |
| 107 | ||
| 5068 | 108 | Goal | 
| 3038 | 109 | "!!t.[|t~=UU|]==> ((t andalso s)=FF)=(t=FF | s=FF)"; | 
| 110 | by (rtac iffI 1); | |
| 111 | by (res_inst_tac [("p","t")] trE 1);
 | |
| 4477 
b3e5857d8d99
New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
 paulson parents: 
4423diff
changeset | 112 | by Auto_tac; | 
| 3038 | 113 | by (res_inst_tac [("p","t")] trE 1);
 | 
| 4477 
b3e5857d8d99
New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
 paulson parents: 
4423diff
changeset | 114 | by Auto_tac; | 
| 3038 | 115 | qed"andalso_or"; | 
| 2640 | 116 | |
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5068diff
changeset | 117 | Goal "[|t~=UU|]==> ((t andalso s)~=FF)=(t~=FF & s~=FF)"; | 
| 2640 | 118 | by (rtac iffI 1); | 
| 3038 | 119 | by (res_inst_tac [("p","t")] trE 1);
 | 
| 4477 
b3e5857d8d99
New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
 paulson parents: 
4423diff
changeset | 120 | by Auto_tac; | 
| 2640 | 121 | by (res_inst_tac [("p","t")] trE 1);
 | 
| 4477 
b3e5857d8d99
New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
 paulson parents: 
4423diff
changeset | 122 | by Auto_tac; | 
| 2640 | 123 | qed"andalso_and"; | 
| 124 | ||
| 5068 | 125 | Goal "(Def x ~=FF)= x"; | 
| 4098 | 126 | by (simp_tac (simpset() addsimps [FF_def]) 1); | 
| 3038 | 127 | qed"Def_bool1"; | 
| 2640 | 128 | |
| 5068 | 129 | Goal "(Def x = FF) = (~x)"; | 
| 4098 | 130 | by (simp_tac (simpset() addsimps [FF_def]) 1); | 
| 3038 | 131 | qed"Def_bool2"; | 
| 2640 | 132 | |
| 5068 | 133 | Goal "(Def x = TT) = x"; | 
| 4098 | 134 | by (simp_tac (simpset() addsimps [TT_def]) 1); | 
| 3038 | 135 | qed"Def_bool3"; | 
| 136 | ||
| 5068 | 137 | Goal "(Def x ~= TT) = (~x)"; | 
| 4098 | 138 | by (simp_tac (simpset() addsimps [TT_def]) 1); | 
| 3038 | 139 | qed"Def_bool4"; | 
| 2640 | 140 | |
| 5068 | 141 | Goal | 
| 2640 | 142 | "(If Def P then A else B fi)= (if P then A else B)"; | 
| 143 | by (res_inst_tac [("p","Def P")]  trE 1);
 | |
| 144 | by (Asm_full_simp_tac 1); | |
| 4098 | 145 | by (asm_full_simp_tac (simpset() addsimps tr_defs@[flift1_def,o_def]) 1); | 
| 146 | by (asm_full_simp_tac (simpset() addsimps tr_defs@[flift1_def,o_def]) 1); | |
| 2640 | 147 | qed"If_and_if"; | 
| 148 | ||
| 3038 | 149 | Addsimps [Def_bool1,Def_bool2,Def_bool3,Def_bool4]; | 
| 150 | ||
| 151 | (* ----------------------------------------------------------------- *) | |
| 152 | section"admissibility"; | |
| 153 | (* ----------------------------------------------------------------- *) | |
| 154 | ||
| 155 | ||
| 156 | (* The following rewrite rules for admissibility should in the future be | |
| 157 | replaced by a more general admissibility test that also checks | |
| 158 | chain-finiteness, of which these lemmata are specific examples *) | |
| 159 | ||
| 5068 | 160 | Goal "x~=FF = (x=TT|x=UU)"; | 
| 3038 | 161 | by (res_inst_tac [("p","x")] trE 1);
 | 
| 162 | by (TRYALL (Asm_full_simp_tac)); | |
| 163 | qed"adm_trick_1"; | |
| 164 | ||
| 5068 | 165 | Goal "x~=TT = (x=FF|x=UU)"; | 
| 3038 | 166 | by (res_inst_tac [("p","x")] trE 1);
 | 
| 167 | by (TRYALL (Asm_full_simp_tac)); | |
| 168 | qed"adm_trick_2"; | |
| 169 | ||
| 7654 | 170 | bind_thms ("adm_tricks", [adm_trick_1,adm_trick_2]);
 | 
| 3038 | 171 | |
| 172 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5068diff
changeset | 173 | Goal "cont(f) ==> adm (%x. (f x)~=TT)"; | 
| 3038 | 174 | by (simp_tac (HOL_basic_ss addsimps adm_tricks) 1); | 
| 175 | by (REPEAT ((resolve_tac (adm_lemmas@cont_lemmas1) 1) ORELSE atac 1)); | |
| 176 | qed"adm_nTT"; | |
| 177 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5068diff
changeset | 178 | Goal "cont(f) ==> adm (%x. (f x)~=FF)"; | 
| 3038 | 179 | by (simp_tac (HOL_basic_ss addsimps adm_tricks) 1); | 
| 180 | by (REPEAT ((resolve_tac (adm_lemmas@cont_lemmas1) 1) ORELSE atac 1)); | |
| 181 | qed"adm_nFF"; | |
| 182 | ||
| 4642 | 183 | Addsimps [adm_nTT,adm_nFF]; |