src/HOL/NumberTheory/BijectionRel.thy
changeset 21404 eb85850d3eb7
parent 19670 2e4a143c73c5
child 22274 ce1459004c8d
--- a/src/HOL/NumberTheory/BijectionRel.thy	Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/NumberTheory/BijectionRel.thy	Fri Nov 17 02:20:03 2006 +0100
@@ -30,13 +30,15 @@
 *}
 
 definition
-  bijP :: "('a => 'a => bool) => 'a set => bool"
+  bijP :: "('a => 'a => bool) => 'a set => bool" where
   "bijP P F = (\<forall>a b. a \<in> F \<and> P a b --> b \<in> F)"
 
-  uniqP :: "('a => 'a => bool) => bool"
+definition
+  uniqP :: "('a => 'a => bool) => bool" where
   "uniqP P = (\<forall>a b c d. P a b \<and> P c d --> (a = c) = (b = d))"
 
-  symP :: "('a => 'a => bool) => bool"
+definition
+  symP :: "('a => 'a => bool) => bool" where
   "symP P = (\<forall>a b. P a b = P b a)"
 
 consts