src/HOL/Library/Disjoint_Sets.thy
author paulson <lp15@cam.ac.uk>
Thu, 10 Dec 2015 13:38:40 +0000
changeset 61824 dcbe9f756ae0
parent 60727 53697011b03a
child 62390 842917225d56
permissions -rw-r--r--
not_leE -> not_le_imp_less and other tidying
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
60727
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
     1
(*  Title:      HOL/Library/Disjoint_Sets.thy
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
     2
    Author:     Johannes Hölzl, TU München
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
     3
*)
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
     4
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
     5
section \<open>Handling Disjoint Sets\<close>
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
     6
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
     7
theory Disjoint_Sets
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
     8
  imports Main
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
     9
begin
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    10
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    11
lemma range_subsetD: "range f \<subseteq> B \<Longrightarrow> f i \<in> B"
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    12
  by blast
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    13
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    14
lemma Int_Diff_disjoint: "A \<inter> B \<inter> (A - B) = {}"
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    15
  by blast
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    16
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    17
lemma Int_Diff_Un: "A \<inter> B \<union> (A - B) = A"
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    18
  by blast
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    19
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    20
lemma mono_Un: "mono A \<Longrightarrow> (\<Union>i\<le>n. A i) = A n"
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    21
  unfolding mono_def by auto
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    22
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    23
subsection \<open>Set of Disjoint Sets\<close>
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    24
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    25
definition "disjoint A \<longleftrightarrow> (\<forall>a\<in>A. \<forall>b\<in>A. a \<noteq> b \<longrightarrow> a \<inter> b = {})"
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    26
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    27
lemma disjointI:
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    28
  "(\<And>a b. a \<in> A \<Longrightarrow> b \<in> A \<Longrightarrow> a \<noteq> b \<Longrightarrow> a \<inter> b = {}) \<Longrightarrow> disjoint A"
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    29
  unfolding disjoint_def by auto
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    30
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    31
lemma disjointD:
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    32
  "disjoint A \<Longrightarrow> a \<in> A \<Longrightarrow> b \<in> A \<Longrightarrow> a \<noteq> b \<Longrightarrow> a \<inter> b = {}"
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    33
  unfolding disjoint_def by auto
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    34
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    35
lemma disjoint_empty[iff]: "disjoint {}"
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    36
  by (auto simp: disjoint_def)
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    37
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    38
lemma disjoint_INT:
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    39
  assumes *: "\<And>i. i \<in> I \<Longrightarrow> disjoint (F i)"
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    40
  shows "disjoint {\<Inter>i\<in>I. X i | X. \<forall>i\<in>I. X i \<in> F i}"
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    41
proof (safe intro!: disjointI del: equalityI)
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    42
  fix A B :: "'a \<Rightarrow> 'b set" assume "(\<Inter>i\<in>I. A i) \<noteq> (\<Inter>i\<in>I. B i)" 
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    43
  then obtain i where "A i \<noteq> B i" "i \<in> I"
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    44
    by auto
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    45
  moreover assume "\<forall>i\<in>I. A i \<in> F i" "\<forall>i\<in>I. B i \<in> F i"
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    46
  ultimately show "(\<Inter>i\<in>I. A i) \<inter> (\<Inter>i\<in>I. B i) = {}"
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    47
    using *[OF \<open>i\<in>I\<close>, THEN disjointD, of "A i" "B i"]
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    48
    by (auto simp: INT_Int_distrib[symmetric])
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    49
qed
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    50
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    51
lemma disjoint_singleton[simp]: "disjoint {A}"
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    52
  by(simp add: disjoint_def)
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    53
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    54
subsubsection "Family of Disjoint Sets"
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    55
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    56
definition disjoint_family_on :: "('i \<Rightarrow> 'a set) \<Rightarrow> 'i set \<Rightarrow> bool" where
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    57
  "disjoint_family_on A S \<longleftrightarrow> (\<forall>m\<in>S. \<forall>n\<in>S. m \<noteq> n \<longrightarrow> A m \<inter> A n = {})"
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    58
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    59
abbreviation "disjoint_family A \<equiv> disjoint_family_on A UNIV"
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    60
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    61
lemma disjoint_family_onD:
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    62
  "disjoint_family_on A I \<Longrightarrow> i \<in> I \<Longrightarrow> j \<in> I \<Longrightarrow> i \<noteq> j \<Longrightarrow> A i \<inter> A j = {}"
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    63
  by (auto simp: disjoint_family_on_def)
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    64
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    65
lemma disjoint_family_subset: "disjoint_family A \<Longrightarrow> (\<And>x. B x \<subseteq> A x) \<Longrightarrow> disjoint_family B"
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    66
  by (force simp add: disjoint_family_on_def)
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    67
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    68
lemma disjoint_family_on_bisimulation:
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    69
  assumes "disjoint_family_on f S"
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    70
  and "\<And>n m. n \<in> S \<Longrightarrow> m \<in> S \<Longrightarrow> n \<noteq> m \<Longrightarrow> f n \<inter> f m = {} \<Longrightarrow> g n \<inter> g m = {}"
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    71
  shows "disjoint_family_on g S"
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    72
  using assms unfolding disjoint_family_on_def by auto
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    73
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    74
lemma disjoint_family_on_mono:
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    75
  "A \<subseteq> B \<Longrightarrow> disjoint_family_on f B \<Longrightarrow> disjoint_family_on f A"
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    76
  unfolding disjoint_family_on_def by auto
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    77
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    78
lemma disjoint_family_Suc:
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    79
  "(\<And>n. A n \<subseteq> A (Suc n)) \<Longrightarrow> disjoint_family (\<lambda>i. A (Suc i) - A i)"
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    80
  using lift_Suc_mono_le[of A]
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    81
  by (auto simp add: disjoint_family_on_def)
61824
dcbe9f756ae0 not_leE -> not_le_imp_less and other tidying
paulson <lp15@cam.ac.uk>
parents: 60727
diff changeset
    82
     (metis insert_absorb insert_subset le_SucE le_antisym not_le_imp_less less_imp_le)
60727
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    83
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    84
lemma disjoint_family_on_disjoint_image:
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    85
  "disjoint_family_on A I \<Longrightarrow> disjoint (A ` I)"
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    86
  unfolding disjoint_family_on_def disjoint_def by force
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    87
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    88
lemma disjoint_family_on_vimageI: "disjoint_family_on F I \<Longrightarrow> disjoint_family_on (\<lambda>i. f -` F i) I"
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    89
  by (auto simp: disjoint_family_on_def)
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    90
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    91
lemma disjoint_image_disjoint_family_on:
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    92
  assumes d: "disjoint (A ` I)" and i: "inj_on A I"
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    93
  shows "disjoint_family_on A I"
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    94
  unfolding disjoint_family_on_def
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    95
proof (intro ballI impI)
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    96
  fix n m assume nm: "m \<in> I" "n \<in> I" and "n \<noteq> m"
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    97
  with i[THEN inj_onD, of n m] show "A n \<inter> A m = {}"
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    98
    by (intro disjointD[OF d]) auto
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    99
qed
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   100
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   101
lemma disjoint_UN:
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   102
  assumes F: "\<And>i. i \<in> I \<Longrightarrow> disjoint (F i)" and *: "disjoint_family_on (\<lambda>i. \<Union>F i) I"
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   103
  shows "disjoint (\<Union>i\<in>I. F i)"
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   104
proof (safe intro!: disjointI del: equalityI)
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   105
  fix A B i j assume "A \<noteq> B" "A \<in> F i" "i \<in> I" "B \<in> F j" "j \<in> I"
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   106
  show "A \<inter> B = {}"
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   107
  proof cases
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   108
    assume "i = j" with F[of i] \<open>i \<in> I\<close> \<open>A \<in> F i\<close> \<open>B \<in> F j\<close> \<open>A \<noteq> B\<close> show "A \<inter> B = {}"
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   109
      by (auto dest: disjointD)
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   110
  next
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   111
    assume "i \<noteq> j"
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   112
    with * \<open>i\<in>I\<close> \<open>j\<in>I\<close> have "(\<Union>F i) \<inter> (\<Union>F j) = {}"
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   113
      by (rule disjoint_family_onD)
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   114
    with \<open>A\<in>F i\<close> \<open>i\<in>I\<close> \<open>B\<in>F j\<close> \<open>j\<in>I\<close>
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   115
    show "A \<inter> B = {}"
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   116
      by auto
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   117
  qed
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   118
qed
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   119
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   120
lemma disjoint_union: "disjoint C \<Longrightarrow> disjoint B \<Longrightarrow> \<Union>C \<inter> \<Union>B = {} \<Longrightarrow> disjoint (C \<union> B)"
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   121
  using disjoint_UN[of "{C, B}" "\<lambda>x. x"] by (auto simp add: disjoint_family_on_def)
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   122
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   123
subsection \<open>Construct Disjoint Sequences\<close>
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   124
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   125
definition disjointed :: "(nat \<Rightarrow> 'a set) \<Rightarrow> nat \<Rightarrow> 'a set" where
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   126
  "disjointed A n = A n - (\<Union>i\<in>{0..<n}. A i)"
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   127
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   128
lemma finite_UN_disjointed_eq: "(\<Union>i\<in>{0..<n}. disjointed A i) = (\<Union>i\<in>{0..<n}. A i)"
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   129
proof (induct n)
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   130
  case 0 show ?case by simp
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   131
next
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   132
  case (Suc n)
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   133
  thus ?case by (simp add: atLeastLessThanSuc disjointed_def)
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   134
qed
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   135
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   136
lemma UN_disjointed_eq: "(\<Union>i. disjointed A i) = (\<Union>i. A i)"
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   137
  by (rule UN_finite2_eq [where k=0])
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   138
     (simp add: finite_UN_disjointed_eq)
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   139
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   140
lemma less_disjoint_disjointed: "m < n \<Longrightarrow> disjointed A m \<inter> disjointed A n = {}"
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   141
  by (auto simp add: disjointed_def)
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   142
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   143
lemma disjoint_family_disjointed: "disjoint_family (disjointed A)"
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   144
  by (simp add: disjoint_family_on_def)
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   145
     (metis neq_iff Int_commute less_disjoint_disjointed)
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   146
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   147
lemma disjointed_subset: "disjointed A n \<subseteq> A n"
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   148
  by (auto simp add: disjointed_def)
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   149
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   150
lemma disjointed_0[simp]: "disjointed A 0 = A 0"
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   151
  by (simp add: disjointed_def)
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   152
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   153
lemma disjointed_mono: "mono A \<Longrightarrow> disjointed A (Suc n) = A (Suc n) - A n"
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   154
  using mono_Un[of A] by (simp add: disjointed_def atLeastLessThanSuc_atLeastAtMost atLeast0AtMost)
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   155
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   156
end