QuotMain.thy
changeset 253 e169a99c6ada
parent 252 e30997c88050
child 255 264c7b9d19f4
--- a/QuotMain.thy	Fri Oct 30 19:03:53 2009 +0100
+++ b/QuotMain.thy	Sat Oct 31 11:20:55 2009 +0100
@@ -535,10 +535,11 @@
       rtac rel_refl,
       atac,
       rtac @{thm universal_twice},
+      (rtac @{thm impI} THEN' atac),
       rtac @{thm implication_twice},
-      (fn i => CHANGED (asm_full_simp_tac ((Simplifier.context ctxt HOL_ss) addsimps
+      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})
      ]);
@@ -848,26 +849,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
 *}
 
@@ -929,8 +929,10 @@
   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;