src/HOL/Library/Disjoint_Sets.thy
author wenzelm
Tue, 19 Jul 2016 09:55:03 +0200
changeset 63520 2803d2b8f85d
parent 63148 6a767355d1a9
child 63928 d81fb5b46a5c
permissions -rw-r--r--
Linux platform base-line is Ubuntu 12.04 LTS;
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
63098
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
     5
section \<open>Partitions and Disjoint Sets\<close>
60727
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
63098
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
    23
lemma disjnt_equiv_class: "equiv A r \<Longrightarrow> disjnt (r``{a}) (r``{b}) \<longleftrightarrow> (a, b) \<notin> r"
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
    24
  by (auto dest: equiv_class_self simp: equiv_class_eq_iff disjnt_def)
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
    25
60727
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    26
subsection \<open>Set of Disjoint Sets\<close>
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    27
62843
313d3b697c9a Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
paulson <lp15@cam.ac.uk>
parents: 62390
diff changeset
    28
abbreviation disjoint :: "'a set set \<Rightarrow> bool" where "disjoint \<equiv> pairwise disjnt"
313d3b697c9a Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
paulson <lp15@cam.ac.uk>
parents: 62390
diff changeset
    29
313d3b697c9a Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
paulson <lp15@cam.ac.uk>
parents: 62390
diff changeset
    30
lemma disjoint_def: "disjoint A \<longleftrightarrow> (\<forall>a\<in>A. \<forall>b\<in>A. a \<noteq> b \<longrightarrow> a \<inter> b = {})"
313d3b697c9a Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
paulson <lp15@cam.ac.uk>
parents: 62390
diff changeset
    31
  unfolding pairwise_def disjnt_def by auto
60727
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    32
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    33
lemma disjointI:
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    34
  "(\<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
    35
  unfolding disjoint_def by auto
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    36
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    37
lemma disjointD:
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    38
  "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
    39
  unfolding disjoint_def by auto
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    40
63099
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
    41
lemma disjoint_image: "inj_on f (\<Union>A) \<Longrightarrow> disjoint A \<Longrightarrow> disjoint (op ` f ` A)"
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
    42
  unfolding inj_on_def disjoint_def by blast
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
    43
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
    44
lemma assumes "disjoint (A \<union> B)"
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
    45
      shows   disjoint_unionD1: "disjoint A" and disjoint_unionD2: "disjoint B"
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
    46
  using assms by (simp_all add: disjoint_def)
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
    47
  
60727
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    48
lemma disjoint_INT:
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    49
  assumes *: "\<And>i. i \<in> I \<Longrightarrow> disjoint (F i)"
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    50
  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
    51
proof (safe intro!: disjointI del: equalityI)
63098
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
    52
  fix A B :: "'a \<Rightarrow> 'b set" assume "(\<Inter>i\<in>I. A i) \<noteq> (\<Inter>i\<in>I. B i)"
60727
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    53
  then obtain i where "A i \<noteq> B i" "i \<in> I"
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    54
    by auto
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    55
  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
    56
  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
    57
    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
    58
    by (auto simp: INT_Int_distrib[symmetric])
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    59
qed
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
subsubsection "Family of Disjoint Sets"
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    62
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    63
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
    64
  "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
    65
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    66
abbreviation "disjoint_family A \<equiv> disjoint_family_on A UNIV"
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_onD:
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    69
  "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
    70
  by (auto simp: disjoint_family_on_def)
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    71
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    72
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
    73
  by (force simp add: disjoint_family_on_def)
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    74
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    75
lemma disjoint_family_on_bisimulation:
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    76
  assumes "disjoint_family_on f S"
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    77
  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
    78
  shows "disjoint_family_on g S"
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    79
  using assms unfolding disjoint_family_on_def by auto
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    80
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    81
lemma disjoint_family_on_mono:
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    82
  "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
    83
  unfolding disjoint_family_on_def by auto
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    84
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    85
lemma disjoint_family_Suc:
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    86
  "(\<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
    87
  using lift_Suc_mono_le[of A]
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    88
  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
    89
     (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
    90
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    91
lemma disjoint_family_on_disjoint_image:
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    92
  "disjoint_family_on A I \<Longrightarrow> disjoint (A ` I)"
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    93
  unfolding disjoint_family_on_def disjoint_def by force
63099
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
    94
 
60727
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    95
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
    96
  by (auto simp: disjoint_family_on_def)
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    97
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    98
lemma disjoint_image_disjoint_family_on:
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    99
  assumes d: "disjoint (A ` I)" and i: "inj_on A I"
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   100
  shows "disjoint_family_on A I"
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   101
  unfolding disjoint_family_on_def
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   102
proof (intro ballI impI)
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   103
  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
   104
  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
   105
    by (intro disjointD[OF d]) auto
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   106
qed
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   107
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   108
lemma disjoint_UN:
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   109
  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
   110
  shows "disjoint (\<Union>i\<in>I. F i)"
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   111
proof (safe intro!: disjointI del: equalityI)
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   112
  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
   113
  show "A \<inter> B = {}"
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   114
  proof cases
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   115
    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
   116
      by (auto dest: disjointD)
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   117
  next
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   118
    assume "i \<noteq> j"
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   119
    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
   120
      by (rule disjoint_family_onD)
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   121
    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
   122
    show "A \<inter> B = {}"
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   123
      by auto
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   124
  qed
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   125
qed
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   126
63099
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
   127
lemma distinct_list_bind: 
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
   128
  assumes "distinct xs" "\<And>x. x \<in> set xs \<Longrightarrow> distinct (f x)" 
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
   129
          "disjoint_family_on (set \<circ> f) (set xs)"
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
   130
  shows   "distinct (List.bind xs f)"
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
   131
  using assms
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
   132
  by (induction xs)
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
   133
     (auto simp: disjoint_family_on_def distinct_map inj_on_def set_list_bind)
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
   134
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
   135
lemma bij_betw_UNION_disjoint:
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
   136
  assumes disj: "disjoint_family_on A' I"
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
   137
  assumes bij: "\<And>i. i \<in> I \<Longrightarrow> bij_betw f (A i) (A' i)"
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
   138
  shows   "bij_betw f (\<Union>i\<in>I. A i) (\<Union>i\<in>I. A' i)"
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
   139
unfolding bij_betw_def
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
   140
proof
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
   141
  from bij show eq: "f ` UNION I A = UNION I A'"
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
   142
    by (auto simp: bij_betw_def image_UN)
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
   143
  show "inj_on f (UNION I A)"
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
   144
  proof (rule inj_onI, clarify)
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
   145
    fix i j x y assume A: "i \<in> I" "j \<in> I" "x \<in> A i" "y \<in> A j" and B: "f x = f y"
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
   146
    from A bij[of i] bij[of j] have "f x \<in> A' i" "f y \<in> A' j"
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
   147
      by (auto simp: bij_betw_def)
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
   148
    with B have "A' i \<inter> A' j \<noteq> {}" by auto
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
   149
    with disj A have "i = j" unfolding disjoint_family_on_def by blast
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
   150
    with A B bij[of i] show "x = y" by (auto simp: bij_betw_def dest: inj_onD)
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
   151
  qed
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
   152
qed
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
   153
60727
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   154
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
   155
  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
   156
63099
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
   157
text \<open>
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
   158
  The union of an infinite disjoint family of non-empty sets is infinite.
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
   159
\<close>
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
   160
lemma infinite_disjoint_family_imp_infinite_UNION:
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
   161
  assumes "\<not>finite A" "\<And>x. x \<in> A \<Longrightarrow> f x \<noteq> {}" "disjoint_family_on f A"
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
   162
  shows   "\<not>finite (UNION A f)"
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
   163
proof -
63148
6a767355d1a9 updated 'define';
wenzelm
parents: 63145
diff changeset
   164
  define g where "g x = (SOME y. y \<in> f x)" for x
63099
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
   165
  have g: "g x \<in> f x" if "x \<in> A" for x
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
   166
    unfolding g_def by (rule someI_ex, insert assms(2) that) blast
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
   167
  have inj_on_g: "inj_on g A"
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
   168
  proof (rule inj_onI, rule ccontr)
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
   169
    fix x y assume A: "x \<in> A" "y \<in> A" "g x = g y" "x \<noteq> y"
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
   170
    with g[of x] g[of y] have "g x \<in> f x" "g x \<in> f y" by auto
63145
703edebd1d92 isabelle update_cartouches -c -t;
wenzelm
parents: 63100
diff changeset
   171
    with A \<open>x \<noteq> y\<close> assms show False
63099
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
   172
      by (auto simp: disjoint_family_on_def inj_on_def)
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
   173
  qed
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
   174
  from g have "g ` A \<subseteq> UNION A f" by blast
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
   175
  moreover from inj_on_g \<open>\<not>finite A\<close> have "\<not>finite (g ` A)"
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
   176
    using finite_imageD by blast
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
   177
  ultimately show ?thesis using finite_subset by blast
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
   178
qed  
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
   179
  
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
   180
60727
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   181
subsection \<open>Construct Disjoint Sequences\<close>
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   182
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   183
definition disjointed :: "(nat \<Rightarrow> 'a set) \<Rightarrow> nat \<Rightarrow> 'a set" where
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   184
  "disjointed A n = A n - (\<Union>i\<in>{0..<n}. A i)"
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   185
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   186
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
   187
proof (induct n)
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   188
  case 0 show ?case by simp
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   189
next
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   190
  case (Suc n)
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   191
  thus ?case by (simp add: atLeastLessThanSuc disjointed_def)
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   192
qed
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   193
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   194
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
   195
  by (rule UN_finite2_eq [where k=0])
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   196
     (simp add: finite_UN_disjointed_eq)
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   197
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   198
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
   199
  by (auto simp add: disjointed_def)
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   200
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   201
lemma disjoint_family_disjointed: "disjoint_family (disjointed A)"
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   202
  by (simp add: disjoint_family_on_def)
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   203
     (metis neq_iff Int_commute less_disjoint_disjointed)
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   204
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   205
lemma disjointed_subset: "disjointed A n \<subseteq> A n"
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   206
  by (auto simp add: disjointed_def)
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   207
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   208
lemma disjointed_0[simp]: "disjointed A 0 = A 0"
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   209
  by (simp add: disjointed_def)
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   210
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   211
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
   212
  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
   213
63098
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   214
subsection \<open>Partitions\<close>
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   215
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   216
text \<open>
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   217
  Partitions @{term P} of a set @{term A}. We explicitly disallow empty sets.
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   218
\<close>
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   219
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   220
definition partition_on :: "'a set \<Rightarrow> 'a set set \<Rightarrow> bool"
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   221
where
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   222
  "partition_on A P \<longleftrightarrow> \<Union>P = A \<and> disjoint P \<and> {} \<notin> P"
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   223
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   224
lemma partition_onI:
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   225
  "\<Union>P = A \<Longrightarrow> (\<And>p q. p \<in> P \<Longrightarrow> q \<in> P \<Longrightarrow> p \<noteq> q \<Longrightarrow> disjnt p q) \<Longrightarrow> {} \<notin> P \<Longrightarrow> partition_on A P"
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   226
  by (auto simp: partition_on_def pairwise_def)
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   227
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   228
lemma partition_onD1: "partition_on A P \<Longrightarrow> A = \<Union>P"
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   229
  by (auto simp: partition_on_def)
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   230
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   231
lemma partition_onD2: "partition_on A P \<Longrightarrow> disjoint P"
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   232
  by (auto simp: partition_on_def)
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   233
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   234
lemma partition_onD3: "partition_on A P \<Longrightarrow> {} \<notin> P"
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   235
  by (auto simp: partition_on_def)
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   236
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   237
subsection \<open>Constructions of partitions\<close>
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   238
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   239
lemma partition_on_empty: "partition_on {} P \<longleftrightarrow> P = {}"
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   240
  unfolding partition_on_def by fastforce
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   241
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   242
lemma partition_on_space: "A \<noteq> {} \<Longrightarrow> partition_on A {A}"
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   243
  by (auto simp: partition_on_def disjoint_def)
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   244
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   245
lemma partition_on_singletons: "partition_on A ((\<lambda>x. {x}) ` A)"
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   246
  by (auto simp: partition_on_def disjoint_def)
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   247
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   248
lemma partition_on_transform:
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   249
  assumes P: "partition_on A P"
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   250
  assumes F_UN: "\<Union>(F ` P) = F (\<Union>P)" and F_disjnt: "\<And>p q. p \<in> P \<Longrightarrow> q \<in> P \<Longrightarrow> disjnt p q \<Longrightarrow> disjnt (F p) (F q)"
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   251
  shows "partition_on (F A) (F ` P - {{}})"
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   252
proof -
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   253
  have "\<Union>(F ` P - {{}}) = F A"
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   254
    unfolding P[THEN partition_onD1] F_UN[symmetric] by auto
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   255
  with P show ?thesis
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   256
    by (auto simp add: partition_on_def pairwise_def intro!: F_disjnt)
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   257
qed
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   258
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   259
lemma partition_on_restrict: "partition_on A P \<Longrightarrow> partition_on (B \<inter> A) (op \<inter> B ` P - {{}})"
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   260
  by (intro partition_on_transform) (auto simp: disjnt_def)
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   261
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   262
lemma partition_on_vimage: "partition_on A P \<Longrightarrow> partition_on (f -` A) (op -` f ` P - {{}})"
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   263
  by (intro partition_on_transform) (auto simp: disjnt_def)
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   264
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   265
lemma partition_on_inj_image:
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   266
  assumes P: "partition_on A P" and f: "inj_on f A"
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   267
  shows "partition_on (f ` A) (op ` f ` P - {{}})"
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   268
proof (rule partition_on_transform[OF P])
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   269
  show "p \<in> P \<Longrightarrow> q \<in> P \<Longrightarrow> disjnt p q \<Longrightarrow> disjnt (f ` p) (f ` q)" for p q
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   270
    using f[THEN inj_onD] P[THEN partition_onD1] by (auto simp: disjnt_def)
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   271
qed auto
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   272
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   273
subsection \<open>Finiteness of partitions\<close>
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   274
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   275
lemma finitely_many_partition_on:
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   276
  assumes "finite A"
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   277
  shows "finite {P. partition_on A P}"
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   278
proof (rule finite_subset)
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   279
  show "{P. partition_on A P} \<subseteq> Pow (Pow A)"
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   280
    unfolding partition_on_def by auto
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   281
  show "finite (Pow (Pow A))"
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   282
    using assms by simp
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   283
qed
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   284
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   285
lemma finite_elements: "finite A \<Longrightarrow> partition_on A P \<Longrightarrow> finite P"
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   286
  using partition_onD1[of A P] by (simp add: finite_UnionD)
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   287
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   288
subsection \<open>Equivalence of partitions and equivalence classes\<close>
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   289
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   290
lemma partition_on_quotient:
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   291
  assumes r: "equiv A r"
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   292
  shows "partition_on A (A // r)"
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   293
proof (rule partition_onI)
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   294
  from r have "refl_on A r"
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   295
    by (auto elim: equivE)
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   296
  then show "\<Union>(A // r) = A" "{} \<notin> A // r"
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   297
    by (auto simp: refl_on_def quotient_def)
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   298
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   299
  fix p q assume "p \<in> A // r" "q \<in> A // r" "p \<noteq> q"
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   300
  then obtain x y where "x \<in> A" "y \<in> A" "p = r `` {x}" "q = r `` {y}"
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   301
    by (auto simp: quotient_def)
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   302
  with r equiv_class_eq_iff[OF r, of x y] \<open>p \<noteq> q\<close> show "disjnt p q"
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   303
    by (auto simp: disjnt_equiv_class)
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   304
qed
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   305
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   306
lemma equiv_partition_on:
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   307
  assumes P: "partition_on A P"
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   308
  shows "equiv A {(x, y). \<exists>p \<in> P. x \<in> p \<and> y \<in> p}"
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   309
proof (rule equivI)
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   310
  have "A = \<Union>P" "disjoint P" "{} \<notin> P"
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   311
    using P by (auto simp: partition_on_def)
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   312
  then show "refl_on A {(x, y). \<exists>p\<in>P. x \<in> p \<and> y \<in> p}"
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   313
    unfolding refl_on_def by auto
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   314
  show "trans {(x, y). \<exists>p\<in>P. x \<in> p \<and> y \<in> p}"
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   315
    using \<open>disjoint P\<close> by (auto simp: trans_def disjoint_def)
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   316
qed (auto simp: sym_def)
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   317
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   318
lemma partition_on_eq_quotient:
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   319
  assumes P: "partition_on A P"
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   320
  shows "A // {(x, y). \<exists>p \<in> P. x \<in> p \<and> y \<in> p} = P"
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   321
  unfolding quotient_def
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   322
proof safe
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   323
  fix x assume "x \<in> A"
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   324
  then obtain p where "p \<in> P" "x \<in> p" "\<And>q. q \<in> P \<Longrightarrow> x \<in> q \<Longrightarrow> p = q"
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   325
    using P by (auto simp: partition_on_def disjoint_def)
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   326
  then have "{y. \<exists>p\<in>P. x \<in> p \<and> y \<in> p} = p"
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   327
    by (safe intro!: bexI[of _ p]) simp
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   328
  then show "{(x, y). \<exists>p\<in>P. x \<in> p \<and> y \<in> p} `` {x} \<in> P"
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   329
    by (simp add: \<open>p \<in> P\<close>)
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   330
next
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   331
  fix p assume "p \<in> P"
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   332
  then have "p \<noteq> {}"
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   333
    using P by (auto simp: partition_on_def)
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   334
  then obtain x where "x \<in> p"
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   335
    by auto
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   336
  then have "x \<in> A" "\<And>q. q \<in> P \<Longrightarrow> x \<in> q \<Longrightarrow> p = q"
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   337
    using P \<open>p \<in> P\<close> by (auto simp: partition_on_def disjoint_def)
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   338
  with \<open>p\<in>P\<close> \<open>x \<in> p\<close> have "{y. \<exists>p\<in>P. x \<in> p \<and> y \<in> p} = p"
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   339
    by (safe intro!: bexI[of _ p]) simp
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   340
  then show "p \<in> (\<Union>x\<in>A. {{(x, y). \<exists>p\<in>P. x \<in> p \<and> y \<in> p} `` {x}})"
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   341
    by (auto intro: \<open>x \<in> A\<close>)
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   342
qed
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   343
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   344
lemma partition_on_alt: "partition_on A P \<longleftrightarrow> (\<exists>r. equiv A r \<and> P = A // r)"
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   345
  by (auto simp: partition_on_eq_quotient intro!: partition_on_quotient intro: equiv_partition_on)
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   346
62390
842917225d56 more canonical names
nipkow
parents: 61824
diff changeset
   347
end