src/HOL/NumberTheory/BijectionRel.thy
changeset 21404 eb85850d3eb7
parent 19670 2e4a143c73c5
child 22274 ce1459004c8d
equal deleted inserted replaced
21403:dd58f13a8eb4 21404:eb85850d3eb7
    28   Add extra condition to @{term insert}: @{term "\<forall>b \<in> B. \<not> P a b"}
    28   Add extra condition to @{term insert}: @{term "\<forall>b \<in> B. \<not> P a b"}
    29   (and similar for @{term A}).
    29   (and similar for @{term A}).
    30 *}
    30 *}
    31 
    31 
    32 definition
    32 definition
    33   bijP :: "('a => 'a => bool) => 'a set => bool"
    33   bijP :: "('a => 'a => bool) => 'a set => bool" where
    34   "bijP P F = (\<forall>a b. a \<in> F \<and> P a b --> b \<in> F)"
    34   "bijP P F = (\<forall>a b. a \<in> F \<and> P a b --> b \<in> F)"
    35 
    35 
    36   uniqP :: "('a => 'a => bool) => bool"
    36 definition
       
    37   uniqP :: "('a => 'a => bool) => bool" where
    37   "uniqP P = (\<forall>a b c d. P a b \<and> P c d --> (a = c) = (b = d))"
    38   "uniqP P = (\<forall>a b c d. P a b \<and> P c d --> (a = c) = (b = d))"
    38 
    39 
    39   symP :: "('a => 'a => bool) => bool"
    40 definition
       
    41   symP :: "('a => 'a => bool) => bool" where
    40   "symP P = (\<forall>a b. P a b = P b a)"
    42   "symP P = (\<forall>a b. P a b = P b a)"
    41 
    43 
    42 consts
    44 consts
    43   bijER :: "('a => 'a => bool) => 'a set set"
    45   bijER :: "('a => 'a => bool) => 'a set set"
    44 
    46