FSet.thy
changeset 508 fac6069d8e80
parent 506 91c374abde06
child 509 d62b6517a0ab
--- a/FSet.thy	Thu Dec 03 15:03:31 2009 +0100
+++ b/FSet.thy	Fri Dec 04 08:18:38 2009 +0100
@@ -135,9 +135,6 @@
        if (a memb A) then (fold1 f g z A) else (f (g a) (fold1 f g z A))
      ) else z)"
 
-(* fold1_def is not usable, but: *)
-thm fold1.simps
-
 lemma fs1_strong_cases:
   fixes X :: "'a list"
   shows "(X = []) \<or> (\<exists>a. \<exists> Y. (~(a memb Y) \<and> (X \<approx> a # Y)))"
@@ -434,8 +431,8 @@
   "INSERT2 \<equiv> op #"
 
 ML {* val quot = @{thms QUOTIENT_fset QUOTIENT_fset2} *}
-ML {* fun inj_repabs_tac_fset lthy = inj_repabs_tac lthy quot [rel_refl] [trans2] *}
-ML {* fun lift_tac_fset lthy t = lift_tac lthy t [rel_eqv] quot *}
+ML {* fun inj_repabs_tac_fset lthy = inj_repabs_tac lthy [rel_refl] [trans2] *}
+ML {* fun lift_tac_fset lthy t = lift_tac lthy t [rel_eqv] *}
 
 lemma "P (x :: 'a fset2) (EMPTY :: 'c fset) \<Longrightarrow> (\<And>e t. P x t \<Longrightarrow> P x (INSERT e t)) \<Longrightarrow> P x l"
 apply (tactic {* lift_tac_fset @{context} @{thm list_induct_part} 1 *})
@@ -469,7 +466,7 @@
   sorry
 
 ML {* val rsp_thms = @{thms list_rec_rsp list_case_rsp} @ rsp_thms *}
-ML {* fun lift_tac_fset lthy t = lift_tac lthy t [rel_eqv] quot *}
+ML {* fun lift_tac_fset lthy t = lift_tac lthy t [rel_eqv] *}
 
 lemma "fset_rec (f1::'t) x (INSERT a xa) = x a xa (fset_rec f1 x xa)"
 apply (tactic {* lift_tac_fset @{context} @{thm list.recs(2)} 1 *})