src/HOL/Relation.thy
author haftmann
Tue, 28 Apr 2009 13:34:48 +0200
changeset 31011 506e57123cd1
parent 30198 922f944f03b2
child 32231 8f9b8d14fc9f
permissions -rw-r--r--
ephermal enforcement of import order to circumvent current problem in merging interpretation morphisms
wenzelm@10358
     1
(*  Title:      HOL/Relation.thy
paulson@1983
     2
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
paulson@1983
     3
    Copyright   1996  University of Cambridge
nipkow@1128
     4
*)
nipkow@1128
     5
berghofe@12905
     6
header {* Relations *}
berghofe@12905
     7
nipkow@15131
     8
theory Relation
haftmann@31011
     9
imports Finite_Set Datatype
haftmann@31011
    10
  (*FIXME order is important, otherwise merge problem for canonical interpretation of class monoid_mult wrt. power!*)
nipkow@15131
    11
begin
paulson@5978
    12
wenzelm@12913
    13
subsection {* Definitions *}
wenzelm@12913
    14
wenzelm@19656
    15
definition
wenzelm@21404
    16
  converse :: "('a * 'b) set => ('b * 'a) set"
wenzelm@21404
    17
    ("(_^-1)" [1000] 999) where
wenzelm@10358
    18
  "r^-1 == {(y, x). (x, y) : r}"
paulson@5978
    19
wenzelm@21210
    20
notation (xsymbols)
wenzelm@19656
    21
  converse  ("(_\<inverse>)" [1000] 999)
wenzelm@19656
    22
wenzelm@19656
    23
definition
wenzelm@21404
    24
  rel_comp  :: "[('b * 'c) set, ('a * 'b) set] => ('a * 'c) set"
wenzelm@21404
    25
    (infixr "O" 75) where
wenzelm@12913
    26
  "r O s == {(x,z). EX y. (x, y) : s & (y, z) : r}"
wenzelm@12913
    27
wenzelm@21404
    28
definition
wenzelm@21404
    29
  Image :: "[('a * 'b) set, 'a set] => 'b set"
wenzelm@21404
    30
    (infixl "``" 90) where
wenzelm@12913
    31
  "r `` s == {y. EX x:s. (x,y):r}"
paulson@7912
    32
wenzelm@21404
    33
definition
wenzelm@21404
    34
  Id :: "('a * 'a) set" where -- {* the identity relation *}
wenzelm@12913
    35
  "Id == {p. EX x. p = (x,x)}"
paulson@7912
    36
wenzelm@21404
    37
definition
nipkow@30198
    38
  Id_on  :: "'a set => ('a * 'a) set" where -- {* diagonal: identity over a set *}
nipkow@30198
    39
  "Id_on A == \<Union>x\<in>A. {(x,x)}"
wenzelm@12913
    40
wenzelm@21404
    41
definition
wenzelm@21404
    42
  Domain :: "('a * 'b) set => 'a set" where
wenzelm@12913
    43
  "Domain r == {x. EX y. (x,y):r}"
paulson@5978
    44
wenzelm@21404
    45
definition
wenzelm@21404
    46
  Range  :: "('a * 'b) set => 'b set" where
wenzelm@12913
    47
  "Range r == Domain(r^-1)"
paulson@5978
    48
wenzelm@21404
    49
definition
wenzelm@21404
    50
  Field :: "('a * 'a) set => 'a set" where
paulson@13830
    51
  "Field r == Domain r \<union> Range r"
paulson@10786
    52
wenzelm@21404
    53
definition
nipkow@30198
    54
  refl_on :: "['a set, ('a * 'a) set] => bool" where -- {* reflexivity over a set *}
nipkow@30198
    55
  "refl_on A r == r \<subseteq> A \<times> A & (ALL x: A. (x,x) : r)"
paulson@6806
    56
nipkow@26297
    57
abbreviation
nipkow@30198
    58
  refl :: "('a * 'a) set => bool" where -- {* reflexivity over a type *}
nipkow@30198
    59
  "refl == refl_on UNIV"
nipkow@26297
    60
wenzelm@21404
    61
definition
wenzelm@21404
    62
  sym :: "('a * 'a) set => bool" where -- {* symmetry predicate *}
wenzelm@12913
    63
  "sym r == ALL x y. (x,y): r --> (y,x): r"
paulson@6806
    64
wenzelm@21404
    65
definition
wenzelm@21404
    66
  antisym :: "('a * 'a) set => bool" where -- {* antisymmetry predicate *}
wenzelm@12913
    67
  "antisym r == ALL x y. (x,y):r --> (y,x):r --> x=y"
paulson@6806
    68
wenzelm@21404
    69
definition
wenzelm@21404
    70
  trans :: "('a * 'a) set => bool" where -- {* transitivity predicate *}
wenzelm@12913
    71
  "trans r == (ALL x y z. (x,y):r --> (y,z):r --> (x,z):r)"
paulson@5978
    72
wenzelm@21404
    73
definition
nipkow@29796
    74
irrefl :: "('a * 'a) set => bool" where
nipkow@29796
    75
"irrefl r \<equiv> \<forall>x. (x,x) \<notin> r"
nipkow@29796
    76
nipkow@29796
    77
definition
nipkow@29796
    78
total_on :: "'a set => ('a * 'a) set => bool" where
nipkow@29796
    79
"total_on A r \<equiv> \<forall>x\<in>A.\<forall>y\<in>A. x\<noteq>y \<longrightarrow> (x,y)\<in>r \<or> (y,x)\<in>r"
nipkow@29796
    80
nipkow@29796
    81
abbreviation "total \<equiv> total_on UNIV"
nipkow@29796
    82
nipkow@29796
    83
definition
wenzelm@21404
    84
  single_valued :: "('a * 'b) set => bool" where
wenzelm@12913
    85
  "single_valued r == ALL x y. (x,y):r --> (ALL z. (x,z):r --> y=z)"
berghofe@7014
    86
wenzelm@21404
    87
definition
wenzelm@21404
    88
  inv_image :: "('b * 'b) set => ('a => 'b) => ('a * 'a) set" where
wenzelm@12913
    89
  "inv_image r f == {(x, y). (f x, f y) : r}"
oheimb@11136
    90
berghofe@12905
    91
wenzelm@12913
    92
subsection {* The identity relation *}
berghofe@12905
    93
berghofe@12905
    94
lemma IdI [intro]: "(a, a) : Id"
nipkow@26271
    95
by (simp add: Id_def)
berghofe@12905
    96
berghofe@12905
    97
lemma IdE [elim!]: "p : Id ==> (!!x. p = (x, x) ==> P) ==> P"
nipkow@26271
    98
by (unfold Id_def) (iprover elim: CollectE)
berghofe@12905
    99
berghofe@12905
   100
lemma pair_in_Id_conv [iff]: "((a, b) : Id) = (a = b)"
nipkow@26271
   101
by (unfold Id_def) blast
berghofe@12905
   102
nipkow@30198
   103
lemma refl_Id: "refl Id"
nipkow@30198
   104
by (simp add: refl_on_def)
berghofe@12905
   105
berghofe@12905
   106
lemma antisym_Id: "antisym Id"
berghofe@12905
   107
  -- {* A strange result, since @{text Id} is also symmetric. *}
nipkow@26271
   108
by (simp add: antisym_def)
berghofe@12905
   109
huffman@19228
   110
lemma sym_Id: "sym Id"
nipkow@26271
   111
by (simp add: sym_def)
huffman@19228
   112
berghofe@12905
   113
lemma trans_Id: "trans Id"
nipkow@26271
   114
by (simp add: trans_def)
berghofe@12905
   115
berghofe@12905
   116
wenzelm@12913
   117
subsection {* Diagonal: identity over a set *}
berghofe@12905
   118
nipkow@30198
   119
lemma Id_on_empty [simp]: "Id_on {} = {}"
nipkow@30198
   120
by (simp add: Id_on_def) 
paulson@13812
   121
nipkow@30198
   122
lemma Id_on_eqI: "a = b ==> a : A ==> (a, b) : Id_on A"
nipkow@30198
   123
by (simp add: Id_on_def)
berghofe@12905
   124
nipkow@30198
   125
lemma Id_onI [intro!,noatp]: "a : A ==> (a, a) : Id_on A"
nipkow@30198
   126
by (rule Id_on_eqI) (rule refl)
berghofe@12905
   127
nipkow@30198
   128
lemma Id_onE [elim!]:
nipkow@30198
   129
  "c : Id_on A ==> (!!x. x : A ==> c = (x, x) ==> P) ==> P"
wenzelm@12913
   130
  -- {* The general elimination rule. *}
nipkow@30198
   131
by (unfold Id_on_def) (iprover elim!: UN_E singletonE)
berghofe@12905
   132
nipkow@30198
   133
lemma Id_on_iff: "((x, y) : Id_on A) = (x = y & x : A)"
nipkow@26271
   134
by blast
berghofe@12905
   135
nipkow@30198
   136
lemma Id_on_subset_Times: "Id_on A \<subseteq> A \<times> A"
nipkow@26271
   137
by blast
berghofe@12905
   138
berghofe@12905
   139
berghofe@12905
   140
subsection {* Composition of two relations *}
berghofe@12905
   141
wenzelm@12913
   142
lemma rel_compI [intro]:
berghofe@12905
   143
  "(a, b) : s ==> (b, c) : r ==> (a, c) : r O s"
nipkow@26271
   144
by (unfold rel_comp_def) blast
berghofe@12905
   145
wenzelm@12913
   146
lemma rel_compE [elim!]: "xz : r O s ==>
berghofe@12905
   147
  (!!x y z. xz = (x, z) ==> (x, y) : s ==> (y, z) : r  ==> P) ==> P"
nipkow@26271
   148
by (unfold rel_comp_def) (iprover elim!: CollectE splitE exE conjE)
berghofe@12905
   149
berghofe@12905
   150
lemma rel_compEpair:
berghofe@12905
   151
  "(a, c) : r O s ==> (!!y. (a, y) : s ==> (y, c) : r ==> P) ==> P"
nipkow@26271
   152
by (iprover elim: rel_compE Pair_inject ssubst)
berghofe@12905
   153
berghofe@12905
   154
lemma R_O_Id [simp]: "R O Id = R"
nipkow@26271
   155
by fast
berghofe@12905
   156
berghofe@12905
   157
lemma Id_O_R [simp]: "Id O R = R"
nipkow@26271
   158
by fast
berghofe@12905
   159
krauss@23185
   160
lemma rel_comp_empty1[simp]: "{} O R = {}"
nipkow@26271
   161
by blast
krauss@23185
   162
krauss@23185
   163
lemma rel_comp_empty2[simp]: "R O {} = {}"
nipkow@26271
   164
by blast
krauss@23185
   165
berghofe@12905
   166
lemma O_assoc: "(R O S) O T = R O (S O T)"
nipkow@26271
   167
by blast
berghofe@12905
   168
wenzelm@12913
   169
lemma trans_O_subset: "trans r ==> r O r \<subseteq> r"
nipkow@26271
   170
by (unfold trans_def) blast
berghofe@12905
   171
wenzelm@12913
   172
lemma rel_comp_mono: "r' \<subseteq> r ==> s' \<subseteq> s ==> (r' O s') \<subseteq> (r O s)"
nipkow@26271
   173
by blast
berghofe@12905
   174
berghofe@12905
   175
lemma rel_comp_subset_Sigma:
wenzelm@12913
   176
    "s \<subseteq> A \<times> B ==> r \<subseteq> B \<times> C ==> (r O s) \<subseteq> A \<times> C"
nipkow@26271
   177
by blast
berghofe@12905
   178
krauss@28008
   179
lemma rel_comp_distrib[simp]: "R O (S \<union> T) = (R O S) \<union> (R O T)" 
krauss@28008
   180
by auto
krauss@28008
   181
krauss@28008
   182
lemma rel_comp_distrib2[simp]: "(S \<union> T) O R = (S O R) \<union> (T O R)"
krauss@28008
   183
by auto
krauss@28008
   184
berghofe@12905
   185
wenzelm@12913
   186
subsection {* Reflexivity *}
wenzelm@12913
   187
nipkow@30198
   188
lemma refl_onI: "r \<subseteq> A \<times> A ==> (!!x. x : A ==> (x, x) : r) ==> refl_on A r"
nipkow@30198
   189
by (unfold refl_on_def) (iprover intro!: ballI)
berghofe@12905
   190
nipkow@30198
   191
lemma refl_onD: "refl_on A r ==> a : A ==> (a, a) : r"
nipkow@30198
   192
by (unfold refl_on_def) blast
berghofe@12905
   193
nipkow@30198
   194
lemma refl_onD1: "refl_on A r ==> (x, y) : r ==> x : A"
nipkow@30198
   195
by (unfold refl_on_def) blast
huffman@19228
   196
nipkow@30198
   197
lemma refl_onD2: "refl_on A r ==> (x, y) : r ==> y : A"
nipkow@30198
   198
by (unfold refl_on_def) blast
huffman@19228
   199
nipkow@30198
   200
lemma refl_on_Int: "refl_on A r ==> refl_on B s ==> refl_on (A \<inter> B) (r \<inter> s)"
nipkow@30198
   201
by (unfold refl_on_def) blast
huffman@19228
   202
nipkow@30198
   203
lemma refl_on_Un: "refl_on A r ==> refl_on B s ==> refl_on (A \<union> B) (r \<union> s)"
nipkow@30198
   204
by (unfold refl_on_def) blast
huffman@19228
   205
nipkow@30198
   206
lemma refl_on_INTER:
nipkow@30198
   207
  "ALL x:S. refl_on (A x) (r x) ==> refl_on (INTER S A) (INTER S r)"
nipkow@30198
   208
by (unfold refl_on_def) fast
huffman@19228
   209
nipkow@30198
   210
lemma refl_on_UNION:
nipkow@30198
   211
  "ALL x:S. refl_on (A x) (r x) \<Longrightarrow> refl_on (UNION S A) (UNION S r)"
nipkow@30198
   212
by (unfold refl_on_def) blast
huffman@19228
   213
nipkow@30198
   214
lemma refl_on_empty[simp]: "refl_on {} {}"
nipkow@30198
   215
by(simp add:refl_on_def)
nipkow@26297
   216
nipkow@30198
   217
lemma refl_on_Id_on: "refl_on A (Id_on A)"
nipkow@30198
   218
by (rule refl_onI [OF Id_on_subset_Times Id_onI])
huffman@19228
   219
wenzelm@12913
   220
wenzelm@12913
   221
subsection {* Antisymmetry *}
berghofe@12905
   222
berghofe@12905
   223
lemma antisymI:
berghofe@12905
   224
  "(!!x y. (x, y) : r ==> (y, x) : r ==> x=y) ==> antisym r"
nipkow@26271
   225
by (unfold antisym_def) iprover
berghofe@12905
   226
berghofe@12905
   227
lemma antisymD: "antisym r ==> (a, b) : r ==> (b, a) : r ==> a = b"
nipkow@26271
   228
by (unfold antisym_def) iprover
berghofe@12905
   229
huffman@19228
   230
lemma antisym_subset: "r \<subseteq> s ==> antisym s ==> antisym r"
nipkow@26271
   231
by (unfold antisym_def) blast
wenzelm@12913
   232
huffman@19228
   233
lemma antisym_empty [simp]: "antisym {}"
nipkow@26271
   234
by (unfold antisym_def) blast
huffman@19228
   235
nipkow@30198
   236
lemma antisym_Id_on [simp]: "antisym (Id_on A)"
nipkow@26271
   237
by (unfold antisym_def) blast
huffman@19228
   238
huffman@19228
   239
huffman@19228
   240
subsection {* Symmetry *}
huffman@19228
   241
huffman@19228
   242
lemma symI: "(!!a b. (a, b) : r ==> (b, a) : r) ==> sym r"
nipkow@26271
   243
by (unfold sym_def) iprover
paulson@15177
   244
paulson@15177
   245
lemma symD: "sym r ==> (a, b) : r ==> (b, a) : r"
nipkow@26271
   246
by (unfold sym_def, blast)
berghofe@12905
   247
huffman@19228
   248
lemma sym_Int: "sym r ==> sym s ==> sym (r \<inter> s)"
nipkow@26271
   249
by (fast intro: symI dest: symD)
huffman@19228
   250
huffman@19228
   251
lemma sym_Un: "sym r ==> sym s ==> sym (r \<union> s)"
nipkow@26271
   252
by (fast intro: symI dest: symD)
huffman@19228
   253
huffman@19228
   254
lemma sym_INTER: "ALL x:S. sym (r x) ==> sym (INTER S r)"
nipkow@26271
   255
by (fast intro: symI dest: symD)
huffman@19228
   256
huffman@19228
   257
lemma sym_UNION: "ALL x:S. sym (r x) ==> sym (UNION S r)"
nipkow@26271
   258
by (fast intro: symI dest: symD)
huffman@19228
   259
nipkow@30198
   260
lemma sym_Id_on [simp]: "sym (Id_on A)"
nipkow@26271
   261
by (rule symI) clarify
huffman@19228
   262
huffman@19228
   263
huffman@19228
   264
subsection {* Transitivity *}
huffman@19228
   265
berghofe@12905
   266
lemma transI:
berghofe@12905
   267
  "(!!x y z. (x, y) : r ==> (y, z) : r ==> (x, z) : r) ==> trans r"
nipkow@26271
   268
by (unfold trans_def) iprover
berghofe@12905
   269
berghofe@12905
   270
lemma transD: "trans r ==> (a, b) : r ==> (b, c) : r ==> (a, c) : r"
nipkow@26271
   271
by (unfold trans_def) iprover
berghofe@12905
   272
huffman@19228
   273
lemma trans_Int: "trans r ==> trans s ==> trans (r \<inter> s)"
nipkow@26271
   274
by (fast intro: transI elim: transD)
huffman@19228
   275
huffman@19228
   276
lemma trans_INTER: "ALL x:S. trans (r x) ==> trans (INTER S r)"
nipkow@26271
   277
by (fast intro: transI elim: transD)
huffman@19228
   278
nipkow@30198
   279
lemma trans_Id_on [simp]: "trans (Id_on A)"
nipkow@26271
   280
by (fast intro: transI elim: transD)
huffman@19228
   281
nipkow@29796
   282
lemma trans_diff_Id: " trans r \<Longrightarrow> antisym r \<Longrightarrow> trans (r-Id)"
nipkow@29796
   283
unfolding antisym_def trans_def by blast
nipkow@29796
   284
nipkow@29796
   285
subsection {* Irreflexivity *}
nipkow@29796
   286
nipkow@29796
   287
lemma irrefl_diff_Id[simp]: "irrefl(r-Id)"
nipkow@29796
   288
by(simp add:irrefl_def)
nipkow@29796
   289
nipkow@29796
   290
subsection {* Totality *}
nipkow@29796
   291
nipkow@29796
   292
lemma total_on_empty[simp]: "total_on {} r"
nipkow@29796
   293
by(simp add:total_on_def)
nipkow@29796
   294
nipkow@29796
   295
lemma total_on_diff_Id[simp]: "total_on A (r-Id) = total_on A r"
nipkow@29796
   296
by(simp add: total_on_def)
berghofe@12905
   297
wenzelm@12913
   298
subsection {* Converse *}
wenzelm@12913
   299
wenzelm@12913
   300
lemma converse_iff [iff]: "((a,b): r^-1) = ((b,a) : r)"
nipkow@26271
   301
by (simp add: converse_def)
berghofe@12905
   302
nipkow@13343
   303
lemma converseI[sym]: "(a, b) : r ==> (b, a) : r^-1"
nipkow@26271
   304
by (simp add: converse_def)
berghofe@12905
   305
nipkow@13343
   306
lemma converseD[sym]: "(a,b) : r^-1 ==> (b, a) : r"
nipkow@26271
   307
by (simp add: converse_def)
berghofe@12905
   308
berghofe@12905
   309
lemma converseE [elim!]:
berghofe@12905
   310
  "yx : r^-1 ==> (!!x y. yx = (y, x) ==> (x, y) : r ==> P) ==> P"
wenzelm@12913
   311
    -- {* More general than @{text converseD}, as it ``splits'' the member of the relation. *}
nipkow@26271
   312
by (unfold converse_def) (iprover elim!: CollectE splitE bexE)
berghofe@12905
   313
berghofe@12905
   314
lemma converse_converse [simp]: "(r^-1)^-1 = r"
nipkow@26271
   315
by (unfold converse_def) blast
berghofe@12905
   316
berghofe@12905
   317
lemma converse_rel_comp: "(r O s)^-1 = s^-1 O r^-1"
nipkow@26271
   318
by blast
berghofe@12905
   319
huffman@19228
   320
lemma converse_Int: "(r \<inter> s)^-1 = r^-1 \<inter> s^-1"
nipkow@26271
   321
by blast
huffman@19228
   322
huffman@19228
   323
lemma converse_Un: "(r \<union> s)^-1 = r^-1 \<union> s^-1"
nipkow@26271
   324
by blast
huffman@19228
   325
huffman@19228
   326
lemma converse_INTER: "(INTER S r)^-1 = (INT x:S. (r x)^-1)"
nipkow@26271
   327
by fast
huffman@19228
   328
huffman@19228
   329
lemma converse_UNION: "(UNION S r)^-1 = (UN x:S. (r x)^-1)"
nipkow@26271
   330
by blast
huffman@19228
   331
berghofe@12905
   332
lemma converse_Id [simp]: "Id^-1 = Id"
nipkow@26271
   333
by blast
berghofe@12905
   334
nipkow@30198
   335
lemma converse_Id_on [simp]: "(Id_on A)^-1 = Id_on A"
nipkow@26271
   336
by blast
berghofe@12905
   337
nipkow@30198
   338
lemma refl_on_converse [simp]: "refl_on A (converse r) = refl_on A r"
nipkow@30198
   339
by (unfold refl_on_def) auto
berghofe@12905
   340
huffman@19228
   341
lemma sym_converse [simp]: "sym (converse r) = sym r"
nipkow@26271
   342
by (unfold sym_def) blast
huffman@19228
   343
huffman@19228
   344
lemma antisym_converse [simp]: "antisym (converse r) = antisym r"
nipkow@26271
   345
by (unfold antisym_def) blast
berghofe@12905
   346
huffman@19228
   347
lemma trans_converse [simp]: "trans (converse r) = trans r"
nipkow@26271
   348
by (unfold trans_def) blast
berghofe@12905
   349
huffman@19228
   350
lemma sym_conv_converse_eq: "sym r = (r^-1 = r)"
nipkow@26271
   351
by (unfold sym_def) fast
huffman@19228
   352
huffman@19228
   353
lemma sym_Un_converse: "sym (r \<union> r^-1)"
nipkow@26271
   354
by (unfold sym_def) blast
huffman@19228
   355
huffman@19228
   356
lemma sym_Int_converse: "sym (r \<inter> r^-1)"
nipkow@26271
   357
by (unfold sym_def) blast
huffman@19228
   358
nipkow@29796
   359
lemma total_on_converse[simp]: "total_on A (r^-1) = total_on A r"
nipkow@29796
   360
by (auto simp: total_on_def)
nipkow@29796
   361
wenzelm@12913
   362
berghofe@12905
   363
subsection {* Domain *}
berghofe@12905
   364
paulson@24286
   365
declare Domain_def [noatp]
paulson@24286
   366
berghofe@12905
   367
lemma Domain_iff: "(a : Domain r) = (EX y. (a, y) : r)"
nipkow@26271
   368
by (unfold Domain_def) blast
berghofe@12905
   369
berghofe@12905
   370
lemma DomainI [intro]: "(a, b) : r ==> a : Domain r"
nipkow@26271
   371
by (iprover intro!: iffD2 [OF Domain_iff])
berghofe@12905
   372
berghofe@12905
   373
lemma DomainE [elim!]:
berghofe@12905
   374
  "a : Domain r ==> (!!y. (a, y) : r ==> P) ==> P"
nipkow@26271
   375
by (iprover dest!: iffD1 [OF Domain_iff])
berghofe@12905
   376
berghofe@12905
   377
lemma Domain_empty [simp]: "Domain {} = {}"
nipkow@26271
   378
by blast
berghofe@12905
   379
berghofe@12905
   380
lemma Domain_insert: "Domain (insert (a, b) r) = insert a (Domain r)"
nipkow@26271
   381
by blast
berghofe@12905
   382
berghofe@12905
   383
lemma Domain_Id [simp]: "Domain Id = UNIV"
nipkow@26271
   384
by blast
berghofe@12905
   385
nipkow@30198
   386
lemma Domain_Id_on [simp]: "Domain (Id_on A) = A"
nipkow@26271
   387
by blast
berghofe@12905
   388
paulson@13830
   389
lemma Domain_Un_eq: "Domain(A \<union> B) = Domain(A) \<union> Domain(B)"
nipkow@26271
   390
by blast
berghofe@12905
   391
paulson@13830
   392
lemma Domain_Int_subset: "Domain(A \<inter> B) \<subseteq> Domain(A) \<inter> Domain(B)"
nipkow@26271
   393
by blast
berghofe@12905
   394
wenzelm@12913
   395
lemma Domain_Diff_subset: "Domain(A) - Domain(B) \<subseteq> Domain(A - B)"
nipkow@26271
   396
by blast
berghofe@12905
   397
paulson@13830
   398
lemma Domain_Union: "Domain (Union S) = (\<Union>A\<in>S. Domain A)"
nipkow@26271
   399
by blast
nipkow@26271
   400
nipkow@26271
   401
lemma Domain_converse[simp]: "Domain(r^-1) = Range r"
nipkow@26271
   402
by(auto simp:Range_def)
berghofe@12905
   403
wenzelm@12913
   404
lemma Domain_mono: "r \<subseteq> s ==> Domain r \<subseteq> Domain s"
nipkow@26271
   405
by blast
berghofe@12905
   406
paulson@22172
   407
lemma fst_eq_Domain: "fst ` R = Domain R";
nipkow@26271
   408
by (auto intro!:image_eqI)
paulson@22172
   409
haftmann@29609
   410
lemma Domain_dprod [simp]: "Domain (dprod r s) = uprod (Domain r) (Domain s)"
haftmann@29609
   411
by auto
haftmann@29609
   412
haftmann@29609
   413
lemma Domain_dsum [simp]: "Domain (dsum r s) = usum (Domain r) (Domain s)"
haftmann@29609
   414
by auto
haftmann@29609
   415
berghofe@12905
   416
berghofe@12905
   417
subsection {* Range *}
berghofe@12905
   418
berghofe@12905
   419
lemma Range_iff: "(a : Range r) = (EX y. (y, a) : r)"
nipkow@26271
   420
by (simp add: Domain_def Range_def)
berghofe@12905
   421
berghofe@12905
   422
lemma RangeI [intro]: "(a, b) : r ==> b : Range r"
nipkow@26271
   423
by (unfold Range_def) (iprover intro!: converseI DomainI)
berghofe@12905
   424
berghofe@12905
   425
lemma RangeE [elim!]: "b : Range r ==> (!!x. (x, b) : r ==> P) ==> P"
nipkow@26271
   426
by (unfold Range_def) (iprover elim!: DomainE dest!: converseD)
berghofe@12905
   427
berghofe@12905
   428
lemma Range_empty [simp]: "Range {} = {}"
nipkow@26271
   429
by blast
berghofe@12905
   430
berghofe@12905
   431
lemma Range_insert: "Range (insert (a, b) r) = insert b (Range r)"
nipkow@26271
   432
by blast
berghofe@12905
   433
berghofe@12905
   434
lemma Range_Id [simp]: "Range Id = UNIV"
nipkow@26271
   435
by blast
berghofe@12905
   436
nipkow@30198
   437
lemma Range_Id_on [simp]: "Range (Id_on A) = A"
nipkow@26271
   438
by auto
berghofe@12905
   439
paulson@13830
   440
lemma Range_Un_eq: "Range(A \<union> B) = Range(A) \<union> Range(B)"
nipkow@26271
   441
by blast
berghofe@12905
   442
paulson@13830
   443
lemma Range_Int_subset: "Range(A \<inter> B) \<subseteq> Range(A) \<inter> Range(B)"
nipkow@26271
   444
by blast
berghofe@12905
   445
wenzelm@12913
   446
lemma Range_Diff_subset: "Range(A) - Range(B) \<subseteq> Range(A - B)"
nipkow@26271
   447
by blast
berghofe@12905
   448
paulson@13830
   449
lemma Range_Union: "Range (Union S) = (\<Union>A\<in>S. Range A)"
nipkow@26271
   450
by blast
nipkow@26271
   451
nipkow@26271
   452
lemma Range_converse[simp]: "Range(r^-1) = Domain r"
nipkow@26271
   453
by blast
berghofe@12905
   454
paulson@22172
   455
lemma snd_eq_Range: "snd ` R = Range R";
nipkow@26271
   456
by (auto intro!:image_eqI)
nipkow@26271
   457
nipkow@26271
   458
nipkow@26271
   459
subsection {* Field *}
nipkow@26271
   460
nipkow@26271
   461
lemma mono_Field: "r \<subseteq> s \<Longrightarrow> Field r \<subseteq> Field s"
nipkow@26271
   462
by(auto simp:Field_def Domain_def Range_def)
nipkow@26271
   463
nipkow@26271
   464
lemma Field_empty[simp]: "Field {} = {}"
nipkow@26271
   465
by(auto simp:Field_def)
nipkow@26271
   466
nipkow@26271
   467
lemma Field_insert[simp]: "Field (insert (a,b) r) = {a,b} \<union> Field r"
nipkow@26271
   468
by(auto simp:Field_def)
nipkow@26271
   469
nipkow@26271
   470
lemma Field_Un[simp]: "Field (r \<union> s) = Field r \<union> Field s"
nipkow@26271
   471
by(auto simp:Field_def)
nipkow@26271
   472
nipkow@26271
   473
lemma Field_Union[simp]: "Field (\<Union>R) = \<Union>(Field ` R)"
nipkow@26271
   474
by(auto simp:Field_def)
nipkow@26271
   475
nipkow@26271
   476
lemma Field_converse[simp]: "Field(r^-1) = Field r"
nipkow@26271
   477
by(auto simp:Field_def)
paulson@22172
   478
berghofe@12905
   479
berghofe@12905
   480
subsection {* Image of a set under a relation *}
berghofe@12905
   481
paulson@24286
   482
declare Image_def [noatp]
paulson@24286
   483
wenzelm@12913
   484
lemma Image_iff: "(b : r``A) = (EX x:A. (x, b) : r)"
nipkow@26271
   485
by (simp add: Image_def)
berghofe@12905
   486
wenzelm@12913
   487
lemma Image_singleton: "r``{a} = {b. (a, b) : r}"
nipkow@26271
   488
by (simp add: Image_def)
berghofe@12905
   489
wenzelm@12913
   490
lemma Image_singleton_iff [iff]: "(b : r``{a}) = ((a, b) : r)"
nipkow@26271
   491
by (rule Image_iff [THEN trans]) simp
berghofe@12905
   492
paulson@24286
   493
lemma ImageI [intro,noatp]: "(a, b) : r ==> a : A ==> b : r``A"
nipkow@26271
   494
by (unfold Image_def) blast
berghofe@12905
   495
berghofe@12905
   496
lemma ImageE [elim!]:
wenzelm@12913
   497
    "b : r `` A ==> (!!x. (x, b) : r ==> x : A ==> P) ==> P"
nipkow@26271
   498
by (unfold Image_def) (iprover elim!: CollectE bexE)
berghofe@12905
   499
berghofe@12905
   500
lemma rev_ImageI: "a : A ==> (a, b) : r ==> b : r `` A"
berghofe@12905
   501
  -- {* This version's more effective when we already have the required @{text a} *}
nipkow@26271
   502
by blast
berghofe@12905
   503
berghofe@12905
   504
lemma Image_empty [simp]: "R``{} = {}"
nipkow@26271
   505
by blast
berghofe@12905
   506
berghofe@12905
   507
lemma Image_Id [simp]: "Id `` A = A"
nipkow@26271
   508
by blast
berghofe@12905
   509
nipkow@30198
   510
lemma Image_Id_on [simp]: "Id_on A `` B = A \<inter> B"
nipkow@26271
   511
by blast
berghofe@12905
   512
paulson@13830
   513
lemma Image_Int_subset: "R `` (A \<inter> B) \<subseteq> R `` A \<inter> R `` B"
nipkow@26271
   514
by blast
berghofe@12905
   515
paulson@13830
   516
lemma Image_Int_eq:
paulson@13830
   517
     "single_valued (converse R) ==> R `` (A \<inter> B) = R `` A \<inter> R `` B"
nipkow@26271
   518
by (simp add: single_valued_def, blast) 
paulson@13830
   519
paulson@13830
   520
lemma Image_Un: "R `` (A \<union> B) = R `` A \<union> R `` B"
nipkow@26271
   521
by blast
berghofe@12905
   522
paulson@13812
   523
lemma Un_Image: "(R \<union> S) `` A = R `` A \<union> S `` A"
nipkow@26271
   524
by blast
paulson@13812
   525
wenzelm@12913
   526
lemma Image_subset: "r \<subseteq> A \<times> B ==> r``C \<subseteq> B"
nipkow@26271
   527
by (iprover intro!: subsetI elim!: ImageE dest!: subsetD SigmaD2)
berghofe@12905
   528
paulson@13830
   529
lemma Image_eq_UN: "r``B = (\<Union>y\<in> B. r``{y})"
berghofe@12905
   530
  -- {* NOT suitable for rewriting *}
nipkow@26271
   531
by blast
berghofe@12905
   532
wenzelm@12913
   533
lemma Image_mono: "r' \<subseteq> r ==> A' \<subseteq> A ==> (r' `` A') \<subseteq> (r `` A)"
nipkow@26271
   534
by blast
berghofe@12905
   535
paulson@13830
   536
lemma Image_UN: "(r `` (UNION A B)) = (\<Union>x\<in>A. r `` (B x))"
nipkow@26271
   537
by blast
berghofe@12905
   538
paulson@13830
   539
lemma Image_INT_subset: "(r `` INTER A B) \<subseteq> (\<Inter>x\<in>A. r `` (B x))"
nipkow@26271
   540
by blast
berghofe@12905
   541
paulson@13830
   542
text{*Converse inclusion requires some assumptions*}
paulson@13830
   543
lemma Image_INT_eq:
paulson@13830
   544
     "[|single_valued (r\<inverse>); A\<noteq>{}|] ==> r `` INTER A B = (\<Inter>x\<in>A. r `` B x)"
paulson@13830
   545
apply (rule equalityI)
paulson@13830
   546
 apply (rule Image_INT_subset) 
paulson@13830
   547
apply  (simp add: single_valued_def, blast)
paulson@13830
   548
done
paulson@13830
   549
wenzelm@12913
   550
lemma Image_subset_eq: "(r``A \<subseteq> B) = (A \<subseteq> - ((r^-1) `` (-B)))"
nipkow@26271
   551
by blast
berghofe@12905
   552
berghofe@12905
   553
wenzelm@12913
   554
subsection {* Single valued relations *}
wenzelm@12913
   555
wenzelm@12913
   556
lemma single_valuedI:
berghofe@12905
   557
  "ALL x y. (x,y):r --> (ALL z. (x,z):r --> y=z) ==> single_valued r"
nipkow@26271
   558
by (unfold single_valued_def)
berghofe@12905
   559
berghofe@12905
   560
lemma single_valuedD:
berghofe@12905
   561
  "single_valued r ==> (x, y) : r ==> (x, z) : r ==> y = z"
nipkow@26271
   562
by (simp add: single_valued_def)
berghofe@12905
   563
huffman@19228
   564
lemma single_valued_rel_comp:
huffman@19228
   565
  "single_valued r ==> single_valued s ==> single_valued (r O s)"
nipkow@26271
   566
by (unfold single_valued_def) blast
huffman@19228
   567
huffman@19228
   568
lemma single_valued_subset:
huffman@19228
   569
  "r \<subseteq> s ==> single_valued s ==> single_valued r"
nipkow@26271
   570
by (unfold single_valued_def) blast
huffman@19228
   571
huffman@19228
   572
lemma single_valued_Id [simp]: "single_valued Id"
nipkow@26271
   573
by (unfold single_valued_def) blast
huffman@19228
   574
nipkow@30198
   575
lemma single_valued_Id_on [simp]: "single_valued (Id_on A)"
nipkow@26271
   576
by (unfold single_valued_def) blast
huffman@19228
   577
berghofe@12905
   578
berghofe@12905
   579
subsection {* Graphs given by @{text Collect} *}
berghofe@12905
   580
berghofe@12905
   581
lemma Domain_Collect_split [simp]: "Domain{(x,y). P x y} = {x. EX y. P x y}"
nipkow@26271
   582
by auto
berghofe@12905
   583
berghofe@12905
   584
lemma Range_Collect_split [simp]: "Range{(x,y). P x y} = {y. EX x. P x y}"
nipkow@26271
   585
by auto
berghofe@12905
   586
berghofe@12905
   587
lemma Image_Collect_split [simp]: "{(x,y). P x y} `` A = {y. EX x:A. P x y}"
nipkow@26271
   588
by auto
berghofe@12905
   589
berghofe@12905
   590
wenzelm@12913
   591
subsection {* Inverse image *}
berghofe@12905
   592
huffman@19228
   593
lemma sym_inv_image: "sym r ==> sym (inv_image r f)"
nipkow@26271
   594
by (unfold sym_def inv_image_def) blast
huffman@19228
   595
wenzelm@12913
   596
lemma trans_inv_image: "trans r ==> trans (inv_image r f)"
berghofe@12905
   597
  apply (unfold trans_def inv_image_def)
berghofe@12905
   598
  apply (simp (no_asm))
berghofe@12905
   599
  apply blast
berghofe@12905
   600
  done
berghofe@12905
   601
haftmann@23709
   602
haftmann@29609
   603
subsection {* Finiteness *}
haftmann@29609
   604
haftmann@29609
   605
lemma finite_converse [iff]: "finite (r^-1) = finite r"
haftmann@29609
   606
  apply (subgoal_tac "r^-1 = (%(x,y). (y,x))`r")
haftmann@29609
   607
   apply simp
haftmann@29609
   608
   apply (rule iffI)
haftmann@29609
   609
    apply (erule finite_imageD [unfolded inj_on_def])
haftmann@29609
   610
    apply (simp split add: split_split)
haftmann@29609
   611
   apply (erule finite_imageI)
haftmann@29609
   612
  apply (simp add: converse_def image_def, auto)
haftmann@29609
   613
  apply (rule bexI)
haftmann@29609
   614
   prefer 2 apply assumption
haftmann@29609
   615
  apply simp
haftmann@29609
   616
  done
haftmann@29609
   617
haftmann@29609
   618
text {* \paragraph{Finiteness of transitive closure} (Thanks to Sidi
haftmann@29609
   619
Ehmety) *}
haftmann@29609
   620
haftmann@29609
   621
lemma finite_Field: "finite r ==> finite (Field r)"
haftmann@29609
   622
  -- {* A finite relation has a finite field (@{text "= domain \<union> range"}. *}
haftmann@29609
   623
  apply (induct set: finite)
haftmann@29609
   624
   apply (auto simp add: Field_def Domain_insert Range_insert)
haftmann@29609
   625
  done
haftmann@29609
   626
haftmann@29609
   627
haftmann@23709
   628
subsection {* Version of @{text lfp_induct} for binary relations *}
haftmann@23709
   629
haftmann@23709
   630
lemmas lfp_induct2 = 
haftmann@23709
   631
  lfp_induct_set [of "(a, b)", split_format (complete)]
haftmann@23709
   632
nipkow@1128
   633
end