diff -r 114bb544ecb9 -r d641c32855f0 QuotMain.thy --- a/QuotMain.thy Sat Dec 05 21:47:48 2009 +0100 +++ b/QuotMain.thy Sat Dec 05 21:50:31 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) ]) *}