| author | kleing | 
| Wed, 31 May 2000 18:06:02 +0200 | |
| changeset 9012 | d1bd2144ab5d | 
| parent 8442 | 96023903c2df | 
| permissions | -rw-r--r-- | 
| 8011 | 1 | (* Title: HOL/MicroJava/BV/Convert.ML | 
| 2 | ID: $Id$ | |
| 3 | Author: Cornelia Pusch | |
| 4 | Copyright 1999 Technische Universitaet Muenchen | |
| 5 | *) | |
| 6 | ||
| 8139 | 7 | |
| 8 | Goalw[sup_ty_opt_def] "G \\<turnstile> t <=o t"; | |
| 8442 
96023903c2df
case_tac now subsumes both boolean and datatype cases;
 wenzelm parents: 
8423diff
changeset | 9 | by (case_tac "t" 1); | 
| 8139 | 10 | by Auto_tac; | 
| 11 | qed "sup_ty_opt_refl"; | |
| 12 | Addsimps [sup_ty_opt_refl]; | |
| 13 | ||
| 14 | Goalw[sup_loc_def] "G \\<turnstile> t <=l t"; | |
| 15 | by (induct_tac "t" 1); | |
| 16 | by Auto_tac; | |
| 17 | qed "sup_loc_refl"; | |
| 18 | Addsimps [sup_loc_refl]; | |
| 19 | ||
| 20 | Goalw[sup_state_def] "G \\<turnstile> s <=s s"; | |
| 21 | by Auto_tac; | |
| 22 | qed "sup_state_refl"; | |
| 23 | Addsimps [sup_state_refl]; | |
| 24 | ||
| 8011 | 25 | val widen_PrimT_conv1 = | 
| 26 | prove_typerel "(G \\<turnstile> PrimT x \\<preceq> T) = (T = PrimT x)" | |
| 27 | [ prove_widen_lemma "G\\<turnstile>S\\<preceq>T \\<Longrightarrow> S = PrimT x \\<longrightarrow> T = PrimT x"]; | |
| 28 | Addsimps [widen_PrimT_conv1]; | |
| 29 | ||
| 8023 | 30 | Goalw [sup_ty_opt_def] "(G \\<turnstile> None <=o any) = (any = None)"; | 
| 8011 | 31 | by(simp_tac (simpset() addsplits [option.split]) 1); | 
| 32 | qed "anyConvNone"; | |
| 33 | Addsimps [anyConvNone]; | |
| 34 | ||
| 8023 | 35 | Goalw [sup_ty_opt_def] "(G \\<turnstile> Some ty' <=o Some ty) = (G \\<turnstile> ty' \\<preceq> ty)"; | 
| 8011 | 36 | by(Simp_tac 1); | 
| 37 | qed "SomeanyConvSome"; | |
| 38 | Addsimps [SomeanyConvSome]; | |
| 39 | ||
| 40 | Goal | |
| 8023 | 41 | "(G \\<turnstile> Some(PrimT Integer) <=o X) = (X=None \\<or> (X=Some(PrimT Integer)))"; | 
| 8011 | 42 | by (simp_tac (simpset() setloop (split_tac [ty.split,option.split]) addsimps [sup_ty_opt_def]) 1); | 
| 43 | qed "sup_PTS_eq"; | |
| 44 | ||
| 45 | ||
| 8119 | 46 | Goalw [sup_loc_def] "CFS \\<turnstile> [] <=l XT = (XT=[])"; | 
| 47 | by(Simp_tac 1); | |
| 8011 | 48 | qed "sup_loc_Nil"; | 
| 8119 | 49 | AddIffs [sup_loc_Nil]; | 
| 50 | ||
| 8011 | 51 | |
| 8119 | 52 | Goalw [sup_loc_def] | 
| 8023 | 53 | "CFS \\<turnstile> (Y#YT) <=l XT = (\\<exists>X XT'. XT=X#XT' \\<and> CFS \\<turnstile> Y <=o X \\<and> CFS \\<turnstile> YT <=l XT')"; | 
| 8119 | 54 | by(simp_tac (simpset() addsimps [list_all2_Cons1]) 1); | 
| 8011 | 55 | qed "sup_loc_Cons"; | 
| 8119 | 56 | AddIffs [sup_loc_Cons]; | 
| 8139 | 57 | |
| 8389 | 58 | |
| 59 | Goal "\\<forall> b. G \\<turnstile> a <=l b \\<longrightarrow> length a = length b"; | |
| 60 | by (induct_tac "a" 1); | |
| 61 | by (Simp_tac 1); | |
| 62 | by (Clarsimp_tac 1); | |
| 63 | qed_spec_mp "sup_loc_length"; | |
| 64 | ||
| 65 | Goal "\\<forall> n b. (G \\<turnstile> a <=l b) \\<longrightarrow> n < length a \\<longrightarrow> (G \\<turnstile> (a!n) <=o (b!n))"; | |
| 66 | by (induct_tac "a" 1); | |
| 67 | by (Simp_tac 1); | |
| 68 | by (Clarsimp_tac 1); | |
| 8442 
96023903c2df
case_tac now subsumes both boolean and datatype cases;
 wenzelm parents: 
8423diff
changeset | 69 | by (case_tac "n" 1); | 
| 8389 | 70 | by (Asm_full_simp_tac 1); | 
| 71 | by (eres_inst_tac [("x","nat")] allE 1);
 | |
| 72 | by (smp_tac 1 1); | |
| 73 | by (Asm_full_simp_tac 1); | |
| 74 | qed_spec_mp "sup_loc_nth"; | |
| 75 | ||
| 76 | Goalw[sup_state_def] | |
| 77 | "(G \\<turnstile> (x, y) <=s (a, b)) \\<Longrightarrow> length x = length a"; | |
| 78 | by (Clarsimp_tac 1); | |
| 79 | bd sup_loc_length 1; | |
| 80 | by Auto_tac; | |
| 81 | qed "sup_state_length_fst"; | |
| 82 | ||
| 83 | Goalw[sup_state_def] | |
| 84 | "(G \\<turnstile> (x, y) <=s (a, b)) \\<longrightarrow> length y = length b"; | |
| 85 | by (clarsimp_tac (claset(), simpset() addsimps [sup_loc_length]) 1); | |
| 86 | qed "sup_state_length_snd"; | |
| 87 | ||
| 8394 | 88 | Goal "\\<forall>b. length a = length b \\<longrightarrow> (G \\<turnstile> (a@x) <=l (b@y)) = ((G \\<turnstile> a <=l b) \\<and> (G \\<turnstile> x <=l y))"; | 
| 89 | by (induct_tac "a" 1); | |
| 90 | by (Clarsimp_tac 1); | |
| 91 | by (Clarify_tac 1); | |
| 8442 
96023903c2df
case_tac now subsumes both boolean and datatype cases;
 wenzelm parents: 
8423diff
changeset | 92 | by (case_tac "b" 1); | 
| 8394 | 93 | by (Clarsimp_tac 1); | 
| 94 | by (Clarsimp_tac 1); | |
| 95 | qed_spec_mp "sup_loc_append"; | |
| 96 | ||
| 97 | ||
| 98 | Goalw[sup_state_def] | |
| 99 | "length a = length b \\<Longrightarrow> (G \\<turnstile> (i,a@x) <=s (j,b@y)) = ((G \\<turnstile> (i,a) <=s (j,b)) \\<and> (G \\<turnstile> (i,x) <=s (j,y)))"; | |
| 100 | by (auto_tac (claset(), simpset() addsimps [sup_loc_append])); | |
| 101 | qed "sup_state_append_snd"; | |
| 102 | ||
| 103 | Goalw[sup_state_def] | |
| 104 | "length a = length b \\<Longrightarrow> (G \\<turnstile> (a@x,i) <=s (b@y,j)) = ((G \\<turnstile> (a,i) <=s (b,j)) \\<and> (G \\<turnstile> (x,i) <=s (y,j)))"; | |
| 105 | by (auto_tac (claset(), simpset() addsimps [sup_loc_append])); | |
| 106 | qed "sup_state_append_fst"; | |
| 107 | ||
| 108 | ||
| 109 | Goal "\\<forall> b. (G \\<turnstile> (rev a) <=l rev b) = (G \\<turnstile> a <=l b)"; | |
| 110 | by (induct_tac "a" 1); | |
| 111 | by (Simp_tac 1); | |
| 112 | by (Clarsimp_tac 1); | |
| 113 | by Safe_tac; | |
| 114 | by (clarsimp_tac (claset(), simpset() addsimps [sup_loc_append, sup_loc_length]) 2); | |
| 8442 
96023903c2df
case_tac now subsumes both boolean and datatype cases;
 wenzelm parents: 
8423diff
changeset | 115 | by (case_tac "b" 1); | 
| 8394 | 116 | bd sup_loc_length 1; | 
| 117 | by (Clarsimp_tac 1); | |
| 118 | by (clarsimp_tac (claset(), simpset() addsimps [sup_loc_append, sup_loc_length]) 1); | |
| 119 | by (subgoal_tac "length (rev list) = length (rev lista)" 1); | |
| 120 | bd sup_loc_length 2; | |
| 121 | by (Clarsimp_tac 2); | |
| 122 | by (dres_inst_tac [("G","G"),("x","[a]"),("y","[aa]")] sup_loc_append 1);
 | |
| 123 | by (Clarsimp_tac 1); | |
| 124 | qed_spec_mp "sup_loc_rev"; | |
| 125 | ||
| 126 | Goalw[sup_state_def] | |
| 127 | "(G \\<turnstile> (rev a,x) <=s (rev b,y)) = (G \\<turnstile> (a,x) <=s (b,y))"; | |
| 128 | by (auto_tac (claset(), simpset() addsimps [sup_loc_rev, rev_map RS sym])); | |
| 129 | qed "sup_state_rev_fst"; | |
| 130 | ||
| 131 | Goal "map x a = Y # YT \\<longrightarrow> map x (tl a) = YT \\<and> x (hd a) = Y \\<and> (\\<exists> y yt. a = y#yt)"; | |
| 8442 
96023903c2df
case_tac now subsumes both boolean and datatype cases;
 wenzelm parents: 
8423diff
changeset | 132 | by (case_tac "a" 1); | 
| 8394 | 133 | by Auto_tac; | 
| 134 | qed_spec_mp "map_hd_tl"; | |
| 135 | ||
| 136 | Goalw[sup_state_def] | |
| 137 | "(G \\<turnstile> (x#xt, a) <=s (yt, b)) = (\\<exists>y yt'. yt=y#yt' \\<and> (G \\<turnstile> x \\<preceq> y) \\<and> (G \\<turnstile> (xt,a) <=s (yt',b)))"; | |
| 138 | by Auto_tac; | |
| 139 | bd map_hd_tl 1; | |
| 140 | by Auto_tac; | |
| 141 | qed "sup_state_Cons1"; | |
| 142 | ||
| 143 | Goalw [sup_loc_def] | |
| 144 | "CFS \\<turnstile> YT <=l (X#XT) = (\\<exists>Y YT'. YT=Y#YT' \\<and> CFS \\<turnstile> Y <=o X \\<and> CFS \\<turnstile> YT' <=l XT)"; | |
| 145 | by(simp_tac (simpset() addsimps [list_all2_Cons2]) 1); | |
| 146 | qed "sup_loc_Cons2"; | |
| 147 | ||
| 148 | Goalw[sup_state_def] | |
| 149 | "(G \\<turnstile> (xt, a) <=s (y#yt, b)) = (\\<exists>x xt'. xt=x#xt' \\<and> (G \\<turnstile> x \\<preceq> y) \\<and> (G \\<turnstile> (xt',a) <=s (yt,b)))"; | |
| 150 | by (auto_tac (claset(), simpset() addsimps [sup_loc_Cons2])); | |
| 151 | bd map_hd_tl 1; | |
| 152 | by Auto_tac; | |
| 153 | qed "sup_state_Cons2"; | |
| 154 | ||
| 155 | Goalw[sup_state_def] | |
| 156 | "G \\<turnstile> (a, x) <=s (b, y) \\<Longrightarrow> G \\<turnstile> (c, x) <=s (c, y)"; | |
| 157 | by Auto_tac; | |
| 158 | qed "sup_state_ignore_fst"; | |
| 159 | ||
| 160 | Goalw[sup_ty_opt_def] | |
| 161 | "\\<lbrakk>G \\<turnstile> a <=o b; G \\<turnstile> b <=o c\\<rbrakk> \\<Longrightarrow> G \\<turnstile> a <=o c"; | |
| 8442 
96023903c2df
case_tac now subsumes both boolean and datatype cases;
 wenzelm parents: 
8423diff
changeset | 162 | by (case_tac "c" 1); | 
| 8394 | 163 | by (Clarsimp_tac 1); | 
| 164 | by (Clarsimp_tac 1); | |
| 8442 
96023903c2df
case_tac now subsumes both boolean and datatype cases;
 wenzelm parents: 
8423diff
changeset | 165 | by (case_tac "a" 1); | 
| 8394 | 166 | by (Clarsimp_tac 1); | 
| 8442 
96023903c2df
case_tac now subsumes both boolean and datatype cases;
 wenzelm parents: 
8423diff
changeset | 167 | by (case_tac "b" 1); | 
| 8394 | 168 | by (Clarsimp_tac 1); | 
| 169 | by (Clarsimp_tac 1); | |
| 8442 
96023903c2df
case_tac now subsumes both boolean and datatype cases;
 wenzelm parents: 
8423diff
changeset | 170 | by (case_tac "b" 1); | 
| 8394 | 171 | by (Clarsimp_tac 1); | 
| 172 | by (Clarsimp_tac 1); | |
| 173 | be widen_trans 1; | |
| 174 | ba 1; | |
| 175 | qed "sup_ty_opt_trans"; | |
| 176 | ||
| 177 | Goal "\\<lbrakk>G \\<turnstile> a <=l b; G \\<turnstile> b <=l c\\<rbrakk> \\<Longrightarrow> \\<forall> n. n < length a \\<longrightarrow> G \\<turnstile> a!n <=o c!n"; | |
| 178 | by (Clarify_tac 1); | |
| 179 | by (forward_tac [sup_loc_length] 1); | |
| 180 | bd sup_loc_nth 1; | |
| 181 | ba 1; | |
| 182 | bd sup_loc_nth 1; | |
| 183 | by (Force_tac 1); | |
| 184 | bd sup_ty_opt_trans 1; | |
| 185 | ba 1; | |
| 186 | ba 1; | |
| 187 | qed "sup_loc_all_trans"; | |
| 188 | ||
| 189 | Goal "\\<forall> b. length a = length b \\<longrightarrow> (\\<forall> n. n < length a \\<longrightarrow> (G \\<turnstile> (a!n) <=o (b!n))) \\<longrightarrow> (G \\<turnstile> a <=l b)"; | |
| 190 | by (induct_tac "a" 1); | |
| 191 | by (Clarsimp_tac 1); | |
| 192 | by (Clarsimp_tac 1); | |
| 8423 | 193 | by (case_tac "b=[]" 1); | 
| 8394 | 194 | by (Clarsimp_tac 1); | 
| 195 | by (Clarsimp_tac 1); | |
| 196 | by (asm_full_simp_tac (simpset() addsimps [neq_Nil_conv]) 1); | |
| 197 | by (Clarify_tac 1); | |
| 198 | by (res_inst_tac [("x","y")] exI 1);
 | |
| 199 | by (res_inst_tac [("x","ys")] exI 1);
 | |
| 200 | by (eres_inst_tac [("x","ys")] allE 1);
 | |
| 201 | by (Full_simp_tac 1); | |
| 202 | by (Clarify_tac 1); | |
| 203 | be impE 1; | |
| 204 | by (Clarsimp_tac 1); | |
| 205 |  by (eres_inst_tac [("x","Suc n")] allE 1);
 | |
| 206 | by (Clarsimp_tac 1); | |
| 207 | by (eres_inst_tac [("x","0")] allE 1);
 | |
| 208 | by (Clarsimp_tac 1); | |
| 209 | qed_spec_mp "all_nth_sup_loc"; | |
| 210 | ||
| 211 | Goal "\\<lbrakk>G \\<turnstile> a <=l b; G \\<turnstile> b <=l c\\<rbrakk> \\<Longrightarrow> G \\<turnstile> a <=l c"; | |
| 212 | by (forward_tac [sup_loc_all_trans] 1); | |
| 213 | ba 1; | |
| 214 | bd sup_loc_length 1; | |
| 215 | bd sup_loc_length 1; | |
| 216 | br all_nth_sup_loc 1; | |
| 217 | by (Clarsimp_tac 1); | |
| 218 | ba 1; | |
| 219 | qed "sup_loc_trans"; | |
| 220 | ||
| 221 | Goalw[sup_state_def] | |
| 222 | "\\<lbrakk>G \\<turnstile> a <=s b; G \\<turnstile> b <=s c\\<rbrakk> \\<Longrightarrow> G \\<turnstile> a <=s c"; | |
| 223 | by Safe_tac; | |
| 224 | br sup_loc_trans 1; | |
| 225 | ba 1; | |
| 226 | ba 1; | |
| 227 | br sup_loc_trans 1; | |
| 228 | ba 1; | |
| 229 | ba 1; | |
| 230 | qed "sup_state_trans"; | |
| 231 | ||
| 232 | Goalw[sup_ty_opt_def] "G \\<turnstile> a <=o Some b \\<Longrightarrow> \\<exists> x. a = Some x"; | |
| 8442 
96023903c2df
case_tac now subsumes both boolean and datatype cases;
 wenzelm parents: 
8423diff
changeset | 233 | by (case_tac "a" 1); | 
| 8394 | 234 | by Auto_tac; | 
| 235 | qed "sup_ty_opt_some"; | |
| 236 | ||
| 237 | ||
| 238 | Goalw[sup_loc_def] | |
| 239 | "\\<forall> n y. (G \\<turnstile> a <=o b) \\<longrightarrow> n < length y \\<longrightarrow> (G \\<turnstile> x <=l y) \\<longrightarrow> (G \\<turnstile> x[n := a] <=l y[n := b])"; | |
| 240 | by (induct_tac "x" 1); | |
| 241 | by (Simp_tac 1); | |
| 242 | by (clarsimp_tac ((claset(), simpset()) addIffs [sup_loc_Cons2]) 1); | |
| 8442 
96023903c2df
case_tac now subsumes both boolean and datatype cases;
 wenzelm parents: 
8423diff
changeset | 243 | by (case_tac "n" 1); | 
| 8394 | 244 | by (clarsimp_tac (claset(), simpset() addsimps [list_all2_Cons1]) 1); | 
| 245 | by (clarsimp_tac (claset(), simpset() addsimps [list_all2_Cons1]) 1); | |
| 246 | qed_spec_mp "sup_loc_update"; |