--- a/src/HOL/Lex/Automata.ML Mon Jun 22 17:13:09 1998 +0200
+++ b/src/HOL/Lex/Automata.ML Mon Jun 22 17:26:46 1998 +0200
@@ -6,14 +6,14 @@
(*** Equivalence of NA and DA --- redundant ***)
-goalw thy [na2da_def]
+Goalw [na2da_def]
"!Q. DA.delta (na2da A) w Q = Union(NA.delta A w `` Q)";
by(induct_tac "w" 1);
by(ALLGOALS Asm_simp_tac);
by(ALLGOALS Blast_tac);
qed_spec_mp "DA_delta_is_lift_NA_delta";
-goalw thy [DA.accepts_def,NA.accepts_def]
+Goalw [DA.accepts_def,NA.accepts_def]
"NA.accepts A w = DA.accepts (na2da A) w";
by(simp_tac (simpset() addsimps [DA_delta_is_lift_NA_delta]) 1);
by(simp_tac (simpset() addsimps [na2da_def]) 1);
@@ -21,7 +21,7 @@
(*** Direct equivalence of NAe and DA ***)
-goalw thy [nae2da_def]
+Goalw [nae2da_def]
"!Q. (eps A)^* ^^ (DA.delta (nae2da A) w Q) = steps A w ^^ Q";
by(induct_tac "w" 1);
by (Simp_tac 1);
@@ -29,12 +29,12 @@
by(Blast_tac 1);
qed_spec_mp "espclosure_DA_delta_is_steps";
-goalw thy [nae2da_def]
+Goalw [nae2da_def]
"fin (nae2da A) Q = (? q : (eps A)^* ^^ Q. fin A q)";
by(Simp_tac 1);
val lemma = result();
-goalw thy [NAe.accepts_def,DA.accepts_def]
+Goalw [NAe.accepts_def,DA.accepts_def]
"DA.accepts (nae2da A) w = NAe.accepts A w";
by(simp_tac (simpset() addsimps [lemma,espclosure_DA_delta_is_steps]) 1);
by(simp_tac (simpset() addsimps [nae2da_def]) 1);