generalized type
authorhaftmann
Fri, 19 Nov 2010 11:44:46 +0100
changeset 40615 ab551d108feb
parent 40614 d6eeffa0d9a0
child 40616 c5ee1e06d795
child 40617 4a1173d21ec4
generalized type
src/HOL/Quotient.thy
--- a/src/HOL/Quotient.thy	Fri Nov 19 10:37:06 2010 +0100
+++ b/src/HOL/Quotient.thy	Fri Nov 19 11:44:46 2010 +0100
@@ -167,7 +167,7 @@
   by (simp add: fun_eq_iff)
 
 definition
-  fun_rel :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> ('b \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> bool" (infixr "===>" 55)
+  fun_rel :: "('a \<Rightarrow> 'c \<Rightarrow> bool) \<Rightarrow> ('b \<Rightarrow> 'd \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> ('c \<Rightarrow> 'd) \<Rightarrow> bool" (infixr "===>" 55)
 where
   "fun_rel E1 E2 = (\<lambda>f g. \<forall>x y. E1 x y \<longrightarrow> E2 (f x) (g y))"