--- a/Quot/Examples/FSet.thy Sun Dec 20 00:15:40 2009 +0100
+++ b/Quot/Examples/FSet.thy Sun Dec 20 00:26:53 2009 +0100
@@ -22,9 +22,9 @@
apply(auto intro: list_eq.intros list_eq_refl)
done
-quotient fset = "'a list" / "list_eq"
- apply(rule equivp_list_eq)
- done
+quotient_type
+ fset = "'a list" / "list_eq"
+ by (rule equivp_list_eq)
quotient_def
"EMPTY :: 'a fset"
@@ -271,7 +271,7 @@
apply (lifting list_induct_part)
done
-quotient fset2 = "'a list" / "list_eq"
+quotient_type fset2 = "'a list" / "list_eq"
by (rule equivp_list_eq)
quotient_def