src/HOL/W0/Maybe.ML
changeset 5069 3ea049f7979d
parent 4831 dae4d63a1318
child 9476 2210dffb9764
--- a/src/HOL/W0/Maybe.ML	Mon Jun 22 17:13:09 1998 +0200
+++ b/src/HOL/W0/Maybe.ML	Mon Jun 22 17:26:46 1998 +0200
@@ -6,25 +6,25 @@
 
 
 (* constructor laws for bind *)
-goalw thy [bind_def] "(Ok s) bind f = (f s)";
+Goalw [bind_def] "(Ok s) bind f = (f s)";
 by (Simp_tac 1);
 qed "bind_Ok";
 
-goalw thy [bind_def] "Fail bind f = Fail";
+Goalw [bind_def] "Fail bind f = Fail";
 by (Simp_tac 1);
 qed "bind_Fail";
 
 Addsimps [bind_Ok,bind_Fail];
 
 (* case splitting of bind *)
-goal thy
+Goal
   "P(res bind f) = ((res = Fail --> P Fail) & (!s. res = Ok s --> P(f s)))";
 by (induct_tac "res" 1);
 by (fast_tac (HOL_cs addss simpset()) 1);
 by (Asm_simp_tac 1);
 qed "split_bind";
 
-goal Maybe.thy
+Goal
   "((m bind f) = Fail) = ((m=Fail) | (? p. m = Ok p & f p = Fail))";
 by (simp_tac (simpset() addsplits [split_bind]) 1);
 qed "bind_eq_Fail";