Experiments in Int
authorCezary Kaliszyk <kaliszyk@in.tum.de>
Wed, 04 Nov 2009 11:08:05 +0100
changeset 276 783d6c940e45
parent 275 34ad627ac5d5
child 278 bdedfb51778d
Experiments in Int
FSet.thy
IntEx.thy
--- a/FSet.thy	Wed Nov 04 10:43:33 2009 +0100
+++ b/FSet.thy	Wed Nov 04 11:08:05 2009 +0100
@@ -333,7 +333,7 @@
 term list_rec
 
 quotient_def
-  fset_rec::"'a \<Rightarrow> ('b \<Rightarrow> 'b fset \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> 'a fset \<Rightarrow> 'b"
+  fset_rec::"'a \<Rightarrow> ('b \<Rightarrow> 'b fset \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> 'b fset \<Rightarrow> 'a"
 where
   "fset_rec \<equiv> list_rec"
 
@@ -351,8 +351,8 @@
 ML {* val (trans2, reps_same, absrep, quot) = lookup_quot_thms @{context} "fset" *}
 
 
-ML {* val ind_r_a = atomize_thm @{thm card1_suc} *}
-(* prove {* build_regularize_goal ind_r_a rty rel @{context}  *}
+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' [
       rtac rel_refl,
--- a/IntEx.thy	Wed Nov 04 10:43:33 2009 +0100
+++ b/IntEx.thy	Wed Nov 04 11:08:05 2009 +0100
@@ -43,6 +43,7 @@
 
 term PLUS
 thm PLUS_def
+ML {* prop_of @{thm PLUS_def} *}
 
 fun
   my_neg :: "(nat \<times> nat) \<Rightarrow> (nat \<times> nat)"
@@ -128,10 +129,6 @@
   lift_thm lthy qty ty_name rsp_thms defs t
 *}
 
-ML {*
-  val (rty, rel, rel_refl, rel_eqv) = lookup_quot_data @{context} qty;
-*}
-
 ML {* lift_thm_my_int @{context} @{thm plus_sym_pre} *}
 
 lemma plus_assoc_pre:
@@ -151,18 +148,43 @@
 
 
 
+
+
 text {* Below is the construction site code used if things do not work *}
-ML {* val (trans2, reps_same, quot) = lookup_quot_thms @{context} "my_int" *}
+ML {* val (rty, rel, rel_refl, rel_eqv) = lookup_quot_data @{context} qty *}
+ML {* val (trans2, reps_same, absrep, quot) = lookup_quot_thms @{context} "my_int" *}
 ML {* val abs = findabs rty (prop_of (atomize_thm @{thm plus_sym_pre})) *}
 ML {* val simp_lam_prs_thms = map (make_simp_prs_thm @{context} quot @{thm LAMBDA_PRS}) abs *}
 ML {* val defs_sym = add_lower_defs @{context} defs *}
+(* ML {* val consts = [@{const_name my_plus}] *}*)
 ML {* val consts = lookup_quot_consts defs *}
-ML {* val t_a = atomize_thm @{thm plus_sym_pre} *}
-ML {* val t_r = regularize t_a rty rel rel_eqv rel_refl @{context} *}
-ML {* val t_t = repabs @{context} t_r consts rty qty quot rel_refl trans2 rsp_thms *}
-ML {* val t_l = repeat_eqsubst_thm @{context} simp_lam_prs_thms t_t *}
-ML {* val t_a = simp_allex_prs quot [] t_l *}
-ML {* val t_d = repeat_eqsubst_thm @{context} defs_sym t_a *}
+ML {* val t_a = atomize_thm @{thm ho_tst} *}
+prove t_r: {* build_regularize_goal t_a rty rel @{context} *}
+ML_prf {*   fun tac ctxt =
+      (ObjectLogic.full_atomize_tac) THEN'
+     REPEAT_ALL_NEW (FIRST' [
+      rtac rel_refl,
+      atac,
+      rtac @{thm universal_twice},
+      (rtac @{thm impI} THEN' atac),
+      (*rtac @{thm equality_twice},*)
+      EqSubst.eqsubst_tac ctxt [0]
+        [(@{thm equiv_res_forall} OF [rel_eqv]),
+         (@{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 *})
+apply (auto)
+sorry
+(*ML {* val t_r = regularize t_a rty rel rel_eqv rel_refl @{context} *}*)
+ML {* val t_t = Toplevel.program (fn () => repabs @{context} @{thm t_r} consts rty qty quot rel_refl trans2 rsp_thms) *}
+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 t_a = MetaSimplifier.rewrite_rule (allthms) t_t *}
+ML {* val t_l = repeat_eqsubst_thm @{context} simp_lam_prs_thms t_a *}
+ML {* val t_d = repeat_eqsubst_thm @{context} defs_sym t_l *}
 ML {* val t_r = MetaSimplifier.rewrite_rule [reps_same] t_d *}
 ML {* ObjectLogic.rulify t_r *}