--- a/src/FOL/ex/Quantifiers_Cla.thy Tue Oct 06 13:31:44 2015 +0200
+++ b/src/FOL/ex/Quantifiers_Cla.thy Tue Oct 06 15:14:28 2015 +0200
@@ -58,11 +58,11 @@
apply fast?
oops
-schematic_lemma "P(?a) --> (ALL x. P(x))"
+schematic_goal "P(?a) --> (ALL x. P(x))"
apply fast?
oops
-schematic_lemma "(P(?a) --> (ALL x. Q(x))) --> (ALL x. P(x) --> Q(x))"
+schematic_goal "(P(?a) --> (ALL x. Q(x))) --> (ALL x. P(x) --> Q(x))"
apply fast?
oops
@@ -76,7 +76,7 @@
lemma "(P --> (EX x. Q(x))) & P --> (EX x. Q(x))"
by fast
-schematic_lemma "(ALL x. P(x)-->Q(f(x))) & (ALL x. Q(x)-->R(g(x))) & P(d) --> R(?a)"
+schematic_goal "(ALL x. P(x)-->Q(f(x))) & (ALL x. Q(x)-->R(g(x))) & P(d) --> R(?a)"
by fast
lemma "(ALL x. Q(x)) --> (EX x. Q(x))"