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