src/CCL/ex/Stream.thy
changeset 41526 54b4686704af
parent 39159 0dec18004e75
child 42155 ffe99b07c9c0
--- a/src/CCL/ex/Stream.thy	Wed Jan 12 15:53:37 2011 +0100
+++ b/src/CCL/ex/Stream.thy	Wed Jan 12 16:33:04 2011 +0100
@@ -63,7 +63,7 @@
   shows "map(f,l@m) = map(f,l) @ map(f,m)"
   apply (tactic {* eq_coinduct3_tac @{context}
     "{p. EX x y. p=<x,y> & (EX l:Lists (A). EX m:Lists (A). x=map (f,l@m) & y=map (f,l) @ map (f,m))}" 1 *})
-  apply (blast intro: prems)
+  apply (blast intro: assms)
   apply safe
   apply (drule ListsXH [THEN iffD1])
   apply (tactic "EQgen_tac @{context} [] 1")
@@ -82,7 +82,7 @@
   shows "k @ l @ m = (k @ l) @ m"
   apply (tactic {* eq_coinduct3_tac @{context}
     "{p. EX x y. p=<x,y> & (EX k:Lists (A). EX l:Lists (A). EX m:Lists (A). x=k @ l @ m & y= (k @ l) @ m) }" 1*})
-  apply (blast intro: prems)
+  apply (blast intro: assms)
   apply safe
   apply (drule ListsXH [THEN iffD1])
   apply (tactic "EQgen_tac @{context} [] 1")
@@ -100,7 +100,7 @@
   shows "l @ m = l"
   apply (tactic {* eq_coinduct3_tac @{context}
     "{p. EX x y. p=<x,y> & (EX l:ILists (A) .EX m. x=l@m & y=l)}" 1 *})
-  apply (blast intro: prems)
+  apply (blast intro: assms)
   apply safe
   apply (drule IListsXH [THEN iffD1])
   apply (tactic "EQgen_tac @{context} [] 1")