| 2525 |      1 | (* Title:     HOL/MiniML/Instance.ML
 | 
|  |      2 |    ID:        $Id$
 | 
|  |      3 |    Author:    Wolfgang Naraschewski and Tobias Nipkow
 | 
|  |      4 |    Copyright  1996 TU Muenchen
 | 
|  |      5 | *)
 | 
|  |      6 | 
 | 
|  |      7 | (* lemmatas for instatiation *)
 | 
|  |      8 | 
 | 
|  |      9 | 
 | 
|  |     10 | (* lemmatas for bound_typ_inst *)
 | 
|  |     11 | 
 | 
| 5069 |     12 | Goal "bound_typ_inst S (mk_scheme t) = t";
 | 
| 5184 |     13 | by (induct_tac "t" 1);
 | 
| 2525 |     14 | by (ALLGOALS Asm_simp_tac);
 | 
|  |     15 | qed "bound_typ_inst_mk_scheme";
 | 
|  |     16 | 
 | 
|  |     17 | Addsimps [bound_typ_inst_mk_scheme];
 | 
| 2625 |     18 | 
 | 
| 5118 |     19 | Goal "bound_typ_inst ($S o R) ($S sch) = $S (bound_typ_inst R sch)";
 | 
| 5184 |     20 | by (induct_tac "sch" 1);
 | 
| 2525 |     21 | by (ALLGOALS Asm_full_simp_tac);
 | 
|  |     22 | qed "bound_typ_inst_composed_subst";
 | 
|  |     23 | 
 | 
|  |     24 | Addsimps [bound_typ_inst_composed_subst];
 | 
|  |     25 | 
 | 
| 5118 |     26 | Goal "S = S' ==> sch = sch' ==> bound_typ_inst S sch = bound_typ_inst S' sch'";
 | 
| 2525 |     27 | by (Asm_full_simp_tac 1);
 | 
|  |     28 | qed "bound_typ_inst_eq";
 | 
|  |     29 | 
 | 
|  |     30 | 
 | 
| 2625 |     31 | 
 | 
| 2525 |     32 | (* lemmatas for bound_scheme_inst *)
 | 
|  |     33 | 
 | 
| 5118 |     34 | Goal "bound_scheme_inst B (mk_scheme t) = mk_scheme t";
 | 
| 5184 |     35 | by (induct_tac "t" 1);
 | 
| 2525 |     36 | by (Simp_tac 1);
 | 
|  |     37 | by (Asm_simp_tac 1);
 | 
|  |     38 | qed "bound_scheme_inst_mk_scheme";
 | 
|  |     39 | 
 | 
|  |     40 | Addsimps [bound_scheme_inst_mk_scheme];
 | 
|  |     41 | 
 | 
| 5118 |     42 | Goal "$S (bound_scheme_inst B sch) = (bound_scheme_inst ($S o B) ($ S sch))";
 | 
| 5184 |     43 | by (induct_tac "sch" 1);
 | 
| 2525 |     44 | by (Simp_tac 1);
 | 
|  |     45 | by (Simp_tac 1);
 | 
|  |     46 | by (Asm_simp_tac 1);
 | 
|  |     47 | qed "substitution_lemma";
 | 
|  |     48 | 
 | 
| 5069 |     49 | Goal "!t. mk_scheme t = bound_scheme_inst B sch --> \
 | 
| 2525 |     50 | \         (? S. !x:bound_tv sch. B x = mk_scheme (S x))";
 | 
| 5184 |     51 | by (induct_tac "sch" 1);
 | 
| 2525 |     52 | by (Simp_tac 1);
 | 
| 4153 |     53 | by Safe_tac;
 | 
| 2525 |     54 | by (rtac exI 1);
 | 
|  |     55 | by (rtac ballI 1);
 | 
|  |     56 | by (rtac sym 1);
 | 
|  |     57 | by (Asm_full_simp_tac 1);
 | 
|  |     58 | by (Asm_full_simp_tac 1);
 | 
|  |     59 | by (dtac mk_scheme_Fun 1);
 | 
|  |     60 | by (REPEAT (etac exE 1));
 | 
|  |     61 | by (etac conjE 1);
 | 
|  |     62 | by (dtac sym 1);
 | 
|  |     63 | by (dtac sym 1);
 | 
|  |     64 | by (REPEAT ((dtac mp 1) THEN (Fast_tac 1)));
 | 
| 4153 |     65 | by Safe_tac;
 | 
| 2525 |     66 | by (rename_tac "S1 S2" 1);
 | 
|  |     67 | by (res_inst_tac [("x","%x. if x:bound_tv type_scheme1 then (S1 x) else (S2 x)")] exI 1);
 | 
| 5521 |     68 | by (Auto_tac);
 | 
| 2525 |     69 | qed_spec_mp "bound_scheme_inst_type";
 | 
|  |     70 | 
 | 
|  |     71 | 
 | 
| 5118 |     72 | (* lemmas for subst_to_scheme *)
 | 
| 2525 |     73 | 
 | 
| 5118 |     74 | Goal "new_tv n sch --> subst_to_scheme (%k. if n <= k then BVar (k - n) else FVar k) \
 | 
| 2525 |     75 | \                                                 (bound_typ_inst (%k. TVar (k + n)) sch) = sch";
 | 
| 5184 |     76 | by (induct_tac "sch" 1);
 | 
| 5350 |     77 | by (simp_tac (simpset() addsimps [le_def]) 1);
 | 
| 4686 |     78 | by (simp_tac (simpset() addsimps [le_add2,diff_add_inverse2]) 1);
 | 
| 2525 |     79 | by (Asm_simp_tac 1);
 | 
|  |     80 | qed_spec_mp "subst_to_scheme_inverse";
 | 
|  |     81 | 
 | 
| 5118 |     82 | Goal "t = t' ==> \
 | 
|  |     83 | \     subst_to_scheme (%k. if n <= k then BVar (k - n) else FVar k) t = \
 | 
|  |     84 | \     subst_to_scheme (%k. if n <= k then BVar (k - n) else FVar k) t'";
 | 
| 2525 |     85 | by (Fast_tac 1);
 | 
|  |     86 | val aux = result ();
 | 
|  |     87 | 
 | 
| 5069 |     88 | Goal "new_tv n sch --> \
 | 
| 5118 |     89 | \     subst_to_scheme (%k. if n <= k then BVar (k - n) else FVar k) (bound_typ_inst S sch) = \
 | 
|  |     90 | \      bound_scheme_inst ((subst_to_scheme (%k. if n <= k then BVar (k - n) else FVar k)) o S) sch";
 | 
| 5184 |     91 | by (induct_tac "sch" 1);
 | 
| 5350 |     92 | by Auto_tac;
 | 
|  |     93 | by (ALLGOALS trans_tac);
 | 
| 2525 |     94 | val aux2 = result () RS mp;
 | 
|  |     95 | 
 | 
|  |     96 | 
 | 
|  |     97 | (* lemmata for <= *)
 | 
|  |     98 | 
 | 
| 5069 |     99 | Goalw [le_type_scheme_def,is_bound_typ_instance]
 | 
| 5118 |    100 |   "!!(sch::type_scheme) sch'. \
 | 
|  |    101 | \  (sch' <= sch) = (? B. sch' = bound_scheme_inst B sch)";
 | 
| 2525 |    102 | by (rtac iffI 1);
 | 
|  |    103 | by (cut_inst_tac [("sch","sch")] fresh_variable_type_schemes 1); 
 | 
|  |    104 | by (cut_inst_tac [("sch","sch'")] fresh_variable_type_schemes 1);
 | 
|  |    105 | by (dtac make_one_new_out_of_two 1);
 | 
| 3018 |    106 | by (assume_tac 1);
 | 
| 2525 |    107 | by (thin_tac "? n. new_tv n sch'" 1); 
 | 
|  |    108 | by (etac exE 1);
 | 
|  |    109 | by (etac allE 1);
 | 
|  |    110 | by (dtac mp 1);
 | 
|  |    111 | by (res_inst_tac [("x","(%k. TVar (k + n))")] exI 1);
 | 
|  |    112 | by (rtac refl 1);
 | 
|  |    113 | by (etac exE 1);
 | 
|  |    114 | by (REPEAT (etac conjE 1));
 | 
|  |    115 | by (dres_inst_tac [("n","n")] aux 1);
 | 
| 4089 |    116 | by (asm_full_simp_tac (simpset() addsimps [subst_to_scheme_inverse]) 1);
 | 
| 2525 |    117 | by (res_inst_tac [("x","(subst_to_scheme (%k. if n <= k then BVar (k - n) else FVar k)) o S")] exI 1);
 | 
| 4089 |    118 | by (asm_simp_tac (simpset() addsimps [aux2]) 1);
 | 
| 4153 |    119 | by Safe_tac;
 | 
| 2525 |    120 | by (res_inst_tac [("x","%n. bound_typ_inst S (B n)")] exI 1);
 | 
| 5184 |    121 | by (induct_tac "sch" 1);
 | 
| 2525 |    122 | by (Simp_tac 1);
 | 
|  |    123 | by (Simp_tac 1);
 | 
|  |    124 | by (Asm_simp_tac 1);
 | 
|  |    125 | qed "le_type_scheme_def2";
 | 
|  |    126 | 
 | 
| 5069 |    127 | Goalw [is_bound_typ_instance] "(mk_scheme t) <= sch = t <| sch";
 | 
| 4089 |    128 | by (simp_tac (simpset() addsimps [le_type_scheme_def2]) 1); 
 | 
| 2525 |    129 | by (rtac iffI 1); 
 | 
|  |    130 | by (etac exE 1); 
 | 
|  |    131 | by (forward_tac [bound_scheme_inst_type] 1);
 | 
|  |    132 | by (etac exE 1);
 | 
|  |    133 | by (rtac exI 1);
 | 
|  |    134 | by (rtac mk_scheme_injective 1); 
 | 
|  |    135 | by (Asm_full_simp_tac 1);
 | 
|  |    136 | by (rotate_tac 1 1);
 | 
|  |    137 | by (rtac mp 1);
 | 
| 3018 |    138 | by (assume_tac 2);
 | 
| 5184 |    139 | by (induct_tac "sch" 1);
 | 
| 2525 |    140 | by (Simp_tac 1);
 | 
|  |    141 | by (Asm_full_simp_tac 1);
 | 
|  |    142 | by (Fast_tac 1);
 | 
|  |    143 | by (strip_tac 1);
 | 
|  |    144 | by (Asm_full_simp_tac 1);
 | 
|  |    145 | by (etac exE 1);
 | 
|  |    146 | by (Asm_full_simp_tac 1);
 | 
|  |    147 | by (rtac exI 1);
 | 
| 5184 |    148 | by (induct_tac "sch" 1);
 | 
| 2525 |    149 | by (Simp_tac 1);
 | 
|  |    150 | by (Simp_tac 1);
 | 
|  |    151 | by (Asm_full_simp_tac 1);
 | 
|  |    152 | qed_spec_mp "le_type_eq_is_bound_typ_instance";
 | 
|  |    153 | 
 | 
| 5069 |    154 | Goalw [le_env_def]
 | 
| 2525 |    155 |   "(sch # A <= sch' # B) = (sch <= (sch'::type_scheme) & A <= B)";
 | 
| 3018 |    156 | by (Simp_tac 1);
 | 
|  |    157 | by (rtac iffI 1);
 | 
| 4153 |    158 |  by (SELECT_GOAL Safe_tac 1);
 | 
| 3018 |    159 |   by (eres_inst_tac [("x","0")] allE 1);
 | 
|  |    160 |   by (Asm_full_simp_tac 1);
 | 
|  |    161 |  by (eres_inst_tac [("x","Suc i")] allE 1);
 | 
|  |    162 |  by (Asm_full_simp_tac 1);
 | 
|  |    163 | by (rtac conjI 1);
 | 
|  |    164 |  by (Fast_tac 1);
 | 
|  |    165 | by (rtac allI 1);
 | 
| 5184 |    166 | by (induct_tac "i" 1);
 | 
| 3018 |    167 | by (ALLGOALS Asm_simp_tac);
 | 
| 2525 |    168 | qed "le_env_Cons";
 | 
|  |    169 | AddIffs [le_env_Cons];
 | 
|  |    170 | 
 | 
| 5118 |    171 | Goalw [is_bound_typ_instance]"t <| sch ==> $S t <| $S sch";
 | 
| 2525 |    172 | by (etac exE 1);
 | 
|  |    173 | by (rename_tac "SA" 1);
 | 
|  |    174 | by (hyp_subst_tac 1);
 | 
|  |    175 | by (res_inst_tac [("x","$S o SA")] exI 1);
 | 
|  |    176 | by (Simp_tac 1);
 | 
|  |    177 | qed "is_bound_typ_instance_closed_subst";
 | 
|  |    178 | 
 | 
| 5069 |    179 | Goal "!!(sch::type_scheme) sch'. sch' <= sch ==> $S sch' <= $ S sch";
 | 
| 4089 |    180 | by (asm_full_simp_tac (simpset() addsimps [le_type_scheme_def2]) 1);
 | 
| 2525 |    181 | by (etac exE 1);
 | 
| 4089 |    182 | by (asm_full_simp_tac (simpset() addsimps [substitution_lemma]) 1);
 | 
| 2525 |    183 | by (Fast_tac 1);
 | 
|  |    184 | qed "S_compatible_le_scheme";
 | 
|  |    185 | 
 | 
| 5118 |    186 | Goalw [le_env_def,app_subst_list]
 | 
|  |    187 |  "!!(A::type_scheme list) A'. A' <= A ==> $S A' <= $ S A";
 | 
| 4089 |    188 | by (simp_tac (simpset() addcongs [conj_cong]) 1);
 | 
|  |    189 | by (fast_tac (claset() addSIs [S_compatible_le_scheme]) 1);
 | 
| 2525 |    190 | qed "S_compatible_le_scheme_lists";
 | 
|  |    191 | 
 | 
| 5118 |    192 | Goalw [le_type_scheme_def] "[| t <| sch; sch <= sch' |] ==> t <| sch'";
 | 
| 3018 |    193 | by (Fast_tac 1);
 | 
| 2525 |    194 | qed "bound_typ_instance_trans";
 | 
|  |    195 | 
 | 
| 5069 |    196 | Goalw [le_type_scheme_def] "sch <= (sch::type_scheme)";
 | 
| 3018 |    197 | by (Fast_tac 1);
 | 
| 2525 |    198 | qed "le_type_scheme_refl";
 | 
|  |    199 | AddIffs [le_type_scheme_refl];
 | 
|  |    200 | 
 | 
| 5069 |    201 | Goalw [le_env_def] "A <= (A::type_scheme list)";
 | 
| 3018 |    202 | by (Fast_tac 1);
 | 
| 2525 |    203 | qed "le_env_refl";
 | 
|  |    204 | AddIffs [le_env_refl];
 | 
|  |    205 | 
 | 
| 5069 |    206 | Goalw [le_type_scheme_def,is_bound_typ_instance] "sch <= BVar n";
 | 
| 3018 |    207 | by (strip_tac 1);
 | 
| 3842 |    208 | by (res_inst_tac [("x","%a. t")]exI 1);
 | 
| 3018 |    209 | by (Simp_tac 1);
 | 
| 2525 |    210 | qed "bound_typ_instance_BVar";
 | 
|  |    211 | AddIffs [bound_typ_instance_BVar];
 | 
|  |    212 | 
 | 
| 5118 |    213 | Goalw [le_type_scheme_def,is_bound_typ_instance]
 | 
|  |    214 |  "(sch <= FVar n) = (sch = FVar n)";
 | 
| 5184 |    215 | by (induct_tac "sch" 1);
 | 
| 3018 |    216 |   by (Simp_tac 1);
 | 
|  |    217 |  by (Simp_tac 1);
 | 
|  |    218 |  by (Fast_tac 1);
 | 
|  |    219 | by (Asm_full_simp_tac 1);
 | 
|  |    220 | by (Fast_tac 1);
 | 
| 2525 |    221 | qed "le_FVar";
 | 
|  |    222 | Addsimps [le_FVar];
 | 
|  |    223 | 
 | 
| 5069 |    224 | Goalw [le_type_scheme_def,is_bound_typ_instance] "~(FVar n <= sch1 =-> sch2)";
 | 
| 3018 |    225 | by (Simp_tac 1);
 | 
| 2525 |    226 | qed "not_FVar_le_Fun";
 | 
|  |    227 | AddIffs [not_FVar_le_Fun];
 | 
|  |    228 | 
 | 
| 5069 |    229 | Goalw [le_type_scheme_def,is_bound_typ_instance] "~(BVar n <= sch1 =-> sch2)";
 | 
| 3018 |    230 | by (Simp_tac 1);
 | 
|  |    231 | by (res_inst_tac [("x","TVar n")] exI 1);
 | 
|  |    232 | by (Simp_tac 1);
 | 
|  |    233 | by (Fast_tac 1);
 | 
| 2525 |    234 | qed "not_BVar_le_Fun";
 | 
|  |    235 | AddIffs [not_BVar_le_Fun];
 | 
|  |    236 | 
 | 
| 5069 |    237 | Goalw [le_type_scheme_def,is_bound_typ_instance]
 | 
| 5118 |    238 |   "(sch1 =-> sch2 <= sch1' =-> sch2') ==> sch1 <= sch1' & sch2 <= sch2'";
 | 
| 4089 |    239 | by (fast_tac (claset() addss simpset()) 1);
 | 
| 2525 |    240 | qed "Fun_le_FunD";
 | 
|  |    241 | 
 | 
| 5069 |    242 | Goal "(sch' <= sch1 =-> sch2) --> (? sch'1 sch'2. sch' = sch'1 =-> sch'2)";
 | 
| 5184 |    243 | by (induct_tac "sch'" 1);
 | 
| 2525 |    244 | by (Asm_simp_tac 1);
 | 
|  |    245 | by (Asm_simp_tac 1);
 | 
|  |    246 | by (Fast_tac 1);
 | 
|  |    247 | qed_spec_mp "scheme_le_Fun";
 | 
|  |    248 | 
 | 
| 5069 |    249 | Goal "!sch'::type_scheme. sch <= sch' --> free_tv sch' <= free_tv sch";
 | 
| 5184 |    250 | by (induct_tac "sch" 1);
 | 
| 3018 |    251 |   by (rtac allI 1);
 | 
| 5184 |    252 |   by (induct_tac "sch'" 1);
 | 
| 3018 |    253 |     by (Simp_tac 1);
 | 
|  |    254 |    by (Simp_tac 1);
 | 
|  |    255 |   by (Simp_tac 1);
 | 
|  |    256 |  by (rtac allI 1);
 | 
| 5184 |    257 |  by (induct_tac "sch'" 1);
 | 
| 3018 |    258 |    by (Simp_tac 1);
 | 
|  |    259 |   by (Simp_tac 1);
 | 
|  |    260 |  by (Simp_tac 1);
 | 
|  |    261 | by (rtac allI 1);
 | 
| 5184 |    262 | by (induct_tac "sch'" 1);
 | 
| 3018 |    263 |   by (Simp_tac 1);
 | 
|  |    264 |  by (Simp_tac 1);
 | 
|  |    265 | by (Asm_full_simp_tac 1);
 | 
|  |    266 | by (strip_tac 1);
 | 
|  |    267 | by (dtac Fun_le_FunD 1);
 | 
|  |    268 | by (Fast_tac 1);
 | 
| 2525 |    269 | qed_spec_mp "le_type_scheme_free_tv";
 | 
|  |    270 | 
 | 
| 5069 |    271 | Goal "!A::type_scheme list. A <= B --> free_tv B <= free_tv A";
 | 
| 5184 |    272 | by (induct_tac "B" 1);
 | 
| 3018 |    273 |  by (Simp_tac 1);
 | 
|  |    274 | by (rtac allI 1);
 | 
| 5184 |    275 | by (induct_tac "A" 1);
 | 
| 4089 |    276 |  by (simp_tac (simpset() addsimps [le_env_def]) 1);
 | 
| 3018 |    277 | by (Simp_tac 1);
 | 
| 4089 |    278 | by (fast_tac (claset() addDs [le_type_scheme_free_tv]) 1);
 | 
| 2525 |    279 | qed_spec_mp "le_env_free_tv";
 |