--- a/QuotMain.thy Mon Nov 02 09:33:48 2009 +0100
+++ b/QuotMain.thy Mon Nov 02 09:39:29 2009 +0100
@@ -506,6 +506,10 @@
else
Const (@{const_name "Ex"}, ty) $ apply_subt (my_reg lthy rel rty) t
end
+ | Const (@{const_name "op ="}, ty) $ t =>
+ if needs_lift rty (fastype_of t) then
+ (tyRel (fastype_of t) rty rel lthy) $ t
+ else Const (@{const_name "op ="}, ty) $ (my_reg lthy rel rty t)
| t1 $ t2 => (my_reg lthy rel rty t1) $ (my_reg lthy rel rty t2)
| _ => trm
*}
@@ -518,16 +522,32 @@
(my_reg lthy rel rty (prop_of thm)))
*}
+lemma universal_twice: "(\<And>x. (P x \<longrightarrow> Q x)) \<Longrightarrow> ((\<forall>x. P x) \<longrightarrow> (\<forall>x. Q x))"
+apply (auto)
+done
+
+lemma implication_twice: "(c \<longrightarrow> a) \<Longrightarrow> (a \<Longrightarrow> b \<longrightarrow> d) \<Longrightarrow> (a \<longrightarrow> b) \<longrightarrow> (c \<longrightarrow> d)"
+apply (auto)
+done
+
ML {*
-fun regularize thm rty rel rel_eqv lthy =
+fun regularize thm rty rel rel_eqv rel_refl lthy =
let
val g = build_regularize_goal thm rty rel lthy;
fun tac ctxt =
- (asm_full_simp_tac ((Simplifier.context ctxt HOL_ss) addsimps
+ (ObjectLogic.full_atomize_tac) THEN'
+ REPEAT_ALL_NEW (FIRST' [
+ rtac rel_refl,
+ atac,
+ rtac @{thm universal_twice},
+ (rtac @{thm impI} THEN' atac),
+ rtac @{thm implication_twice},
+ EqSubst.eqsubst_tac ctxt [0]
[(@{thm equiv_res_forall} OF [rel_eqv]),
- (@{thm equiv_res_exists} OF [rel_eqv])])) THEN_ALL_NEW
- (((rtac @{thm RIGHT_RES_FORALL_REGULAR}) THEN' (RANGE [fn _ => all_tac, atac]) THEN'
- (MetisTools.metis_tac ctxt [])) ORELSE' (MetisTools.metis_tac ctxt []));
+ (@{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})
+ ]);
val cthm = Goal.prove lthy [] [] g (fn x => tac (#context x) 1);
in
cthm OF [thm]
@@ -551,14 +571,10 @@
)
*}
-ML {* make_const_def *}
-
ML {*
fun old_get_fun flag rty qty lthy ty =
get_fun flag [(qty, rty)] lthy ty
-*}
-ML {*
fun old_make_const_def nconst_bname otrm mx rty qty lthy =
make_const_def nconst_bname otrm mx [(qty, rty)] lthy
*}
@@ -838,26 +854,25 @@
Abs(_, T, b) =>
findaps_all rty (subst_bound ((Free ("x", T)), b))
| (f $ a) => (findaps_all rty f @ findaps_all rty a)
- | Free (_, (T as (Type (_, (_ :: _))))) => (if needs_lift rty T then [T] else [])
+ | Free (_, (T as (Type ("fun", (_ :: _))))) => (if needs_lift rty T then [T] else [])
| _ => [];
fun findaps rty tm = distinct (op =) (findaps_all rty tm)
*}
ML {*
-fun make_simp_lam_prs_thm lthy quot_thm typ =
+fun make_simp_prs_thm lthy quot_thm thm typ =
let
val (_, [lty, rty]) = dest_Type typ;
val thy = ProofContext.theory_of lthy;
val (lcty, rcty) = (ctyp_of thy lty, ctyp_of thy rty)
val inst = [SOME lcty, NONE, SOME rcty];
- val lpi = Drule.instantiate' inst [] @{thm LAMBDA_PRS};
+ val lpi = Drule.instantiate' inst [] thm;
val tac =
- (compose_tac (false, @{thm LAMBDA_PRS}, 2)) THEN_ALL_NEW
+ (compose_tac (false, lpi, 2)) THEN_ALL_NEW
(quotient_tac quot_thm);
val t = Goal.prove lthy [] [] (concl_of lpi) (fn _ => tac 1);
- val ts = @{thm HOL.sym} OF [t]
in
- MetaSimplifier.rewrite_rule [@{thm eq_reflection} OF @{thms id_apply}] ts
+ MetaSimplifier.rewrite_rule [@{thm eq_reflection} OF @{thms id_apply}] t
end
*}
@@ -916,11 +931,13 @@
val (trans2, reps_same, quot) = lookup_quot_thms lthy qty_name;
val consts = lookup_quot_consts defs;
val t_a = atomize_thm t;
- val t_r = regularize t_a rty rel rel_eqv lthy;
+ val t_r = regularize t_a rty rel rel_eqv rel_refl lthy;
val t_t = repabs lthy t_r consts rty qty quot rel_refl trans2 rsp_thms;
val abs = findabs rty (prop_of t_a);
- val simp_lam_prs_thms = map (make_simp_lam_prs_thm lthy quot) abs;
- val t_l = repeat_eqsubst_thm lthy simp_lam_prs_thms t_t;
+ val aps = findaps rty (prop_of t_a);
+ val app_prs_thms = map (make_simp_prs_thm lthy quot @{thm APP_PRS}) aps;
+ val lam_prs_thms = map (make_simp_prs_thm lthy quot @{thm LAMBDA_PRS}) abs;
+ val t_l = repeat_eqsubst_thm lthy (lam_prs_thms @ app_prs_thms) t_t;
val t_a = simp_allex_prs lthy quot t_l;
val defs_sym = add_lower_defs lthy defs;
val t_d = repeat_eqsubst_thm lthy defs_sym t_a;