# HG changeset patch # User Cezary Kaliszyk # Date 1260045939 -3600 # Node ID b460565dbb58074913d98a7b70d13d6526323fee # Parent 8395fc6a694596b125ee077c1f166b3df2578a22 Handling of respects in the fast inj_repabs_tac; includes respects with quotient assumptions. diff -r 8395fc6a6945 -r b460565dbb58 IntEx.thy --- a/IntEx.thy Sat Dec 05 21:28:24 2009 +0100 +++ b/IntEx.thy Sat Dec 05 21:45:39 2009 +0100 @@ -176,28 +176,8 @@ apply(rule ballI) apply(rule apply_rsp[OF Quotient_my_int ho_plus_rsp]) apply(simp only: in_respects) - -apply(tactic {* inj_repabs_tac_intex @{context} 1*}) -apply(tactic {* inj_repabs_tac_intex @{context} 1*}) -apply(tactic {* inj_repabs_tac_intex @{context} 1*}) -apply(tactic {* inj_repabs_tac_intex @{context} 1*}) -apply(tactic {* inj_repabs_tac_intex @{context} 1*}) -apply(rule quot_rel_rsp) -apply(tactic {* quotient_tac @{context} 1 *}) -apply(tactic {* inj_repabs_tac_intex @{context} 1*}) -apply(tactic {* inj_repabs_tac_intex @{context} 1*}) -apply(tactic {* inj_repabs_tac_intex @{context} 1*}) -apply(tactic {* inj_repabs_tac_intex @{context} 1*}) -apply(tactic {* inj_repabs_tac_intex @{context} 1*}) -apply(tactic {* inj_repabs_tac_intex @{context} 1*}) -apply(tactic {* inj_repabs_tac_intex @{context} 1*}) -apply(tactic {* inj_repabs_tac_intex @{context} 1*}) -apply(tactic {* inj_repabs_tac_intex @{context} 1*}) -apply(tactic {* inj_repabs_tac_intex @{context} 1*}) -apply(fold PLUS_def) +apply(tactic {* all_inj_repabs_tac_intex @{context} 1*}) apply(tactic {* clean_tac @{context} 1 *}) -thm Quotient_rel_rep -thm Quotient_rel_rep[OF fun_quotient[OF Quotient_my_int Quotient_my_int]] apply(simp only: Quotient_rel_rep[OF fun_quotient[OF Quotient_my_int Quotient_my_int]]) done @@ -208,14 +188,7 @@ lemma "PLUS = PLUS" apply(tactic {* procedure_tac @{context} @{thm test3} 1 *}) apply(rule ho_plus_rsp) -apply(tactic {* inj_repabs_tac_intex @{context} 1*}) -apply(tactic {* inj_repabs_tac_intex @{context} 1*}) -apply(rule quot_rel_rsp) -apply(tactic {* quotient_tac @{context} 1 *}) -apply(tactic {* inj_repabs_tac_intex @{context} 1*}) -apply(tactic {* inj_repabs_tac_intex @{context} 1*}) -apply(tactic {* inj_repabs_tac_intex @{context} 1*}) -apply(tactic {* inj_repabs_tac_intex @{context} 1*}) +apply(tactic {* all_inj_repabs_tac_intex @{context} 1*}) apply(tactic {* clean_tac @{context} 1 *}) apply(simp only: Quotient_rel_rep[OF fun_quotient[OF Quotient_my_int fun_quotient[OF Quotient_my_int Quotient_my_int]]]) done diff -r 8395fc6a6945 -r b460565dbb58 QuotMain.thy --- a/QuotMain.thy Sat Dec 05 21:28:24 2009 +0100 +++ b/QuotMain.thy Sat Dec 05 21:45:39 2009 +0100 @@ -148,6 +148,9 @@ lemmas [quotient_thm] = fun_quotient list_quotient +lemmas [quotient_rsp] = + quot_rel_rsp + ML {* maps_lookup @{theory} "List.list" *} ML {* maps_lookup @{theory} "*" *} ML {* maps_lookup @{theory} "fun" *} @@ -981,6 +984,13 @@ (Const(@{const_name Ball},_) $ (Const (@{const_name Respects}, _) $ _) $ _) => rtac @{thm fun_rel_id} THEN' quot_true_tac ctxt unlam (* (op =) (Bex\) (Bex\) ----> (op =) (\) (\) *) +| (Const (@{const_name fun_rel}, _) $ _ $ _) $ + (Const (@{const_name fun_rel}, _) $ _ $ _) $ + (Const (@{const_name fun_rel}, _) $ _ $ _) + => rtac @{thm quot_rel_rsp} THEN_ALL_NEW quotient_tac ctxt +| (Const (@{const_name fun_rel}, _) $ _ $ _) $ (Const _) $ (Const _) => + (* respectfulness of constants; in particular of a simple relation *) + resolve_tac (rsp_rules_get ctxt) THEN_ALL_NEW quotient_tac ctxt | Const (@{const_name "op ="},_) $ (Const(@{const_name Bex},_) $ (Const (@{const_name Respects}, _) $ _) $ _) $ (Const(@{const_name Bex},_) $ (Const (@{const_name Respects}, _) $ _) $ _) @@ -1020,9 +1030,7 @@ NDT ctxt "E" (atac), (* reflexivity of the basic relations *) (* R \ \ *) - NDT ctxt "D" (resolve_tac rel_refl), - (* respectfulness of constants *) - NDT ctxt "7" (resolve_tac (rsp_rules_get ctxt)) + NDT ctxt "D" (resolve_tac rel_refl) ]) *}