diff -r b5d293e0b9bb -r 9cb8f9a59402 QuotScript.thy --- a/QuotScript.thy Sat Oct 17 15:42:57 2009 +0200 +++ b/QuotScript.thy Sat Oct 17 16:06:54 2009 +0200 @@ -81,8 +81,8 @@ by auto lemma IDENTITY_QUOTIENT: - shows "QUOTIENT (op =) (\x. x) (\x. x)" -unfolding QUOTIENT_def + shows "QUOTIENT (op =) id id" +unfolding QUOTIENT_def id_def by blast lemma QUOTIENT_SYM: @@ -114,8 +114,8 @@ "f ---> g \ fun_map f g" lemma FUN_MAP_I: - shows "((\x. x) ---> (\x. x)) = (\x. x)" -by (simp add: expand_fun_eq) + shows "(id ---> id) = id" +by (simp add: expand_fun_eq id_def) lemma IN_FUN: shows "x \ ((f ---> g) s) = g (f x \ s)" @@ -382,13 +382,13 @@ lemma I_PRS: assumes q: "QUOTIENT R Abs Rep" - shows "(\x. x) e = Abs ((\ x. x) (Rep e))" + shows "id e = Abs (id (Rep e))" using QUOTIENT_ABS_REP[OF q] by auto lemma I_RSP: assumes q: "QUOTIENT R Abs Rep" and a: "R e1 e2" - shows "R ((\x. x) e1) ((\ x. x) e2)" + shows "R (id e1) (id e2)" using a by auto lemma o_PRS: