QuotMain.thy
changeset 487 f5db9ede89b0
parent 486 7c26b31d2371
child 488 508f3106b89c
--- a/QuotMain.thy	Wed Dec 02 17:16:44 2009 +0100
+++ b/QuotMain.thy	Wed Dec 02 20:54:59 2009 +0100
@@ -604,16 +604,15 @@
       (let
         val gl = Const (@{const_name "EQUIV"}, dummyT) $ R2;
         val glc = HOLogic.mk_Trueprop (Syntax.check_term ctxt gl);
-        val _ = tracing (Syntax.string_of_term ctxt glc);
+(*        val _ = tracing (Syntax.string_of_term ctxt glc);*)
         val eqv = Goal.prove ctxt [] [] glc (fn _ => equiv_tac rel_eqvs 1);
         val thm = (@{thm eq_reflection} OF [@{thm ball_reg_eqv_range} OF [eqv]]);
         val R1c = cterm_of thy R1;
         val thmi = Drule.instantiate' [] [SOME R1c] thm;
-(*        val _ = tracing (Syntax.string_of_term ctxt (prop_of thmi)); *)
+(*        val _ = tracing (Syntax.string_of_term ctxt (prop_of thmi));*)
         val inst = matching_prs thy (term_of (Thm.lhs_of thmi)) ogl
-        val _ = tracing "AAA";
         val thm2 = Drule.eta_contraction_rule (Drule.instantiate inst thmi);
-        val _ = tracing (Syntax.string_of_term ctxt (prop_of thm2));
+(*        val _ = tracing (Syntax.string_of_term ctxt (prop_of thm2)); *)
       in
         SOME thm2
       end
@@ -636,16 +635,15 @@
       (let
         val gl = Const (@{const_name "EQUIV"}, dummyT) $ R2;
         val glc = HOLogic.mk_Trueprop (Syntax.check_term ctxt gl);
-        val _ = tracing (Syntax.string_of_term ctxt glc);
+(*        val _ = tracing (Syntax.string_of_term ctxt glc); *)
         val eqv = Goal.prove ctxt [] [] glc (fn _ => equiv_tac rel_eqvs 1);
         val thm = (@{thm eq_reflection} OF [@{thm bex_reg_eqv_range} OF [eqv]]);
         val R1c = cterm_of thy R1;
         val thmi = Drule.instantiate' [] [SOME R1c] thm;
 (*        val _ = tracing (Syntax.string_of_term ctxt (prop_of thmi)); *)
         val inst = matching_prs thy (term_of (Thm.lhs_of thmi)) ogl
-        val _ = tracing "AAA";
         val thm2 = Drule.eta_contraction_rule (Drule.instantiate inst thmi);
-        val _ = tracing (Syntax.string_of_term ctxt (prop_of thm2));
+(*        val _ = tracing (Syntax.string_of_term ctxt (prop_of thm2));*)
       in
         SOME thm2
       end