FSet.thy
changeset 285 8ebdef196fd5
parent 276 783d6c940e45
child 289 7e8617f20b59
child 290 a0be84b0c707
--- a/FSet.thy	Wed Nov 04 16:10:39 2009 +0100
+++ b/FSet.thy	Thu Nov 05 09:38:34 2009 +0100
@@ -317,9 +317,6 @@
 ML {* fun lift_thm_fset lthy t = lift_thm lthy qty "fset" rsp_thms defs t *}
 
 (* ML {* lift_thm_fset @{context} @{thm neq_Nil_conv} *} *)
-
-
-
 ML {* lift_thm_fset @{context} @{thm m1} *}
 ML {* lift_thm_fset @{context} @{thm m2} *}
 ML {* lift_thm_fset @{context} @{thm list_eq.intros(4)} *}
@@ -351,49 +348,59 @@
 ML {* val (trans2, reps_same, absrep, quot) = lookup_quot_thms @{context} "fset" *}
 
 
-ML {* val ind_r_a = atomize_thm @{thm map_append} *}
- prove {* build_regularize_goal ind_r_a rty rel @{context}  *}
- ML_prf {*  fun tac ctxt =
-     (FIRST' [
+ML {* val t_a = atomize_thm @{thm map_append} *}
+(* prove {* build_regularize_goal t_a rty rel @{context}  *}
+ ML_prf {*  fun tac ctxt = FIRST' [
       rtac rel_refl,
       atac,
-      rtac @{thm get_rid},
-      rtac @{thm get_rid2},
-      (fn i => CHANGED (asm_full_simp_tac ((Simplifier.context ctxt HOL_ss) addsimps
+      rtac @{thm universal_twice},
+      (rtac @{thm impI} THEN' atac),
+      rtac @{thm implication_twice},
+      (*rtac @{thm equality_twice},*)
+      EqSubst.eqsubst_tac ctxt [0]
         [(@{thm equiv_res_forall} OF [rel_eqv]),
-         (@{thm equiv_res_exists} OF [rel_eqv])]) i)),
+         (@{thm equiv_res_exists} OF [rel_eqv])],
       (rtac @{thm impI} THEN' (asm_full_simp_tac (Simplifier.context ctxt HOL_ss)) THEN' rtac rel_refl),
       (rtac @{thm RIGHT_RES_FORALL_REGULAR})
-    ]);
- *}
+     ]; *}
   apply (atomize(full))
-  apply (tactic {* tac @{context} 1 *}) *)
-ML {* val ind_r_r = regularize ind_r_a rty rel rel_eqv rel_refl @{context} *}
+  apply (tactic {* REPEAT_ALL_NEW (tac @{context}) 1 *})
+  done*)
+
+ML {* val t_r = regularize t_a rty rel rel_eqv rel_refl @{context} *}
 ML {*
-  val rt = build_repabs_term @{context} ind_r_r consts rty qty
-  val rg = Logic.mk_equals ((Thm.prop_of ind_r_r), rt);
+  val rt = build_repabs_term @{context} t_r consts rty qty
+  val rg = Logic.mk_equals ((Thm.prop_of t_r), rt);
 *}
+
+ML {* val ttt = snd (dest_Free (term_of @{cpat "f :: ?'a list \<Rightarrow> ?'a list"})) *}
+ML {* old_get_fun repF @{typ "'a list"} @{typ "'a fset"} @{context} @{typ "'a list \<Rightarrow> 'a list"} *}
+
 prove rg
 apply(atomize(full))
 ML_prf {* fun r_mk_comb_tac_fset lthy = r_mk_comb_tac lthy rty quot rel_refl trans2 rsp_thms *}
 apply (tactic {* REPEAT_ALL_NEW (r_mk_comb_tac_fset @{context}) 1 *})
 done
-ML {* val ind_r_t =
+ML {* val t_t =
   Toplevel.program (fn () =>
-  repabs @{context} ind_r_r consts rty qty quot rel_refl trans2 rsp_thms
+  repabs @{context} t_r consts rty qty quot rel_refl trans2 rsp_thms
   )
 *}
 
 ML {* val abs = findabs rty (prop_of (atomize_thm @{thm list.induct})) *}
 ML {* val aps = findaps rty (prop_of (atomize_thm @{thm list.induct})) *}
-ML {* val simp_lam_prs_thms = map (make_simp_prs_thm @{context} quot @{thm LAMBDA_PRS}) abs *}
-ML {* val ind_r_l = repeat_eqsubst_thm @{context} (simp_app_prs_thms @ simp_lam_prs_thms) ind_r_t *}
-ML {* val thm = @{thm FORALL_PRS[OF FUN_QUOTIENT[OF QUOTIENT_fset IDENTITY_QUOTIENT]]} *}
-ML {* val ind_r_a = simp_allex_prs quot [thm] ind_r_l *}
+ML {* val lam_prs_thms = map (make_simp_prs_thm @{context} quot @{thm LAMBDA_PRS}) abs *}
+ML {* val app_prs_thms = map (applic_prs @{context} rty qty absrep) aps *}
+ML {* val app_prs_thms = map Thm.freezeT app_prs_thms *}
+ML {* val t_l = repeat_eqsubst_thm @{context} (app_prs_thms @ lam_prs_thms) t_t *}
+ML {* val (alls, exs) = findallex rty qty (prop_of t_a); *}
+ML {* val allthms = map (make_allex_prs_thm @{context} quot @{thm FORALL_PRS}) alls *}
+ML {* val allthmsv = map Thm.varifyT allthms *}
+ML {* val t_a = MetaSimplifier.rewrite_rule (allthmsv) t_l *}
 ML {* val defs_sym = add_lower_defs @{context} defs *}
-ML {* val ind_r_d = repeat_eqsubst_thm @{context} defs_sym ind_r_a *}
-ML {* val ind_r_s = MetaSimplifier.rewrite_rule @{thms QUOT_TYPE_I_fset.REPS_same} ind_r_d *}
-ML {* ObjectLogic.rulify ind_r_s *}
+ML {* val t_d = repeat_eqsubst_thm @{context} defs_sym t_a *}
+ML {* val t_s = MetaSimplifier.rewrite_rule @{thms QUOT_TYPE_I_fset.REPS_same} t_d *}
+ML {* ObjectLogic.rulify t_s *}
 
 ML {*
   fun lift_thm_fset_note name thm lthy =