src/HOL/Library/FuncSet.thy
author paulson
Fri, 14 May 2004 16:49:42 +0200
changeset 14745 94be403deb84
parent 14706 71590b7733b7
child 14762 bd349ff7907a
permissions -rw-r--r--
new lemmas
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
13586
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
     1
(*  Title:      HOL/Library/FuncSet.thy
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
     2
    ID:         $Id$
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
     3
    Author:     Florian Kammueller and Lawrence C Paulson
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
     4
*)
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
     5
14706
71590b7733b7 tuned document;
wenzelm
parents: 14565
diff changeset
     6
header {* Pi and Function Sets *}
13586
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
     7
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
     8
theory FuncSet = Main:
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
     9
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
    10
constdefs
14706
71590b7733b7 tuned document;
wenzelm
parents: 14565
diff changeset
    11
  Pi :: "['a set, 'a => 'b set] => ('a => 'b) set"
71590b7733b7 tuned document;
wenzelm
parents: 14565
diff changeset
    12
  "Pi A B == {f. \<forall>x. x \<in> A --> f x \<in> B x}"
13586
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
    13
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
    14
  extensional :: "'a set => ('a => 'b) set"
14706
71590b7733b7 tuned document;
wenzelm
parents: 14565
diff changeset
    15
  "extensional A == {f. \<forall>x. x~:A --> f x = arbitrary}"
13586
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
    16
14706
71590b7733b7 tuned document;
wenzelm
parents: 14565
diff changeset
    17
  "restrict" :: "['a => 'b, 'a set] => ('a => 'b)"
71590b7733b7 tuned document;
wenzelm
parents: 14565
diff changeset
    18
  "restrict f A == (%x. if x \<in> A then f x else arbitrary)"
13586
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
    19
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
    20
syntax
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
    21
  "@Pi"  :: "[pttrn, 'a set, 'b set] => ('a => 'b) set"  ("(3PI _:_./ _)" 10)
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
    22
  funcset :: "['a set, 'b set] => ('a => 'b) set"      (infixr "->" 60)
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
    23
  "@lam" :: "[pttrn, 'a set, 'a => 'b] => ('a=>'b)"  ("(3%_:_./ _)" [0,0,3] 3)
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
    24
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
    25
syntax (xsymbols)
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
    26
  "@Pi" :: "[pttrn, 'a set, 'b set] => ('a => 'b) set"  ("(3\<Pi> _\<in>_./ _)"   10)
14706
71590b7733b7 tuned document;
wenzelm
parents: 14565
diff changeset
    27
  funcset :: "['a set, 'b set] => ('a => 'b) set"  (infixr "\<rightarrow>" 60)
13586
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
    28
  "@lam" :: "[pttrn, 'a set, 'a => 'b] => ('a=>'b)"  ("(3\<lambda>_\<in>_./ _)" [0,0,3] 3)
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
    29
14565
c6dc17aab88a use more symbols in HTML output
kleing
parents: 13595
diff changeset
    30
syntax (HTML output)
c6dc17aab88a use more symbols in HTML output
kleing
parents: 13595
diff changeset
    31
  "@Pi" :: "[pttrn, 'a set, 'b set] => ('a => 'b) set"  ("(3\<Pi> _\<in>_./ _)"   10)
c6dc17aab88a use more symbols in HTML output
kleing
parents: 13595
diff changeset
    32
  "@lam" :: "[pttrn, 'a set, 'a => 'b] => ('a=>'b)"  ("(3\<lambda>_\<in>_./ _)" [0,0,3] 3)
c6dc17aab88a use more symbols in HTML output
kleing
parents: 13595
diff changeset
    33
13586
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
    34
translations
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
    35
  "PI x:A. B" => "Pi A (%x. B)"
14706
71590b7733b7 tuned document;
wenzelm
parents: 14565
diff changeset
    36
  "A -> B" => "Pi A (_K B)"
71590b7733b7 tuned document;
wenzelm
parents: 14565
diff changeset
    37
  "%x:A. f" == "restrict (%x. f) A"
13586
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
    38
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
    39
constdefs
14706
71590b7733b7 tuned document;
wenzelm
parents: 14565
diff changeset
    40
  "compose" :: "['a set, 'b => 'c, 'a => 'b] => ('a => 'c)"
13586
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
    41
  "compose A g f == \<lambda>x\<in>A. g (f x)"
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
    42
13595
7e6cdcd113a2 Proof tidying
paulson
parents: 13593
diff changeset
    43
print_translation {* [("Pi", dependent_tr' ("@Pi", "funcset"))] *}
13586
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
    44
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
    45
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
    46
subsection{*Basic Properties of @{term Pi}*}
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
    47
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
    48
lemma Pi_I: "(!!x. x \<in> A ==> f x \<in> B x) ==> f \<in> Pi A B"
14706
71590b7733b7 tuned document;
wenzelm
parents: 14565
diff changeset
    49
  by (simp add: Pi_def)
13586
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
    50
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
    51
lemma funcsetI: "(!!x. x \<in> A ==> f x \<in> B) ==> f \<in> A -> B"
14706
71590b7733b7 tuned document;
wenzelm
parents: 14565
diff changeset
    52
  by (simp add: Pi_def)
13586
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
    53
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
    54
lemma Pi_mem: "[|f: Pi A B; x \<in> A|] ==> f x \<in> B x"
14706
71590b7733b7 tuned document;
wenzelm
parents: 14565
diff changeset
    55
  by (simp add: Pi_def)
13586
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
    56
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
    57
lemma funcset_mem: "[|f \<in> A -> B; x \<in> A|] ==> f x \<in> B"
14706
71590b7733b7 tuned document;
wenzelm
parents: 14565
diff changeset
    58
  by (simp add: Pi_def)
13586
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
    59
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
    60
lemma Pi_eq_empty: "((PI x: A. B x) = {}) = (\<exists>x\<in>A. B(x) = {})"
13593
e39f0751e4bf Tidied. New Pi-theorem.
paulson
parents: 13586
diff changeset
    61
apply (simp add: Pi_def, auto)
13586
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
    62
txt{*Converse direction requires Axiom of Choice to exhibit a function
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
    63
picking an element from each non-empty @{term "B x"}*}
13593
e39f0751e4bf Tidied. New Pi-theorem.
paulson
parents: 13586
diff changeset
    64
apply (drule_tac x = "%u. SOME y. y \<in> B u" in spec, auto)
14706
71590b7733b7 tuned document;
wenzelm
parents: 14565
diff changeset
    65
apply (cut_tac P= "%y. y \<in> B x" in some_eq_ex, auto)
13586
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
    66
done
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
    67
13593
e39f0751e4bf Tidied. New Pi-theorem.
paulson
parents: 13586
diff changeset
    68
lemma Pi_empty [simp]: "Pi {} B = UNIV"
14706
71590b7733b7 tuned document;
wenzelm
parents: 14565
diff changeset
    69
  by (simp add: Pi_def)
13593
e39f0751e4bf Tidied. New Pi-theorem.
paulson
parents: 13586
diff changeset
    70
e39f0751e4bf Tidied. New Pi-theorem.
paulson
parents: 13586
diff changeset
    71
lemma Pi_UNIV [simp]: "A -> UNIV = UNIV"
14706
71590b7733b7 tuned document;
wenzelm
parents: 14565
diff changeset
    72
  by (simp add: Pi_def)
13586
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
    73
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
    74
text{*Covariance of Pi-sets in their second argument*}
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
    75
lemma Pi_mono: "(!!x. x \<in> A ==> B x <= C x) ==> Pi A B <= Pi A C"
14706
71590b7733b7 tuned document;
wenzelm
parents: 14565
diff changeset
    76
  by (simp add: Pi_def, blast)
13586
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
    77
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
    78
text{*Contravariance of Pi-sets in their first argument*}
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
    79
lemma Pi_anti_mono: "A' <= A ==> Pi A B <= Pi A' B"
14706
71590b7733b7 tuned document;
wenzelm
parents: 14565
diff changeset
    80
  by (simp add: Pi_def, blast)
13586
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
    81
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
    82
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
    83
subsection{*Composition With a Restricted Domain: @{term compose}*}
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
    84
14706
71590b7733b7 tuned document;
wenzelm
parents: 14565
diff changeset
    85
lemma funcset_compose:
71590b7733b7 tuned document;
wenzelm
parents: 14565
diff changeset
    86
    "[| f \<in> A -> B; g \<in> B -> C |]==> compose A g f \<in> A -> C"
71590b7733b7 tuned document;
wenzelm
parents: 14565
diff changeset
    87
  by (simp add: Pi_def compose_def restrict_def)
13586
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
    88
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
    89
lemma compose_assoc:
14706
71590b7733b7 tuned document;
wenzelm
parents: 14565
diff changeset
    90
    "[| f \<in> A -> B; g \<in> B -> C; h \<in> C -> D |]
13586
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
    91
      ==> compose A h (compose A g f) = compose A (compose B h g) f"
14706
71590b7733b7 tuned document;
wenzelm
parents: 14565
diff changeset
    92
  by (simp add: expand_fun_eq Pi_def compose_def restrict_def)
13586
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
    93
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
    94
lemma compose_eq: "x \<in> A ==> compose A g f x = g(f(x))"
14706
71590b7733b7 tuned document;
wenzelm
parents: 14565
diff changeset
    95
  by (simp add: compose_def restrict_def)
13586
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
    96
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
    97
lemma surj_compose: "[| f ` A = B; g ` B = C |] ==> compose A g f ` A = C"
14706
71590b7733b7 tuned document;
wenzelm
parents: 14565
diff changeset
    98
  by (auto simp add: image_def compose_eq)
13586
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
    99
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
   100
lemma inj_on_compose:
14706
71590b7733b7 tuned document;
wenzelm
parents: 14565
diff changeset
   101
    "[| f ` A = B; inj_on f A; inj_on g B |] ==> inj_on (compose A g f) A"
71590b7733b7 tuned document;
wenzelm
parents: 14565
diff changeset
   102
  by (auto simp add: inj_on_def compose_eq)
13586
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
   103
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
   104
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
   105
subsection{*Bounded Abstraction: @{term restrict}*}
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
   106
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
   107
lemma restrict_in_funcset: "(!!x. x \<in> A ==> f x \<in> B) ==> (\<lambda>x\<in>A. f x) \<in> A -> B"
14706
71590b7733b7 tuned document;
wenzelm
parents: 14565
diff changeset
   108
  by (simp add: Pi_def restrict_def)
13586
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
   109
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
   110
lemma restrictI: "(!!x. x \<in> A ==> f x \<in> B x) ==> (\<lambda>x\<in>A. f x) \<in> Pi A B"
14706
71590b7733b7 tuned document;
wenzelm
parents: 14565
diff changeset
   111
  by (simp add: Pi_def restrict_def)
13586
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
   112
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
   113
lemma restrict_apply [simp]:
14706
71590b7733b7 tuned document;
wenzelm
parents: 14565
diff changeset
   114
    "(\<lambda>y\<in>A. f y) x = (if x \<in> A then f x else arbitrary)"
71590b7733b7 tuned document;
wenzelm
parents: 14565
diff changeset
   115
  by (simp add: restrict_def)
13586
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
   116
14706
71590b7733b7 tuned document;
wenzelm
parents: 14565
diff changeset
   117
lemma restrict_ext:
13586
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
   118
    "(!!x. x \<in> A ==> f x = g x) ==> (\<lambda>x\<in>A. f x) = (\<lambda>x\<in>A. g x)"
14706
71590b7733b7 tuned document;
wenzelm
parents: 14565
diff changeset
   119
  by (simp add: expand_fun_eq Pi_def Pi_def restrict_def)
13586
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
   120
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
   121
lemma inj_on_restrict_eq: "inj_on (restrict f A) A = inj_on f A"
14706
71590b7733b7 tuned document;
wenzelm
parents: 14565
diff changeset
   122
  by (simp add: inj_on_def restrict_def)
13586
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
   123
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
   124
lemma Id_compose:
14706
71590b7733b7 tuned document;
wenzelm
parents: 14565
diff changeset
   125
    "[|f \<in> A -> B;  f \<in> extensional A|] ==> compose A (\<lambda>y\<in>B. y) f = f"
71590b7733b7 tuned document;
wenzelm
parents: 14565
diff changeset
   126
  by (auto simp add: expand_fun_eq compose_def extensional_def Pi_def)
13586
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
   127
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
   128
lemma compose_Id:
14706
71590b7733b7 tuned document;
wenzelm
parents: 14565
diff changeset
   129
    "[|g \<in> A -> B;  g \<in> extensional A|] ==> compose A g (\<lambda>x\<in>A. x) = g"
71590b7733b7 tuned document;
wenzelm
parents: 14565
diff changeset
   130
  by (auto simp add: expand_fun_eq compose_def extensional_def Pi_def)
13586
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
   131
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
   132
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
   133
subsection{*Extensionality*}
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
   134
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
   135
lemma extensional_arb: "[|f \<in> extensional A; x\<notin> A|] ==> f x = arbitrary"
14706
71590b7733b7 tuned document;
wenzelm
parents: 14565
diff changeset
   136
  by (simp add: extensional_def)
13586
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
   137
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
   138
lemma restrict_extensional [simp]: "restrict f A \<in> extensional A"
14706
71590b7733b7 tuned document;
wenzelm
parents: 14565
diff changeset
   139
  by (simp add: restrict_def extensional_def)
13586
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
   140
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
   141
lemma compose_extensional [simp]: "compose A f g \<in> extensional A"
14706
71590b7733b7 tuned document;
wenzelm
parents: 14565
diff changeset
   142
  by (simp add: compose_def)
13586
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
   143
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
   144
lemma extensionalityI:
14706
71590b7733b7 tuned document;
wenzelm
parents: 14565
diff changeset
   145
    "[| f \<in> extensional A; g \<in> extensional A;
71590b7733b7 tuned document;
wenzelm
parents: 14565
diff changeset
   146
      !!x. x\<in>A ==> f x = g x |] ==> f = g"
71590b7733b7 tuned document;
wenzelm
parents: 14565
diff changeset
   147
  by (force simp add: expand_fun_eq extensional_def)
13586
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
   148
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
   149
lemma Inv_funcset: "f ` A = B ==> (\<lambda>x\<in>B. Inv A f x) : B -> A"
14706
71590b7733b7 tuned document;
wenzelm
parents: 14565
diff changeset
   150
  by (unfold Inv_def) (fast intro: restrict_in_funcset someI2)
13586
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
   151
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
   152
lemma compose_Inv_id:
14706
71590b7733b7 tuned document;
wenzelm
parents: 14565
diff changeset
   153
    "[| inj_on f A;  f ` A = B |]
13586
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
   154
      ==> compose A (\<lambda>y\<in>B. Inv A f y) f = (\<lambda>x\<in>A. x)"
14706
71590b7733b7 tuned document;
wenzelm
parents: 14565
diff changeset
   155
  apply (simp add: compose_def)
71590b7733b7 tuned document;
wenzelm
parents: 14565
diff changeset
   156
  apply (rule restrict_ext, auto)
71590b7733b7 tuned document;
wenzelm
parents: 14565
diff changeset
   157
  apply (erule subst)
71590b7733b7 tuned document;
wenzelm
parents: 14565
diff changeset
   158
  apply (simp add: Inv_f_f)
71590b7733b7 tuned document;
wenzelm
parents: 14565
diff changeset
   159
  done
13586
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
   160
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
   161
lemma compose_id_Inv:
14706
71590b7733b7 tuned document;
wenzelm
parents: 14565
diff changeset
   162
    "f ` A = B ==> compose B f (\<lambda>y\<in>B. Inv A f y) = (\<lambda>x\<in>B. x)"
71590b7733b7 tuned document;
wenzelm
parents: 14565
diff changeset
   163
  apply (simp add: compose_def)
71590b7733b7 tuned document;
wenzelm
parents: 14565
diff changeset
   164
  apply (rule restrict_ext)
71590b7733b7 tuned document;
wenzelm
parents: 14565
diff changeset
   165
  apply (simp add: f_Inv_f)
71590b7733b7 tuned document;
wenzelm
parents: 14565
diff changeset
   166
  done
13586
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
   167
14745
94be403deb84 new lemmas
paulson
parents: 14706
diff changeset
   168
94be403deb84 new lemmas
paulson
parents: 14706
diff changeset
   169
subsection{*Cardinality*}
94be403deb84 new lemmas
paulson
parents: 14706
diff changeset
   170
94be403deb84 new lemmas
paulson
parents: 14706
diff changeset
   171
lemma card_inj: "[|f \<in> A\<rightarrow>B; inj_on f A; finite B|] ==> card(A) \<le> card(B)"
94be403deb84 new lemmas
paulson
parents: 14706
diff changeset
   172
apply (rule card_inj_on_le)
94be403deb84 new lemmas
paulson
parents: 14706
diff changeset
   173
apply (auto simp add: Pi_def)
94be403deb84 new lemmas
paulson
parents: 14706
diff changeset
   174
done
94be403deb84 new lemmas
paulson
parents: 14706
diff changeset
   175
94be403deb84 new lemmas
paulson
parents: 14706
diff changeset
   176
lemma card_bij:
94be403deb84 new lemmas
paulson
parents: 14706
diff changeset
   177
     "[|f \<in> A\<rightarrow>B; inj_on f A;
94be403deb84 new lemmas
paulson
parents: 14706
diff changeset
   178
        g \<in> B\<rightarrow>A; inj_on g B; finite A; finite B|] ==> card(A) = card(B)"
94be403deb84 new lemmas
paulson
parents: 14706
diff changeset
   179
by (blast intro: card_inj order_antisym)
94be403deb84 new lemmas
paulson
parents: 14706
diff changeset
   180
13586
0f339348df0e new theory for Pi-sets, restrict, etc.
paulson
parents:
diff changeset
   181
end