# HG changeset patch # User Cezary Kaliszyk # Date 1264071827 -3600 # Node ID 95ee248b3832ce3ae09efd6ffc45c6b4cbaf836a # Parent b91782991dc864339aeda5d64fb61e6625e1d4e1 Automatic injection of Bexeq diff -r b91782991dc8 -r 95ee248b3832 Quot/Examples/IntEx.thy --- 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': "\! (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" diff -r b91782991dc8 -r 95ee248b3832 Quot/QuotMain.thy --- 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. *) diff -r b91782991dc8 -r 95ee248b3832 Quot/quotient_tacs.ML --- 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}, _) $ _) $ _))