diff -r dd58f13a8eb4 -r eb85850d3eb7 src/HOL/NumberTheory/BijectionRel.thy --- 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 = (\a b. a \ F \ P a b --> b \ F)" - uniqP :: "('a => 'a => bool) => bool" +definition + uniqP :: "('a => 'a => bool) => bool" where "uniqP P = (\a b c d. P a b \ P c d --> (a = c) = (b = d))" - symP :: "('a => 'a => bool) => bool" +definition + symP :: "('a => 'a => bool) => bool" where "symP P = (\a b. P a b = P b a)" consts