diff -r a88e16b69d0f -r 9d35c6145dd2 Quot/QuotBase.thy --- a/Quot/QuotBase.thy Thu Jan 28 10:52:10 2010 +0100 +++ b/Quot/QuotBase.thy Thu Jan 28 12:24:49 2010 +0100 @@ -451,19 +451,19 @@ using a unfolding Quotient_def Bex_def in_respects fun_map_def id_apply by metis -section {* Bexeq quantifier *} +section {* Bex1_rel quantifier *} definition - Bexeq :: "('a \ 'a \ bool) \ ('a \ bool) \ bool" + Bex1_rel :: "('a \ 'a \ bool) \ ('a \ bool) \ bool" where - "Bexeq R P \ (\x \ Respects R. P x) \ (\x \ Respects R. \y \ Respects R. ((P x \ P y) \ (R x y)))" + "Bex1_rel R P \ (\x \ Respects R. P x) \ (\x \ Respects R. \y \ Respects R. ((P x \ P y) \ (R x y)))" (* TODO/FIXME: simplify the repetitions in this proof *) lemma bexeq_rsp: assumes a: "Quotient R absf repf" - shows "((R ===> op =) ===> op =) (Bexeq R) (Bexeq R)" + shows "((R ===> op =) ===> op =) (Bex1_rel R) (Bex1_rel R)" apply simp -unfolding Bexeq_def +unfolding Bex1_rel_def apply rule apply rule apply rule @@ -500,9 +500,9 @@ lemma ex1_prs: assumes a: "Quotient R absf repf" - shows "((absf ---> id) ---> id) (Bexeq R) f = Ex1 f" + shows "((absf ---> id) ---> id) (Bex1_rel R) f = Ex1 f" apply simp -apply (subst Bexeq_def) +apply (subst Bex1_rel_def) apply (subst Bex_def) apply (subst Ex1_def) apply simp @@ -535,8 +535,8 @@ apply metis done -lemma bex1_bexeq_reg: "(\!x\Respects R. P x) \ (Bexeq R (\x. P x))" - apply (simp add: Bex1_def Ex1_def Bexeq_def in_respects) +lemma bex1_bexeq_reg: "(\!x\Respects R. P x) \ (Bex1_rel R (\x. P x))" + apply (simp add: Bex1_def Ex1_def Bex1_rel_def in_respects) apply clarify apply auto apply (rule bexI)