| author | nipkow | 
| Fri, 23 Aug 2002 07:41:05 +0200 | |
| changeset 13517 | 42efec18f5b2 | 
| parent 13145 | 59bc43b51aa2 | 
| child 14401 | 477380c74c1d | 
| permissions | -rw-r--r-- | 
| 5323 | 1 | (* Title: HOL/Lex/RegExp2NA.ML | 
| 2 | ID: $Id$ | |
| 3 | Author: Tobias Nipkow | |
| 4 | Copyright 1998 TUM | |
| 5 | *) | |
| 6 | ||
| 7 | (******************************************************) | |
| 8 | (* atom *) | |
| 9 | (******************************************************) | |
| 10 | ||
| 11 | Goalw [atom_def] "(fin (atom a) q) = (q = [False])"; | |
| 6162 | 12 | by (Simp_tac 1); | 
| 5323 | 13 | qed "fin_atom"; | 
| 14 | ||
| 15 | Goalw [atom_def] "start (atom a) = [True]"; | |
| 6162 | 16 | by (Simp_tac 1); | 
| 5323 | 17 | qed "start_atom"; | 
| 18 | ||
| 19 | Goalw [atom_def,step_def] | |
| 20 | "(p,q) : step (atom a) b = (p=[True] & q=[False] & b=a)"; | |
| 6162 | 21 | by (Simp_tac 1); | 
| 5323 | 22 | qed "in_step_atom_Some"; | 
| 23 | Addsimps [in_step_atom_Some]; | |
| 24 | ||
| 25 | Goal | |
| 26 | "([False],[False]) : steps (atom a) w = (w = [])"; | |
| 27 | by (induct_tac "w" 1); | |
| 6162 | 28 | by (Simp_tac 1); | 
| 12487 | 29 | by (asm_simp_tac (simpset() addsimps [rel_comp_def]) 1); | 
| 5323 | 30 | qed "False_False_in_steps_atom"; | 
| 31 | ||
| 32 | Goal | |
| 33 | "(start (atom a), [False]) : steps (atom a) w = (w = [a])"; | |
| 34 | by (induct_tac "w" 1); | |
| 6162 | 35 | by (asm_simp_tac (simpset() addsimps [start_atom]) 1); | 
| 36 | by (asm_full_simp_tac (simpset() | |
| 12487 | 37 | addsimps [False_False_in_steps_atom,rel_comp_def,start_atom]) 1); | 
| 5323 | 38 | qed "start_fin_in_steps_atom"; | 
| 39 | ||
| 40 | Goal | |
| 41 | "accepts (atom a) w = (w = [a])"; | |
| 6162 | 42 | by (simp_tac(simpset() addsimps | 
| 5323 | 43 | [accepts_conv_steps,start_fin_in_steps_atom,fin_atom]) 1); | 
| 44 | qed "accepts_atom"; | |
| 45 | ||
| 46 | ||
| 47 | (******************************************************) | |
| 12792 | 48 | (* or *) | 
| 5323 | 49 | (******************************************************) | 
| 50 | ||
| 51 | (***** True/False ueber fin anheben *****) | |
| 52 | ||
| 12792 | 53 | Goalw [or_def] | 
| 54 | "!L R. fin (or L R) (True#p) = fin L p"; | |
| 5323 | 55 | by (Simp_tac 1); | 
| 12792 | 56 | qed_spec_mp "fin_or_True"; | 
| 5323 | 57 | |
| 12792 | 58 | Goalw [or_def] | 
| 59 | "!L R. fin (or L R) (False#p) = fin R p"; | |
| 5323 | 60 | by (Simp_tac 1); | 
| 12792 | 61 | qed_spec_mp "fin_or_False"; | 
| 5323 | 62 | |
| 12792 | 63 | AddIffs [fin_or_True,fin_or_False]; | 
| 5323 | 64 | |
| 65 | (***** True/False ueber step anheben *****) | |
| 66 | ||
| 12792 | 67 | Goalw [or_def,step_def] | 
| 68 | "!L R. (True#p,q) : step (or L R) a = (? r. q = True#r & (p,r) : step L a)"; | |
| 5323 | 69 | by (Simp_tac 1); | 
| 6162 | 70 | by (Blast_tac 1); | 
| 12792 | 71 | qed_spec_mp "True_in_step_or"; | 
| 5323 | 72 | |
| 12792 | 73 | Goalw [or_def,step_def] | 
| 74 | "!L R. (False#p,q) : step (or L R) a = (? r. q = False#r & (p,r) : step R a)"; | |
| 5323 | 75 | by (Simp_tac 1); | 
| 6162 | 76 | by (Blast_tac 1); | 
| 12792 | 77 | qed_spec_mp "False_in_step_or"; | 
| 5323 | 78 | |
| 12792 | 79 | AddIffs [True_in_step_or,False_in_step_or]; | 
| 5323 | 80 | |
| 81 | ||
| 82 | (***** True/False ueber steps anheben *****) | |
| 83 | ||
| 84 | Goal | |
| 12792 | 85 | "!p. (True#p,q):steps (or L R) w = (? r. q = True # r & (p,r):steps L w)"; | 
| 5323 | 86 | by (induct_tac "w" 1); | 
| 5758 
27a2b36efd95
corrected auto_tac (applications of unsafe wrappers)
 oheimb parents: 
5525diff
changeset | 87 | by (ALLGOALS Force_tac); | 
| 12792 | 88 | qed_spec_mp "lift_True_over_steps_or"; | 
| 5323 | 89 | |
| 90 | Goal | |
| 12792 | 91 | "!p. (False#p,q):steps (or L R) w = (? r. q = False#r & (p,r):steps R w)"; | 
| 5323 | 92 | by (induct_tac "w" 1); | 
| 5758 
27a2b36efd95
corrected auto_tac (applications of unsafe wrappers)
 oheimb parents: 
5525diff
changeset | 93 | by (ALLGOALS Force_tac); | 
| 12792 | 94 | qed_spec_mp "lift_False_over_steps_or"; | 
| 5323 | 95 | |
| 12792 | 96 | AddIffs [lift_True_over_steps_or,lift_False_over_steps_or]; | 
| 5323 | 97 | |
| 98 | ||
| 99 | (** From the start **) | |
| 100 | ||
| 12792 | 101 | Goalw [or_def,step_def] | 
| 102 | "!L R. (start(or L R),q) : step(or L R) a = \ | |
| 5323 | 103 | \ (? p. (q = True#p & (start L,p) : step L a) | \ | 
| 104 | \ (q = False#p & (start R,p) : step R a))"; | |
| 6162 | 105 | by (Simp_tac 1); | 
| 106 | by (Blast_tac 1); | |
| 12792 | 107 | qed_spec_mp "start_step_or"; | 
| 108 | AddIffs [start_step_or]; | |
| 5323 | 109 | |
| 110 | Goal | |
| 12792 | 111 | "(start(or L R), q) : steps (or L R) w = \ | 
| 112 | \ ( (w = [] & q = start(or L R)) | \ | |
| 5323 | 113 | \ (w ~= [] & (? p. q = True # p & (start L,p) : steps L w | \ | 
| 114 | \ q = False # p & (start R,p) : steps R w)))"; | |
| 8442 
96023903c2df
case_tac now subsumes both boolean and datatype cases;
 wenzelm parents: 
8423diff
changeset | 115 | by (case_tac "w" 1); | 
| 5323 | 116 | by (Asm_simp_tac 1); | 
| 6162 | 117 | by (Blast_tac 1); | 
| 5323 | 118 | by (Asm_simp_tac 1); | 
| 6162 | 119 | by (Blast_tac 1); | 
| 12792 | 120 | qed "steps_or"; | 
| 5323 | 121 | |
| 12792 | 122 | Goalw [or_def] | 
| 123 | "!L R. fin (or L R) (start(or L R)) = \ | |
| 5323 | 124 | \ (fin L (start L) | fin R (start R))"; | 
| 6162 | 125 | by (Simp_tac 1); | 
| 12792 | 126 | qed_spec_mp "fin_start_or"; | 
| 127 | AddIffs [fin_start_or]; | |
| 5323 | 128 | |
| 129 | Goal | |
| 12792 | 130 | "accepts (or L R) w = (accepts L w | accepts R w)"; | 
| 131 | by (simp_tac (simpset() addsimps [accepts_conv_steps,steps_or]) 1); | |
| 5323 | 132 | (* get rid of case_tac: *) | 
| 6162 | 133 | by (case_tac "w = []" 1); | 
| 134 | by (Auto_tac); | |
| 12792 | 135 | qed "accepts_or"; | 
| 136 | AddIffs [accepts_or]; | |
| 5323 | 137 | |
| 138 | (******************************************************) | |
| 139 | (* conc *) | |
| 140 | (******************************************************) | |
| 141 | ||
| 142 | (** True/False in fin **) | |
| 143 | ||
| 144 | Goalw [conc_def] | |
| 145 | "!L R. fin (conc L R) (True#p) = (fin L p & fin R (start R))"; | |
| 146 | by (Simp_tac 1); | |
| 147 | qed_spec_mp "fin_conc_True"; | |
| 148 | ||
| 149 | Goalw [conc_def] | |
| 150 | "!L R. fin (conc L R) (False#p) = fin R p"; | |
| 151 | by (Simp_tac 1); | |
| 152 | qed "fin_conc_False"; | |
| 153 | ||
| 154 | AddIffs [fin_conc_True,fin_conc_False]; | |
| 155 | ||
| 156 | (** True/False in step **) | |
| 157 | ||
| 158 | Goalw [conc_def,step_def] | |
| 159 | "!L R. (True#p,q) : step (conc L R) a = \ | |
| 160 | \ ((? r. q=True#r & (p,r): step L a) | \ | |
| 161 | \ (fin L p & (? r. q=False#r & (start R,r) : step R a)))"; | |
| 162 | by (Simp_tac 1); | |
| 6162 | 163 | by (Blast_tac 1); | 
| 5323 | 164 | qed_spec_mp "True_step_conc"; | 
| 165 | ||
| 166 | Goalw [conc_def,step_def] | |
| 167 | "!L R. (False#p,q) : step (conc L R) a = \ | |
| 168 | \ (? r. q = False#r & (p,r) : step R a)"; | |
| 169 | by (Simp_tac 1); | |
| 6162 | 170 | by (Blast_tac 1); | 
| 5323 | 171 | qed_spec_mp "False_step_conc"; | 
| 172 | ||
| 173 | AddIffs [True_step_conc, False_step_conc]; | |
| 174 | ||
| 175 | (** False in steps **) | |
| 176 | ||
| 177 | Goal | |
| 178 | "!p. (False#p,q): steps (conc L R) w = (? r. q=False#r & (p,r): steps R w)"; | |
| 179 | by (induct_tac "w" 1); | |
| 5758 
27a2b36efd95
corrected auto_tac (applications of unsafe wrappers)
 oheimb parents: 
5525diff
changeset | 180 | by (ALLGOALS Force_tac); | 
| 5323 | 181 | qed_spec_mp "False_steps_conc"; | 
| 182 | AddIffs [False_steps_conc]; | |
| 183 | ||
| 184 | (** True in steps **) | |
| 185 | ||
| 186 | Goal | |
| 187 | "!!L R. !p. (p,q) : steps L w --> (True#p,True#q) : steps (conc L R) w"; | |
| 6162 | 188 | by (induct_tac "w" 1); | 
| 5323 | 189 | by (Simp_tac 1); | 
| 190 | by (Simp_tac 1); | |
| 6162 | 191 | by (Blast_tac 1); | 
| 5323 | 192 | qed_spec_mp "True_True_steps_concI"; | 
| 193 | ||
| 194 | Goal | |
| 195 | "!L R. (True#p,False#q) : step (conc L R) a = \ | |
| 196 | \ (fin L p & (start R,q) : step R a)"; | |
| 6162 | 197 | by (Simp_tac 1); | 
| 5323 | 198 | qed "True_False_step_conc"; | 
| 199 | AddIffs [True_False_step_conc]; | |
| 200 | ||
| 201 | Goal | |
| 202 | "!p. (True#p,q) : steps (conc L R) w --> \ | |
| 203 | \ ((? r. (p,r) : steps L w & q = True#r) | \ | |
| 204 | \ (? u a v. w = u@a#v & \ | |
| 205 | \ (? r. (p,r) : steps L u & fin L r & \ | |
| 206 | \ (? s. (start R,s) : step R a & \ | |
| 207 | \ (? t. (s,t) : steps R v & q = False#t)))))"; | |
| 6162 | 208 | by (induct_tac "w" 1); | 
| 209 | by (Simp_tac 1); | |
| 210 | by (Simp_tac 1); | |
| 211 | by (clarify_tac (claset() delrules [disjCI]) 1); | |
| 212 | by (etac disjE 1); | |
| 213 | by (clarify_tac (claset() delrules [disjCI]) 1); | |
| 214 | by (etac allE 1 THEN mp_tac 1); | |
| 215 | by (etac disjE 1); | |
| 5323 | 216 | by (Blast_tac 1); | 
| 6162 | 217 | by (rtac disjI2 1); | 
| 5323 | 218 | by (Clarify_tac 1); | 
| 6162 | 219 | by (Simp_tac 1); | 
| 220 |  by (res_inst_tac[("x","a#u")] exI 1);
 | |
| 221 | by (Simp_tac 1); | |
| 5323 | 222 | by (Blast_tac 1); | 
| 6162 | 223 | by (rtac disjI2 1); | 
| 5323 | 224 | by (Clarify_tac 1); | 
| 6162 | 225 | by (Simp_tac 1); | 
| 226 | by (res_inst_tac[("x","[]")] exI 1);
 | |
| 227 | by (Simp_tac 1); | |
| 5323 | 228 | by (Blast_tac 1); | 
| 229 | qed_spec_mp "True_steps_concD"; | |
| 230 | ||
| 231 | Goal | |
| 232 | "(True#p,q) : steps (conc L R) w = \ | |
| 233 | \ ((? r. (p,r) : steps L w & q = True#r) | \ | |
| 234 | \ (? u a v. w = u@a#v & \ | |
| 235 | \ (? r. (p,r) : steps L u & fin L r & \ | |
| 236 | \ (? s. (start R,s) : step R a & \ | |
| 237 | \ (? t. (s,t) : steps R v & q = False#t)))))"; | |
| 6162 | 238 | by (force_tac (claset() addDs [True_steps_concD] | 
| 5525 | 239 | addIs [True_True_steps_concI],simpset()) 1); | 
| 5323 | 240 | qed "True_steps_conc"; | 
| 241 | ||
| 242 | (** starting from the start **) | |
| 243 | ||
| 244 | Goalw [conc_def] | |
| 245 | "!L R. start(conc L R) = True#start L"; | |
| 6162 | 246 | by (Simp_tac 1); | 
| 5323 | 247 | qed_spec_mp "start_conc"; | 
| 248 | ||
| 249 | Goalw [conc_def] | |
| 250 | "!L R. fin(conc L R) p = ((fin R (start R) & (? s. p = True#s & fin L s)) | \ | |
| 251 | \ (? s. p = False#s & fin R s))"; | |
| 252 | by (simp_tac (simpset() addsplits [list.split]) 1); | |
| 253 | by (Blast_tac 1); | |
| 254 | qed_spec_mp "final_conc"; | |
| 255 | ||
| 256 | Goal | |
| 257 | "accepts (conc L R) w = (? u v. w = u@v & accepts L u & accepts R v)"; | |
| 258 | by (simp_tac (simpset() addsimps | |
| 259 | [accepts_conv_steps,True_steps_conc,final_conc,start_conc]) 1); | |
| 6162 | 260 | by (rtac iffI 1); | 
| 5323 | 261 | by (Clarify_tac 1); | 
| 6162 | 262 | by (etac disjE 1); | 
| 5323 | 263 | by (Clarify_tac 1); | 
| 6162 | 264 | by (etac disjE 1); | 
| 265 |    by (res_inst_tac [("x","w")] exI 1);
 | |
| 266 | by (Simp_tac 1); | |
| 267 | by (Blast_tac 1); | |
| 268 | by (Blast_tac 1); | |
| 269 | by (etac disjE 1); | |
| 270 | by (Blast_tac 1); | |
| 5323 | 271 | by (Clarify_tac 1); | 
| 6162 | 272 |  by (res_inst_tac [("x","u")] exI 1);
 | 
| 273 | by (Simp_tac 1); | |
| 274 | by (Blast_tac 1); | |
| 5323 | 275 | by (Clarify_tac 1); | 
| 8442 
96023903c2df
case_tac now subsumes both boolean and datatype cases;
 wenzelm parents: 
8423diff
changeset | 276 | by (case_tac "v" 1); | 
| 6162 | 277 | by (Asm_full_simp_tac 1); | 
| 278 | by (Blast_tac 1); | |
| 279 | by (Asm_full_simp_tac 1); | |
| 280 | by (Blast_tac 1); | |
| 5323 | 281 | qed "accepts_conc"; | 
| 282 | ||
| 283 | (******************************************************) | |
| 284 | (* epsilon *) | |
| 285 | (******************************************************) | |
| 286 | ||
| 287 | Goalw [epsilon_def,step_def] "step epsilon a = {}";
 | |
| 6162 | 288 | by (Simp_tac 1); | 
| 5323 | 289 | qed "step_epsilon"; | 
| 290 | Addsimps [step_epsilon]; | |
| 291 | ||
| 292 | Goal "((p,q) : steps epsilon w) = (w=[] & p=q)"; | |
| 6162 | 293 | by (induct_tac "w" 1); | 
| 294 | by (Auto_tac); | |
| 5323 | 295 | qed "steps_epsilon"; | 
| 296 | ||
| 297 | Goal "accepts epsilon w = (w = [])"; | |
| 6162 | 298 | by (simp_tac (simpset() addsimps [steps_epsilon,accepts_conv_steps]) 1); | 
| 299 | by (simp_tac (simpset() addsimps [epsilon_def]) 1); | |
| 5323 | 300 | qed "accepts_epsilon"; | 
| 301 | AddIffs [accepts_epsilon]; | |
| 302 | ||
| 303 | (******************************************************) | |
| 304 | (* plus *) | |
| 305 | (******************************************************) | |
| 306 | ||
| 307 | Goalw [plus_def] "!A. start (plus A) = start A"; | |
| 6162 | 308 | by (Simp_tac 1); | 
| 5323 | 309 | qed_spec_mp "start_plus"; | 
| 310 | Addsimps [start_plus]; | |
| 311 | ||
| 312 | Goalw [plus_def] "!A. fin (plus A) = fin A"; | |
| 6162 | 313 | by (Simp_tac 1); | 
| 5323 | 314 | qed_spec_mp "fin_plus"; | 
| 315 | AddIffs [fin_plus]; | |
| 316 | ||
| 317 | Goalw [plus_def,step_def] | |
| 318 | "!A. (p,q) : step A a --> (p,q) : step (plus A) a"; | |
| 6162 | 319 | by (Simp_tac 1); | 
| 5323 | 320 | qed_spec_mp "step_plusI"; | 
| 321 | ||
| 322 | Goal "!p. (p,q) : steps A w --> (p,q) : steps (plus A) w"; | |
| 6162 | 323 | by (induct_tac "w" 1); | 
| 324 | by (Simp_tac 1); | |
| 325 | by (Simp_tac 1); | |
| 326 | by (blast_tac (claset() addIs [step_plusI]) 1); | |
| 5323 | 327 | qed_spec_mp "steps_plusI"; | 
| 328 | ||
| 329 | Goalw [plus_def,step_def] | |
| 330 | "!A. (p,r): step (plus A) a = \ | |
| 331 | \ ( (p,r): step A a | fin A p & (start A,r) : step A a )"; | |
| 6162 | 332 | by (Simp_tac 1); | 
| 5323 | 333 | qed_spec_mp "step_plus_conv"; | 
| 334 | AddIffs [step_plus_conv]; | |
| 335 | ||
| 336 | Goal | |
| 337 | "[| (start A,q) : steps A u; u ~= []; fin A p |] \ | |
| 338 | \ ==> (p,q) : steps (plus A) u"; | |
| 8442 
96023903c2df
case_tac now subsumes both boolean and datatype cases;
 wenzelm parents: 
8423diff
changeset | 339 | by (case_tac "u" 1); | 
| 6162 | 340 | by (Blast_tac 1); | 
| 341 | by (Asm_full_simp_tac 1); | |
| 342 | by (blast_tac (claset() addIs [steps_plusI]) 1); | |
| 5323 | 343 | qed "fin_steps_plusI"; | 
| 344 | ||
| 345 | (* reverse list induction! Complicates matters for conc? *) | |
| 346 | Goal | |
| 347 | "!r. (start A,r) : steps (plus A) w --> \ | |
| 348 | \ (? us v. w = concat us @ v & \ | |
| 349 | \ (!u:set us. accepts A u) & \ | |
| 350 | \ (start A,r) : steps A v)"; | |
| 13145 | 351 | by (res_inst_tac [("xs","w")] rev_induct 1);
 | 
| 5323 | 352 | by (Simp_tac 1); | 
| 6162 | 353 |  by (res_inst_tac [("x","[]")] exI 1);
 | 
| 5323 | 354 | by (Simp_tac 1); | 
| 355 | by (Simp_tac 1); | |
| 356 | by (Clarify_tac 1); | |
| 6162 | 357 | by (etac allE 1 THEN mp_tac 1); | 
| 5323 | 358 | by (Clarify_tac 1); | 
| 6162 | 359 | by (etac disjE 1); | 
| 360 |  by (res_inst_tac [("x","us")] exI 1);
 | |
| 361 | by (Asm_simp_tac 1); | |
| 362 | by (Blast_tac 1); | |
| 363 | by (res_inst_tac [("x","us@[v]")] exI 1);
 | |
| 364 | by (asm_full_simp_tac (simpset() addsimps [accepts_conv_steps]) 1); | |
| 365 | by (Blast_tac 1); | |
| 5323 | 366 | qed_spec_mp "start_steps_plusD"; | 
| 367 | ||
| 368 | Goal | |
| 369 | "us ~= [] --> (!u : set us. accepts A u) --> accepts (plus A) (concat us)"; | |
| 6162 | 370 | by (simp_tac (simpset() addsimps [accepts_conv_steps]) 1); | 
| 13145 | 371 | by (res_inst_tac [("xs","us")] rev_induct 1);
 | 
| 6162 | 372 | by (Simp_tac 1); | 
| 373 | by (rename_tac "u us" 1); | |
| 374 | by (Simp_tac 1); | |
| 5323 | 375 | by (Clarify_tac 1); | 
| 6162 | 376 | by (case_tac "us = []" 1); | 
| 377 | by (Asm_full_simp_tac 1); | |
| 378 | by (blast_tac (claset() addIs [steps_plusI,fin_steps_plusI]) 1); | |
| 5323 | 379 | by (Clarify_tac 1); | 
| 6162 | 380 | by (case_tac "u = []" 1); | 
| 381 | by (Asm_full_simp_tac 1); | |
| 382 | by (blast_tac (claset() addIs [steps_plusI,fin_steps_plusI]) 1); | |
| 383 | by (Asm_full_simp_tac 1); | |
| 384 | by (blast_tac (claset() addIs [steps_plusI,fin_steps_plusI]) 1); | |
| 5323 | 385 | qed_spec_mp "steps_star_cycle"; | 
| 386 | ||
| 387 | Goal | |
| 388 | "accepts (plus A) w = \ | |
| 389 | \ (? us. us ~= [] & w = concat us & (!u : set us. accepts A u))"; | |
| 6162 | 390 | by (rtac iffI 1); | 
| 391 | by (asm_full_simp_tac (simpset() addsimps [accepts_conv_steps]) 1); | |
| 5323 | 392 | by (Clarify_tac 1); | 
| 6162 | 393 | by (dtac start_steps_plusD 1); | 
| 5323 | 394 | by (Clarify_tac 1); | 
| 6162 | 395 |  by (res_inst_tac [("x","us@[v]")] exI 1);
 | 
| 396 | by (asm_full_simp_tac (simpset() addsimps [accepts_conv_steps]) 1); | |
| 397 | by (Blast_tac 1); | |
| 398 | by (blast_tac (claset() addIs [steps_star_cycle]) 1); | |
| 5323 | 399 | qed "accepts_plus"; | 
| 400 | AddIffs [accepts_plus]; | |
| 401 | ||
| 402 | (******************************************************) | |
| 403 | (* star *) | |
| 404 | (******************************************************) | |
| 405 | ||
| 406 | Goalw [star_def] | |
| 407 | "accepts (star A) w = \ | |
| 408 | \ (? us. (!u : set us. accepts A u) & w = concat us)"; | |
| 6162 | 409 | by (rtac iffI 1); | 
| 5323 | 410 | by (Clarify_tac 1); | 
| 6162 | 411 | by (etac disjE 1); | 
| 412 |   by (res_inst_tac [("x","[]")] exI 1);
 | |
| 413 | by (Simp_tac 1); | |
| 414 | by (Blast_tac 1); | |
| 415 | by (Blast_tac 1); | |
| 416 | by (Force_tac 1); | |
| 5323 | 417 | qed "accepts_star"; | 
| 418 | ||
| 419 | (***** Correctness of r2n *****) | |
| 420 | ||
| 421 | Goal | |
| 422 | "!w. accepts (rexp2na r) w = (w : lang r)"; | |
| 6162 | 423 | by (induct_tac "r" 1); | 
| 424 | by (simp_tac (simpset() addsimps [accepts_conv_steps]) 1); | |
| 425 | by (simp_tac(simpset() addsimps [accepts_atom]) 1); | |
| 426 | by (Asm_simp_tac 1); | |
| 427 | by (asm_simp_tac (simpset() addsimps [accepts_conc,RegSet.conc_def]) 1); | |
| 428 | by (asm_simp_tac (simpset() addsimps [accepts_star,in_star]) 1); | |
| 5323 | 429 | qed_spec_mp "accepts_rexp2na"; |