diff -r 86fba2c4eeef -r 57bde65f6eb2 IntEx.thy --- a/IntEx.thy Wed Nov 25 10:52:21 2009 +0100 +++ b/IntEx.thy Wed Nov 25 11:41:42 2009 +0100 @@ -141,10 +141,6 @@ ML {* val consts = lookup_quot_consts defs *} ML {* -fun lift_thm_my_int lthy t = - lift_thm lthy qty ty_name rsp_thms defs t -fun lift_thm_g_my_int lthy t g = - lift_thm_goal lthy qty ty_name rsp_thms defs t g fun lift_tac_fset lthy t = lift_tac lthy t rel_eqv rel_refl rty quot trans2 rsp_thms reps_same defs *} @@ -164,6 +160,21 @@ by (tactic {* lift_tac_fset @{context} @{thm plus_assoc_pre} 1 *}) +lemma ho_tst: "foldl my_plus x [] = x" +apply simp +done + +lemma "foldl PLUS x [] = x" +apply (tactic {* lift_tac_fset @{context} @{thm ho_tst} 1 *}) +prefer 3 +apply(tactic {* clean_tac @{context} quot @{thms PLUS_def} reps_same 1 *}) +sorry + +(* + FIXME: All below is your construction code; mostly commented out as it + does not work. +*) + ML {* mk_REGULARIZE_goal @{context} @{term "\i j k. my_plus (my_plus i j) k \ my_plus i (my_plus j k)"} @@ -172,13 +183,6 @@ |> writeln *} - -ML {* -val (rty, rel, rel_refl, rel_eqv) = lookup_quot_data @{context} @{typ "my_int"} -val (trans2, reps_same, absrep, quot) = lookup_quot_thms @{context} "my_int" -*} - - lemma "PLUS (PLUS i j) k = PLUS i (PLUS j k)" apply(tactic {* procedure_tac @{thm plus_assoc_pre} @{context} 1 *}) apply(tactic {* regularize_tac @{context} rel_eqv rel_refl 1 *}) @@ -188,7 +192,7 @@ (* -does not work. + ML {* fun r_mk_comb_tac ctxt rty quot_thm reflex_thm trans_thm rsp_thms = (REPEAT1 o FIRST1) @@ -209,14 +213,6 @@ DT ctxt "E" (WEAK_LAMBDA_RES_TAC ctxt), DT ctxt "F" (CHANGED' (asm_full_simp_tac (HOL_ss addsimps @{thms FUN_REL_EQ})))] *} -*) - -ML {* -val (rty, rel, rel_refl, rel_eqv) = lookup_quot_data @{context} @{typ "my_int"} -val (trans2, reps_same, absrep, quot) = lookup_quot_thms @{context} "my_int" -val consts = lookup_quot_consts defs -*} - ML {* mk_inj_REPABS_goal @{context} (reg_atrm, aqtrm) @@ -238,66 +234,6 @@ |> writeln *} - -lemma ho_tst: "foldl my_plus x [] = x" -apply simp -done - -text {* Below is the construction site code used if things do not work *} -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 consts = [@{const_name my_plus}] *}*) -ML {* val consts = lookup_quot_consts defs *} -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) -done -ML {* val t_r = @{thm t_r} OF [t_a] *}*) -ML {* val t_r = regularize t_a rty rel rel_eqv rel_refl @{context} *} -ML {* - val rt = build_repabs_term @{context} t_r consts rty qty - val rg = Logic.mk_equals ((Thm.prop_of t_r), rt); -*} +*) - -prove t_t: rg -apply(atomize(full)) -ML_prf {* fun r_mk_comb_tac_int lthy = r_mk_comb_tac lthy rty quot rel_refl trans2 rsp_thms *} -apply (tactic {* REPEAT_ALL_NEW (r_mk_comb_tac_int @{context}) 1 *}) -done -ML {* val t_t = @{thm Pure.equal_elim_rule1} OF [@{thm t_t},t_r] *} -ML {* val abs = findabs rty (prop_of t_a) *} -ML {* val aps = findaps rty (prop_of t_a); *} -ML {* val simp_lam_prs_thms = map (make_simp_prs_thm @{context} quot @{thm LAMBDA_PRS}) abs *} - -(*ML {* val t_t = repabs @{context} @{thm t_r} consts rty qty quot rel_refl trans2 rsp_thms *}*) -ML findallex -ML {* val (alls, exs) = findallex @{context} 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 defs_sym = flat (map (add_lower_defs @{context}) defs) *} -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 *} -ML {* @{term "Trueprop"} *} -