Automatic injection of Bexeq
authorCezary Kaliszyk <kaliszyk@in.tum.de>
Thu, 21 Jan 2010 12:03:47 +0100
changeset 911 95ee248b3832
parent 910 b91782991dc8
child 912 aa960d16570f
Automatic injection of Bexeq
Quot/Examples/IntEx.thy
Quot/QuotMain.thy
Quot/quotient_tacs.ML
--- a/Quot/Examples/IntEx.thy	Thu Jan 21 11:11:22 2010 +0100
+++ b/Quot/Examples/IntEx.thy	Thu Jan 21 12:03:47 2010 +0100
@@ -191,10 +191,6 @@
 
 lemma ex1tst': "\<exists>! (x :: my_int). x = x"
 apply(lifting ex1tst)
-
-apply injection
-apply (rule quot_respect(9))
-apply (rule Quotient_my_int)
 done
 
 lemma ho_tst: "foldl my_plus x [] = x"
--- a/Quot/QuotMain.thy	Thu Jan 21 11:11:22 2010 +0100
+++ b/Quot/QuotMain.thy	Thu Jan 21 12:03:47 2010 +0100
@@ -101,7 +101,7 @@
 declare [[map "fun" = (fun_map, fun_rel)]]
 
 lemmas [quot_thm] = fun_quotient
-lemmas [quot_respect] = quot_rel_rsp bexeq_rsp
+lemmas [quot_respect] = quot_rel_rsp
 lemmas [quot_equiv] = identity_equivp
 
 (* Lemmas about simplifying id's. *)
--- a/Quot/quotient_tacs.ML	Thu Jan 21 11:11:22 2010 +0100
+++ b/Quot/quotient_tacs.ML	Thu Jan 21 12:03:47 2010 +0100
@@ -358,6 +358,10 @@
     (Const(@{const_name Bex1},_) $ (Const (@{const_name Respects}, _) $ _))
       => rtac @{thm fun_rel_id} THEN' quot_true_tac ctxt unlam
 
+| (Const (@{const_name fun_rel}, _) $ _ $ _) $
+    (Const(@{const_name Bexeq},_) $ _) $ (Const(@{const_name Bexeq},_) $ _)
+      => rtac @{thm bexeq_rsp} THEN' quotient_tac ctxt
+
 | (_ $
     (Const(@{const_name Babs},_) $ (Const (@{const_name Respects}, _) $ _) $ _) $
     (Const(@{const_name Babs},_) $ (Const (@{const_name Respects}, _) $ _) $ _))