--- a/doc-src/TutorialI/CTL/CTL.thy Thu Oct 12 17:54:22 2000 +0200
+++ b/doc-src/TutorialI/CTL/CTL.thy Thu Oct 12 18:06:31 2000 +0200
@@ -84,7 +84,7 @@
apply(rule equalityI);
apply(rule subsetI);
apply(simp);
- apply(erule Lfp.induct);
+ apply(erule lfp_induct);
apply(rule mono_ef);
apply(simp);
apply(blast intro: r_into_rtrancl rtrancl_trans);
@@ -112,7 +112,7 @@
*};
apply(rule subsetI);
-apply(erule Lfp.induct[OF _ mono_af]);
+apply(erule lfp_induct[OF _ mono_af]);
apply(clarsimp simp add: af_def Paths_def);
(*ML"Pretty.setmargin 70";
pr(latex xsymbols symbols);*)
--- a/doc-src/TutorialI/CTL/PDL.thy Thu Oct 12 17:54:22 2000 +0200
+++ b/doc-src/TutorialI/CTL/PDL.thy Thu Oct 12 18:06:31 2000 +0200
@@ -106,7 +106,7 @@
which is proved by @{term lfp}-induction:
*}
- apply(erule Lfp.induct)
+ apply(erule lfp_induct)
apply(rule mono_ef)
apply(simp)
(*pr(latex xsymbols symbols);*)