src/HOL/Library/Disjoint_Sets.thy
author desharna
Wed, 27 Mar 2024 18:29:32 +0100
changeset 80067 c40bdfc84640
parent 74598 5d91897a8e54
permissions -rw-r--r--
tuned proofs of Equiv_Relations.equiv
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
73477
1d8a79aa2a99 dedicated session for combinatorial material
haftmann
parents: 69745
diff changeset
     1
(*  Author:     Johannes Hölzl, TU München
60727
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
     2
*)
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
     3
63098
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
     4
section \<open>Partitions and Disjoint Sets\<close>
60727
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
     5
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
     6
theory Disjoint_Sets
74590
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
     7
  imports FuncSet
60727
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
     8
begin
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
     9
74590
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
    10
lemma mono_imp_UN_eq_last: "mono A \<Longrightarrow> (\<Union>i\<le>n. A i) = A n"
60727
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    11
  unfolding mono_def by auto
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    12
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    13
subsection \<open>Set of Disjoint Sets\<close>
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    14
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
    15
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
    16
313d3b697c9a Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
paulson <lp15@cam.ac.uk>
parents: 62390
diff changeset
    17
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
    18
  unfolding pairwise_def disjnt_def by auto
60727
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 disjointI:
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    21
  "(\<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
    22
  unfolding disjoint_def by auto
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    23
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    24
lemma disjointD:
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    25
  "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
    26
  unfolding disjoint_def by auto
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    27
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 63928
diff changeset
    28
lemma disjoint_image: "inj_on f (\<Union>A) \<Longrightarrow> disjoint A \<Longrightarrow> disjoint ((`) f ` A)"
63099
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
    29
  unfolding inj_on_def disjoint_def by blast
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
    30
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
    31
lemma assumes "disjoint (A \<union> B)"
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
    32
      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
    33
  using assms by (simp_all add: disjoint_def)
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
    34
  
60727
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    35
lemma disjoint_INT:
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    36
  assumes *: "\<And>i. i \<in> I \<Longrightarrow> disjoint (F i)"
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    37
  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
    38
proof (safe intro!: disjointI del: equalityI)
63098
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
    39
  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
    40
  then obtain i where "A i \<noteq> B i" "i \<in> I"
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    41
    by auto
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    42
  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
    43
  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
    44
    using *[OF \<open>i\<in>I\<close>, THEN disjointD, of "A i" "B i"]
68406
6beb45f6cf67 utilize 'flip'
nipkow
parents: 67399
diff changeset
    45
    by (auto simp flip: INT_Int_distrib)
60727
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    46
qed
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    47
69712
dc85b5b3a532 renamings and new material
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    48
lemma diff_Union_pairwise_disjoint:
dc85b5b3a532 renamings and new material
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    49
  assumes "pairwise disjnt \<A>" "\<B> \<subseteq> \<A>"
dc85b5b3a532 renamings and new material
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    50
  shows "\<Union>\<A> - \<Union>\<B> = \<Union>(\<A> - \<B>)"
dc85b5b3a532 renamings and new material
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    51
proof -
dc85b5b3a532 renamings and new material
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    52
  have "False"
dc85b5b3a532 renamings and new material
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    53
    if x: "x \<in> A" "x \<in> B" and AB: "A \<in> \<A>" "A \<notin> \<B>" "B \<in> \<B>" for x A B
dc85b5b3a532 renamings and new material
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    54
  proof -
dc85b5b3a532 renamings and new material
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    55
    have "A \<inter> B = {}"
dc85b5b3a532 renamings and new material
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    56
      using assms disjointD AB by blast
dc85b5b3a532 renamings and new material
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    57
  with x show ?thesis
dc85b5b3a532 renamings and new material
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    58
    by blast
dc85b5b3a532 renamings and new material
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    59
  qed
dc85b5b3a532 renamings and new material
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    60
  then show ?thesis by auto
dc85b5b3a532 renamings and new material
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    61
qed
dc85b5b3a532 renamings and new material
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    62
dc85b5b3a532 renamings and new material
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    63
lemma Int_Union_pairwise_disjoint:
dc85b5b3a532 renamings and new material
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    64
  assumes "pairwise disjnt (\<A> \<union> \<B>)"
dc85b5b3a532 renamings and new material
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    65
  shows "\<Union>\<A> \<inter> \<Union>\<B> = \<Union>(\<A> \<inter> \<B>)"
dc85b5b3a532 renamings and new material
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    66
proof -
dc85b5b3a532 renamings and new material
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    67
  have "False"
dc85b5b3a532 renamings and new material
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    68
    if x: "x \<in> A" "x \<in> B" and AB: "A \<in> \<A>" "A \<notin> \<B>" "B \<in> \<B>" for x A B
dc85b5b3a532 renamings and new material
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    69
  proof -
dc85b5b3a532 renamings and new material
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    70
    have "A \<inter> B = {}"
dc85b5b3a532 renamings and new material
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    71
      using assms disjointD AB by blast
dc85b5b3a532 renamings and new material
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    72
  with x show ?thesis
dc85b5b3a532 renamings and new material
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    73
    by blast
dc85b5b3a532 renamings and new material
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    74
  qed
dc85b5b3a532 renamings and new material
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    75
  then show ?thesis by auto
dc85b5b3a532 renamings and new material
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    76
qed
dc85b5b3a532 renamings and new material
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    77
dc85b5b3a532 renamings and new material
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    78
lemma psubset_Union_pairwise_disjoint:
dc85b5b3a532 renamings and new material
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    79
  assumes \<B>: "pairwise disjnt \<B>" and "\<A> \<subset> \<B> - {{}}"
dc85b5b3a532 renamings and new material
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    80
  shows "\<Union>\<A> \<subset> \<Union>\<B>"
dc85b5b3a532 renamings and new material
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    81
  unfolding psubset_eq
dc85b5b3a532 renamings and new material
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    82
proof
dc85b5b3a532 renamings and new material
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    83
  show "\<Union>\<A> \<subseteq> \<Union>\<B>"
dc85b5b3a532 renamings and new material
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    84
    using assms by blast
dc85b5b3a532 renamings and new material
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    85
  have "\<A> \<subseteq> \<B>" "\<Union>(\<B> - \<A> \<inter> (\<B> - {{}})) \<noteq> {}"
dc85b5b3a532 renamings and new material
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    86
    using assms by blast+
dc85b5b3a532 renamings and new material
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    87
  then show "\<Union>\<A> \<noteq> \<Union>\<B>"
dc85b5b3a532 renamings and new material
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    88
    using diff_Union_pairwise_disjoint [OF \<B>] by blast
dc85b5b3a532 renamings and new material
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    89
qed
dc85b5b3a532 renamings and new material
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    90
60727
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    91
subsubsection "Family of Disjoint Sets"
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    92
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    93
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
    94
  "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
    95
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    96
abbreviation "disjoint_family A \<equiv> disjoint_family_on A UNIV"
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
    97
63928
d81fb5b46a5c new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents: 63148
diff changeset
    98
lemma disjoint_family_elem_disjnt:
d81fb5b46a5c new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents: 63148
diff changeset
    99
  assumes "infinite A" "finite C"
d81fb5b46a5c new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents: 63148
diff changeset
   100
      and df: "disjoint_family_on B A"
d81fb5b46a5c new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents: 63148
diff changeset
   101
  obtains x where "x \<in> A" "disjnt C (B x)"
d81fb5b46a5c new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents: 63148
diff changeset
   102
proof -
d81fb5b46a5c new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents: 63148
diff changeset
   103
  have "False" if *: "\<forall>x \<in> A. \<exists>y. y \<in> C \<and> y \<in> B x"
d81fb5b46a5c new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents: 63148
diff changeset
   104
  proof -
d81fb5b46a5c new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents: 63148
diff changeset
   105
    obtain g where g: "\<forall>x \<in> A. g x \<in> C \<and> g x \<in> B x"
d81fb5b46a5c new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents: 63148
diff changeset
   106
      using * by metis
d81fb5b46a5c new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents: 63148
diff changeset
   107
    with df have "inj_on g A"
d81fb5b46a5c new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents: 63148
diff changeset
   108
      by (fastforce simp add: inj_on_def disjoint_family_on_def)
d81fb5b46a5c new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents: 63148
diff changeset
   109
    then have "infinite (g ` A)"
d81fb5b46a5c new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents: 63148
diff changeset
   110
      using \<open>infinite A\<close> finite_image_iff by blast
d81fb5b46a5c new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents: 63148
diff changeset
   111
    then show False
d81fb5b46a5c new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents: 63148
diff changeset
   112
      by (meson \<open>finite C\<close> finite_subset g image_subset_iff)
d81fb5b46a5c new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents: 63148
diff changeset
   113
  qed
d81fb5b46a5c new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents: 63148
diff changeset
   114
  then show ?thesis
d81fb5b46a5c new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents: 63148
diff changeset
   115
    by (force simp: disjnt_iff intro: that)
d81fb5b46a5c new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents: 63148
diff changeset
   116
qed
d81fb5b46a5c new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents: 63148
diff changeset
   117
60727
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   118
lemma disjoint_family_onD:
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   119
  "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
   120
  by (auto simp: disjoint_family_on_def)
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   121
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   122
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
   123
  by (force simp add: disjoint_family_on_def)
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   124
74598
5d91897a8e54 moved a theorem to a sensible place
paulson <lp15@cam.ac.uk>
parents: 74590
diff changeset
   125
lemma disjoint_family_on_insert:
5d91897a8e54 moved a theorem to a sensible place
paulson <lp15@cam.ac.uk>
parents: 74590
diff changeset
   126
  "i \<notin> I \<Longrightarrow> disjoint_family_on A (insert i I) \<longleftrightarrow> A i \<inter> (\<Union>i\<in>I. A i) = {} \<and> disjoint_family_on A I"
5d91897a8e54 moved a theorem to a sensible place
paulson <lp15@cam.ac.uk>
parents: 74590
diff changeset
   127
  by (fastforce simp: disjoint_family_on_def)
5d91897a8e54 moved a theorem to a sensible place
paulson <lp15@cam.ac.uk>
parents: 74590
diff changeset
   128
60727
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   129
lemma disjoint_family_on_bisimulation:
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   130
  assumes "disjoint_family_on f S"
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   131
  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
   132
  shows "disjoint_family_on g S"
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   133
  using assms unfolding disjoint_family_on_def by auto
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   134
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   135
lemma disjoint_family_on_mono:
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   136
  "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
   137
  unfolding disjoint_family_on_def by auto
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   138
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   139
lemma disjoint_family_Suc:
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   140
  "(\<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
   141
  using lift_Suc_mono_le[of A]
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   142
  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
   143
     (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
   144
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   145
lemma disjoint_family_on_disjoint_image:
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   146
  "disjoint_family_on A I \<Longrightarrow> disjoint (A ` I)"
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   147
  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
   148
 
60727
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   149
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
   150
  by (auto simp: disjoint_family_on_def)
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   151
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   152
lemma disjoint_image_disjoint_family_on:
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   153
  assumes d: "disjoint (A ` I)" and i: "inj_on A I"
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   154
  shows "disjoint_family_on A I"
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   155
  unfolding disjoint_family_on_def
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   156
proof (intro ballI impI)
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   157
  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
   158
  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
   159
    by (intro disjointD[OF d]) auto
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   160
qed
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   161
74438
5827b91ef30e new material from the Roth development, mostly about finite sets, disjoint famillies and partitions
paulson <lp15@cam.ac.uk>
parents: 74223
diff changeset
   162
lemma disjoint_family_on_iff_disjoint_image:
5827b91ef30e new material from the Roth development, mostly about finite sets, disjoint famillies and partitions
paulson <lp15@cam.ac.uk>
parents: 74223
diff changeset
   163
  assumes "\<And>i. i \<in> I \<Longrightarrow> A i \<noteq> {}"
5827b91ef30e new material from the Roth development, mostly about finite sets, disjoint famillies and partitions
paulson <lp15@cam.ac.uk>
parents: 74223
diff changeset
   164
  shows "disjoint_family_on A I \<longleftrightarrow> disjoint (A ` I) \<and> inj_on A I"
5827b91ef30e new material from the Roth development, mostly about finite sets, disjoint famillies and partitions
paulson <lp15@cam.ac.uk>
parents: 74223
diff changeset
   165
proof
5827b91ef30e new material from the Roth development, mostly about finite sets, disjoint famillies and partitions
paulson <lp15@cam.ac.uk>
parents: 74223
diff changeset
   166
  assume "disjoint_family_on A I"
5827b91ef30e new material from the Roth development, mostly about finite sets, disjoint famillies and partitions
paulson <lp15@cam.ac.uk>
parents: 74223
diff changeset
   167
  then show "disjoint (A ` I) \<and> inj_on A I"
5827b91ef30e new material from the Roth development, mostly about finite sets, disjoint famillies and partitions
paulson <lp15@cam.ac.uk>
parents: 74223
diff changeset
   168
    by (metis (mono_tags, lifting) assms disjoint_family_onD disjoint_family_on_disjoint_image inf.idem inj_onI)
5827b91ef30e new material from the Roth development, mostly about finite sets, disjoint famillies and partitions
paulson <lp15@cam.ac.uk>
parents: 74223
diff changeset
   169
qed (use disjoint_image_disjoint_family_on in metis)
5827b91ef30e new material from the Roth development, mostly about finite sets, disjoint famillies and partitions
paulson <lp15@cam.ac.uk>
parents: 74223
diff changeset
   170
5827b91ef30e new material from the Roth development, mostly about finite sets, disjoint famillies and partitions
paulson <lp15@cam.ac.uk>
parents: 74223
diff changeset
   171
lemma card_UN_disjoint':
5827b91ef30e new material from the Roth development, mostly about finite sets, disjoint famillies and partitions
paulson <lp15@cam.ac.uk>
parents: 74223
diff changeset
   172
  assumes "disjoint_family_on A I" "\<And>i. i \<in> I \<Longrightarrow> finite (A i)" "finite I"
5827b91ef30e new material from the Roth development, mostly about finite sets, disjoint famillies and partitions
paulson <lp15@cam.ac.uk>
parents: 74223
diff changeset
   173
  shows "card (\<Union>i\<in>I. A i) = (\<Sum>i\<in>I. card (A i))"
5827b91ef30e new material from the Roth development, mostly about finite sets, disjoint famillies and partitions
paulson <lp15@cam.ac.uk>
parents: 74223
diff changeset
   174
  using assms   by (simp add: card_UN_disjoint disjoint_family_on_def)
5827b91ef30e new material from the Roth development, mostly about finite sets, disjoint famillies and partitions
paulson <lp15@cam.ac.uk>
parents: 74223
diff changeset
   175
60727
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   176
lemma disjoint_UN:
69745
aec42cee2521 more canonical and less specialized syntax
nipkow
parents: 69712
diff changeset
   177
  assumes F: "\<And>i. i \<in> I \<Longrightarrow> disjoint (F i)" and *: "disjoint_family_on (\<lambda>i. \<Union>(F i)) I"
60727
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   178
  shows "disjoint (\<Union>i\<in>I. F i)"
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   179
proof (safe intro!: disjointI del: equalityI)
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   180
  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
   181
  show "A \<inter> B = {}"
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   182
  proof cases
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   183
    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
   184
      by (auto dest: disjointD)
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   185
  next
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   186
    assume "i \<noteq> j"
69745
aec42cee2521 more canonical and less specialized syntax
nipkow
parents: 69712
diff changeset
   187
    with * \<open>i\<in>I\<close> \<open>j\<in>I\<close> have "(\<Union>(F i)) \<inter> (\<Union>(F j)) = {}"
60727
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   188
      by (rule disjoint_family_onD)
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   189
    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
   190
    show "A \<inter> B = {}"
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   191
      by auto
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
qed
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   194
63099
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
   195
lemma distinct_list_bind: 
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
   196
  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
   197
          "disjoint_family_on (set \<circ> f) (set xs)"
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
   198
  shows   "distinct (List.bind xs f)"
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
   199
  using assms
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
   200
  by (induction xs)
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
   201
     (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
   202
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
   203
lemma bij_betw_UNION_disjoint:
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
   204
  assumes disj: "disjoint_family_on A' I"
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
   205
  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
   206
  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
   207
unfolding bij_betw_def
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
   208
proof
69313
b021008c5397 removed legacy input syntax
haftmann
parents: 68406
diff changeset
   209
  from bij show eq: "f ` \<Union>(A ` I) = \<Union>(A' ` I)"
63099
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
   210
    by (auto simp: bij_betw_def image_UN)
69313
b021008c5397 removed legacy input syntax
haftmann
parents: 68406
diff changeset
   211
  show "inj_on f (\<Union>(A ` I))"
63099
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
   212
  proof (rule inj_onI, clarify)
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
   213
    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
   214
    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
   215
      by (auto simp: bij_betw_def)
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
   216
    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
   217
    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
   218
    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
   219
  qed
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
   220
qed
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
   221
60727
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   222
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
   223
  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
   224
63099
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
   225
text \<open>
74438
5827b91ef30e new material from the Roth development, mostly about finite sets, disjoint famillies and partitions
paulson <lp15@cam.ac.uk>
parents: 74223
diff changeset
   226
  Sum/product of the union of a finite disjoint family
5827b91ef30e new material from the Roth development, mostly about finite sets, disjoint famillies and partitions
paulson <lp15@cam.ac.uk>
parents: 74223
diff changeset
   227
\<close>
5827b91ef30e new material from the Roth development, mostly about finite sets, disjoint famillies and partitions
paulson <lp15@cam.ac.uk>
parents: 74223
diff changeset
   228
context comm_monoid_set
5827b91ef30e new material from the Roth development, mostly about finite sets, disjoint famillies and partitions
paulson <lp15@cam.ac.uk>
parents: 74223
diff changeset
   229
begin
5827b91ef30e new material from the Roth development, mostly about finite sets, disjoint famillies and partitions
paulson <lp15@cam.ac.uk>
parents: 74223
diff changeset
   230
5827b91ef30e new material from the Roth development, mostly about finite sets, disjoint famillies and partitions
paulson <lp15@cam.ac.uk>
parents: 74223
diff changeset
   231
lemma UNION_disjoint_family:
5827b91ef30e new material from the Roth development, mostly about finite sets, disjoint famillies and partitions
paulson <lp15@cam.ac.uk>
parents: 74223
diff changeset
   232
  assumes "finite I" and "\<forall>i\<in>I. finite (A i)"
5827b91ef30e new material from the Roth development, mostly about finite sets, disjoint famillies and partitions
paulson <lp15@cam.ac.uk>
parents: 74223
diff changeset
   233
    and "disjoint_family_on A I"
5827b91ef30e new material from the Roth development, mostly about finite sets, disjoint famillies and partitions
paulson <lp15@cam.ac.uk>
parents: 74223
diff changeset
   234
  shows "F g (\<Union>(A ` I)) = F (\<lambda>x. F g (A x)) I"
5827b91ef30e new material from the Roth development, mostly about finite sets, disjoint famillies and partitions
paulson <lp15@cam.ac.uk>
parents: 74223
diff changeset
   235
  using assms unfolding disjoint_family_on_def  by (rule UNION_disjoint)
5827b91ef30e new material from the Roth development, mostly about finite sets, disjoint famillies and partitions
paulson <lp15@cam.ac.uk>
parents: 74223
diff changeset
   236
5827b91ef30e new material from the Roth development, mostly about finite sets, disjoint famillies and partitions
paulson <lp15@cam.ac.uk>
parents: 74223
diff changeset
   237
lemma Union_disjoint_sets:
5827b91ef30e new material from the Roth development, mostly about finite sets, disjoint famillies and partitions
paulson <lp15@cam.ac.uk>
parents: 74223
diff changeset
   238
  assumes "\<forall>A\<in>C. finite A" and "disjoint C"
5827b91ef30e new material from the Roth development, mostly about finite sets, disjoint famillies and partitions
paulson <lp15@cam.ac.uk>
parents: 74223
diff changeset
   239
  shows "F g (\<Union>C) = (F \<circ> F) g C"
5827b91ef30e new material from the Roth development, mostly about finite sets, disjoint famillies and partitions
paulson <lp15@cam.ac.uk>
parents: 74223
diff changeset
   240
  using assms unfolding disjoint_def  by (rule Union_disjoint)
5827b91ef30e new material from the Roth development, mostly about finite sets, disjoint famillies and partitions
paulson <lp15@cam.ac.uk>
parents: 74223
diff changeset
   241
5827b91ef30e new material from the Roth development, mostly about finite sets, disjoint famillies and partitions
paulson <lp15@cam.ac.uk>
parents: 74223
diff changeset
   242
end
5827b91ef30e new material from the Roth development, mostly about finite sets, disjoint famillies and partitions
paulson <lp15@cam.ac.uk>
parents: 74223
diff changeset
   243
5827b91ef30e new material from the Roth development, mostly about finite sets, disjoint famillies and partitions
paulson <lp15@cam.ac.uk>
parents: 74223
diff changeset
   244
text \<open>
63099
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
   245
  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
   246
\<close>
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
   247
lemma infinite_disjoint_family_imp_infinite_UNION:
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
   248
  assumes "\<not>finite A" "\<And>x. x \<in> A \<Longrightarrow> f x \<noteq> {}" "disjoint_family_on f A"
69313
b021008c5397 removed legacy input syntax
haftmann
parents: 68406
diff changeset
   249
  shows   "\<not>finite (\<Union>(f ` A))"
63099
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
   250
proof -
63148
6a767355d1a9 updated 'define';
wenzelm
parents: 63145
diff changeset
   251
  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
   252
  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
   253
    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
   254
  have inj_on_g: "inj_on g A"
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
   255
  proof (rule inj_onI, rule ccontr)
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
   256
    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
   257
    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
   258
    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
   259
      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
   260
  qed
69313
b021008c5397 removed legacy input syntax
haftmann
parents: 68406
diff changeset
   261
  from g have "g ` A \<subseteq> \<Union>(f ` A)" by blast
63099
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
   262
  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
   263
    using finite_imageD by blast
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
   264
  ultimately show ?thesis using finite_subset by blast
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
   265
qed  
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
   266
  
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 62843
diff changeset
   267
60727
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   268
subsection \<open>Construct Disjoint Sequences\<close>
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   269
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   270
definition disjointed :: "(nat \<Rightarrow> 'a set) \<Rightarrow> nat \<Rightarrow> 'a set" where
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   271
  "disjointed A n = A n - (\<Union>i\<in>{0..<n}. A i)"
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   272
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   273
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
   274
proof (induct n)
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   275
  case 0 show ?case by simp
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   276
next
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   277
  case (Suc n)
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   278
  thus ?case by (simp add: atLeastLessThanSuc disjointed_def)
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   279
qed
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   280
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   281
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
   282
  by (rule UN_finite2_eq [where k=0])
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   283
     (simp add: finite_UN_disjointed_eq)
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   284
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   285
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
   286
  by (auto simp add: disjointed_def)
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   287
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   288
lemma disjoint_family_disjointed: "disjoint_family (disjointed A)"
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   289
  by (simp add: disjoint_family_on_def)
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   290
     (metis neq_iff Int_commute less_disjoint_disjointed)
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   291
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   292
lemma disjointed_subset: "disjointed A n \<subseteq> A n"
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   293
  by (auto simp add: disjointed_def)
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   294
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   295
lemma disjointed_0[simp]: "disjointed A 0 = A 0"
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   296
  by (simp add: disjointed_def)
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   297
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   298
lemma disjointed_mono: "mono A \<Longrightarrow> disjointed A (Suc n) = A (Suc n) - A n"
74590
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   299
  using mono_imp_UN_eq_last[of A] by (simp add: disjointed_def atLeastLessThanSuc_atLeastAtMost atLeast0AtMost)
60727
53697011b03a move disjoint sets to their own theory
hoelzl
parents:
diff changeset
   300
63098
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   301
subsection \<open>Partitions\<close>
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   302
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   303
text \<open>
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69313
diff changeset
   304
  Partitions \<^term>\<open>P\<close> of a set \<^term>\<open>A\<close>. We explicitly disallow empty sets.
63098
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   305
\<close>
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   306
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   307
definition partition_on :: "'a set \<Rightarrow> 'a set set \<Rightarrow> bool"
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   308
where
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   309
  "partition_on A P \<longleftrightarrow> \<Union>P = A \<and> disjoint P \<and> {} \<notin> P"
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   310
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   311
lemma partition_onI:
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   312
  "\<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
   313
  by (auto simp: partition_on_def pairwise_def)
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   314
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   315
lemma partition_onD1: "partition_on A P \<Longrightarrow> A = \<Union>P"
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   316
  by (auto simp: partition_on_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_onD2: "partition_on A P \<Longrightarrow> disjoint P"
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   319
  by (auto simp: partition_on_def)
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   320
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   321
lemma partition_onD3: "partition_on A P \<Longrightarrow> {} \<notin> P"
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   322
  by (auto simp: partition_on_def)
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   323
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   324
subsection \<open>Constructions of partitions\<close>
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   325
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   326
lemma partition_on_empty: "partition_on {} P \<longleftrightarrow> P = {}"
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   327
  unfolding partition_on_def by fastforce
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   328
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   329
lemma partition_on_space: "A \<noteq> {} \<Longrightarrow> partition_on A {A}"
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   330
  by (auto simp: partition_on_def disjoint_def)
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   331
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   332
lemma partition_on_singletons: "partition_on A ((\<lambda>x. {x}) ` A)"
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   333
  by (auto simp: partition_on_def disjoint_def)
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   334
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   335
lemma partition_on_transform:
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   336
  assumes P: "partition_on A P"
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   337
  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
   338
  shows "partition_on (F A) (F ` P - {{}})"
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   339
proof -
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   340
  have "\<Union>(F ` P - {{}}) = F A"
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   341
    unfolding P[THEN partition_onD1] F_UN[symmetric] by auto
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   342
  with P show ?thesis
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   343
    by (auto simp add: partition_on_def pairwise_def intro!: F_disjnt)
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   344
qed
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   345
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 63928
diff changeset
   346
lemma partition_on_restrict: "partition_on A P \<Longrightarrow> partition_on (B \<inter> A) ((\<inter>) B ` P - {{}})"
63098
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   347
  by (intro partition_on_transform) (auto simp: disjnt_def)
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   348
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 63928
diff changeset
   349
lemma partition_on_vimage: "partition_on A P \<Longrightarrow> partition_on (f -` A) ((-`) f ` P - {{}})"
63098
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   350
  by (intro partition_on_transform) (auto simp: disjnt_def)
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   351
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   352
lemma partition_on_inj_image:
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   353
  assumes P: "partition_on A P" and f: "inj_on f A"
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 63928
diff changeset
   354
  shows "partition_on (f ` A) ((`) f ` P - {{}})"
63098
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   355
proof (rule partition_on_transform[OF P])
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   356
  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
   357
    using f[THEN inj_onD] P[THEN partition_onD1] by (auto simp: disjnt_def)
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   358
qed auto
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   359
74223
527088d4a89b strengthened a few lemmas about finite sets and added a code equation for complex_of_real
paulson <lp15@cam.ac.uk>
parents: 73477
diff changeset
   360
lemma partition_on_insert:
527088d4a89b strengthened a few lemmas about finite sets and added a code equation for complex_of_real
paulson <lp15@cam.ac.uk>
parents: 73477
diff changeset
   361
  assumes "disjnt p (\<Union>P)"
527088d4a89b strengthened a few lemmas about finite sets and added a code equation for complex_of_real
paulson <lp15@cam.ac.uk>
parents: 73477
diff changeset
   362
  shows "partition_on A (insert p P) \<longleftrightarrow> partition_on (A-p) P \<and> p \<subseteq> A \<and> p \<noteq> {}"
527088d4a89b strengthened a few lemmas about finite sets and added a code equation for complex_of_real
paulson <lp15@cam.ac.uk>
parents: 73477
diff changeset
   363
  using assms
527088d4a89b strengthened a few lemmas about finite sets and added a code equation for complex_of_real
paulson <lp15@cam.ac.uk>
parents: 73477
diff changeset
   364
  by (auto simp: partition_on_def disjnt_iff pairwise_insert)
527088d4a89b strengthened a few lemmas about finite sets and added a code equation for complex_of_real
paulson <lp15@cam.ac.uk>
parents: 73477
diff changeset
   365
63098
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   366
subsection \<open>Finiteness of partitions\<close>
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   367
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   368
lemma finitely_many_partition_on:
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   369
  assumes "finite A"
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   370
  shows "finite {P. partition_on A P}"
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   371
proof (rule finite_subset)
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   372
  show "{P. partition_on A P} \<subseteq> Pow (Pow A)"
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   373
    unfolding partition_on_def by auto
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   374
  show "finite (Pow (Pow A))"
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   375
    using assms by simp
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   376
qed
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   377
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   378
lemma finite_elements: "finite A \<Longrightarrow> partition_on A P \<Longrightarrow> finite P"
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   379
  using partition_onD1[of A P] by (simp add: finite_UnionD)
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   380
74438
5827b91ef30e new material from the Roth development, mostly about finite sets, disjoint famillies and partitions
paulson <lp15@cam.ac.uk>
parents: 74223
diff changeset
   381
lemma product_partition:
5827b91ef30e new material from the Roth development, mostly about finite sets, disjoint famillies and partitions
paulson <lp15@cam.ac.uk>
parents: 74223
diff changeset
   382
  assumes "partition_on A P" and "\<And>p. p \<in> P \<Longrightarrow> finite p" 
5827b91ef30e new material from the Roth development, mostly about finite sets, disjoint famillies and partitions
paulson <lp15@cam.ac.uk>
parents: 74223
diff changeset
   383
  shows "card A = (\<Sum>p\<in>P. card p)"
5827b91ef30e new material from the Roth development, mostly about finite sets, disjoint famillies and partitions
paulson <lp15@cam.ac.uk>
parents: 74223
diff changeset
   384
  using assms unfolding partition_on_def  by (meson card_Union_disjoint)
5827b91ef30e new material from the Roth development, mostly about finite sets, disjoint famillies and partitions
paulson <lp15@cam.ac.uk>
parents: 74223
diff changeset
   385
63098
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   386
subsection \<open>Equivalence of partitions and equivalence classes\<close>
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   387
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   388
lemma partition_on_quotient:
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   389
  assumes r: "equiv A r"
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   390
  shows "partition_on A (A // r)"
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   391
proof (rule partition_onI)
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   392
  from r have "refl_on A r"
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   393
    by (auto elim: equivE)
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   394
  then show "\<Union>(A // r) = A" "{} \<notin> A // r"
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   395
    by (auto simp: refl_on_def quotient_def)
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   396
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   397
  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
   398
  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
   399
    by (auto simp: quotient_def)
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   400
  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
   401
    by (auto simp: disjnt_equiv_class)
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   402
qed
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   403
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   404
lemma equiv_partition_on:
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   405
  assumes P: "partition_on A P"
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   406
  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
   407
proof (rule equivI)
80067
c40bdfc84640 tuned proofs of Equiv_Relations.equiv
desharna
parents: 74598
diff changeset
   408
  have "A = \<Union>P"
63098
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   409
    using P by (auto simp: partition_on_def)
80067
c40bdfc84640 tuned proofs of Equiv_Relations.equiv
desharna
parents: 74598
diff changeset
   410
c40bdfc84640 tuned proofs of Equiv_Relations.equiv
desharna
parents: 74598
diff changeset
   411
  have "{(x, y). \<exists>p \<in> P. x \<in> p \<and> y \<in> p} \<subseteq> A \<times> A"
c40bdfc84640 tuned proofs of Equiv_Relations.equiv
desharna
parents: 74598
diff changeset
   412
    unfolding \<open>A = \<Union>P\<close> by blast
63098
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   413
  then show "refl_on A {(x, y). \<exists>p\<in>P. x \<in> p \<and> y \<in> p}"
80067
c40bdfc84640 tuned proofs of Equiv_Relations.equiv
desharna
parents: 74598
diff changeset
   414
    unfolding refl_on_def \<open>A = \<Union>P\<close> by auto
c40bdfc84640 tuned proofs of Equiv_Relations.equiv
desharna
parents: 74598
diff changeset
   415
next
63098
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   416
  show "trans {(x, y). \<exists>p\<in>P. x \<in> p \<and> y \<in> p}"
80067
c40bdfc84640 tuned proofs of Equiv_Relations.equiv
desharna
parents: 74598
diff changeset
   417
    using P by (auto simp only: trans_def disjoint_def partition_on_def)
c40bdfc84640 tuned proofs of Equiv_Relations.equiv
desharna
parents: 74598
diff changeset
   418
next
c40bdfc84640 tuned proofs of Equiv_Relations.equiv
desharna
parents: 74598
diff changeset
   419
  show "sym {(x, y). \<exists>p\<in>P. x \<in> p \<and> y \<in> p}"
c40bdfc84640 tuned proofs of Equiv_Relations.equiv
desharna
parents: 74598
diff changeset
   420
    by (auto simp only: sym_def)
c40bdfc84640 tuned proofs of Equiv_Relations.equiv
desharna
parents: 74598
diff changeset
   421
qed
63098
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   422
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   423
lemma partition_on_eq_quotient:
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   424
  assumes P: "partition_on A P"
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   425
  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
   426
  unfolding quotient_def
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   427
proof safe
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   428
  fix x assume "x \<in> A"
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   429
  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
   430
    using P by (auto simp: partition_on_def disjoint_def)
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   431
  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
   432
    by (safe intro!: bexI[of _ p]) simp
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   433
  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
   434
    by (simp add: \<open>p \<in> P\<close>)
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   435
next
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   436
  fix p assume "p \<in> P"
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   437
  then have "p \<noteq> {}"
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   438
    using P by (auto simp: partition_on_def)
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   439
  then obtain x where "x \<in> p"
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   440
    by auto
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   441
  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
   442
    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
   443
  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
   444
    by (safe intro!: bexI[of _ p]) simp
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   445
  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
   446
    by (auto intro: \<open>x \<in> A\<close>)
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   447
qed
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   448
56f03591898b Library: add partition_on
hoelzl
parents: 62843
diff changeset
   449
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
   450
  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
   451
74574
cff477b6d015 Refinement of partitions
paulson <lp15@cam.ac.uk>
parents: 74438
diff changeset
   452
subsection \<open>Refinement of partitions\<close>
cff477b6d015 Refinement of partitions
paulson <lp15@cam.ac.uk>
parents: 74438
diff changeset
   453
cff477b6d015 Refinement of partitions
paulson <lp15@cam.ac.uk>
parents: 74438
diff changeset
   454
definition refines :: "'a set \<Rightarrow> 'a set set \<Rightarrow> 'a set set \<Rightarrow> bool"
cff477b6d015 Refinement of partitions
paulson <lp15@cam.ac.uk>
parents: 74438
diff changeset
   455
  where "refines A P Q \<equiv>
cff477b6d015 Refinement of partitions
paulson <lp15@cam.ac.uk>
parents: 74438
diff changeset
   456
        partition_on A P \<and> partition_on A Q \<and> (\<forall>X\<in>P. \<exists>Y\<in>Q. X \<subseteq> Y)" 
cff477b6d015 Refinement of partitions
paulson <lp15@cam.ac.uk>
parents: 74438
diff changeset
   457
cff477b6d015 Refinement of partitions
paulson <lp15@cam.ac.uk>
parents: 74438
diff changeset
   458
lemma refines_refl: "partition_on A P \<Longrightarrow> refines A P P"
cff477b6d015 Refinement of partitions
paulson <lp15@cam.ac.uk>
parents: 74438
diff changeset
   459
  using refines_def by blast
cff477b6d015 Refinement of partitions
paulson <lp15@cam.ac.uk>
parents: 74438
diff changeset
   460
cff477b6d015 Refinement of partitions
paulson <lp15@cam.ac.uk>
parents: 74438
diff changeset
   461
lemma refines_asym1:
cff477b6d015 Refinement of partitions
paulson <lp15@cam.ac.uk>
parents: 74438
diff changeset
   462
  assumes "refines A P Q" "refines A Q P"
cff477b6d015 Refinement of partitions
paulson <lp15@cam.ac.uk>
parents: 74438
diff changeset
   463
  shows "P \<subseteq> Q"
cff477b6d015 Refinement of partitions
paulson <lp15@cam.ac.uk>
parents: 74438
diff changeset
   464
proof 
cff477b6d015 Refinement of partitions
paulson <lp15@cam.ac.uk>
parents: 74438
diff changeset
   465
  fix X
cff477b6d015 Refinement of partitions
paulson <lp15@cam.ac.uk>
parents: 74438
diff changeset
   466
  assume "X \<in> P"
cff477b6d015 Refinement of partitions
paulson <lp15@cam.ac.uk>
parents: 74438
diff changeset
   467
  then obtain Y X' where "Y \<in> Q" "X \<subseteq> Y" "X' \<in> P" "Y \<subseteq> X'"
cff477b6d015 Refinement of partitions
paulson <lp15@cam.ac.uk>
parents: 74438
diff changeset
   468
    by (meson assms refines_def)
cff477b6d015 Refinement of partitions
paulson <lp15@cam.ac.uk>
parents: 74438
diff changeset
   469
  then have "X' = X"
cff477b6d015 Refinement of partitions
paulson <lp15@cam.ac.uk>
parents: 74438
diff changeset
   470
    using assms(2) unfolding partition_on_def refines_def
cff477b6d015 Refinement of partitions
paulson <lp15@cam.ac.uk>
parents: 74438
diff changeset
   471
    by (metis \<open>X \<in> P\<close> \<open>X \<subseteq> Y\<close> disjnt_self_iff_empty disjnt_subset1 pairwiseD)
cff477b6d015 Refinement of partitions
paulson <lp15@cam.ac.uk>
parents: 74438
diff changeset
   472
  then show "X \<in> Q"
cff477b6d015 Refinement of partitions
paulson <lp15@cam.ac.uk>
parents: 74438
diff changeset
   473
    using \<open>X \<subseteq> Y\<close> \<open>Y \<in> Q\<close> \<open>Y \<subseteq> X'\<close> by force
cff477b6d015 Refinement of partitions
paulson <lp15@cam.ac.uk>
parents: 74438
diff changeset
   474
qed
cff477b6d015 Refinement of partitions
paulson <lp15@cam.ac.uk>
parents: 74438
diff changeset
   475
cff477b6d015 Refinement of partitions
paulson <lp15@cam.ac.uk>
parents: 74438
diff changeset
   476
lemma refines_asym: "\<lbrakk>refines A P Q; refines A Q P\<rbrakk> \<Longrightarrow> P=Q"
cff477b6d015 Refinement of partitions
paulson <lp15@cam.ac.uk>
parents: 74438
diff changeset
   477
  by (meson antisym_conv refines_asym1)
cff477b6d015 Refinement of partitions
paulson <lp15@cam.ac.uk>
parents: 74438
diff changeset
   478
cff477b6d015 Refinement of partitions
paulson <lp15@cam.ac.uk>
parents: 74438
diff changeset
   479
lemma refines_trans: "\<lbrakk>refines A P Q; refines A Q R\<rbrakk> \<Longrightarrow> refines A P R"
cff477b6d015 Refinement of partitions
paulson <lp15@cam.ac.uk>
parents: 74438
diff changeset
   480
  by (meson order.trans refines_def)
cff477b6d015 Refinement of partitions
paulson <lp15@cam.ac.uk>
parents: 74438
diff changeset
   481
cff477b6d015 Refinement of partitions
paulson <lp15@cam.ac.uk>
parents: 74438
diff changeset
   482
lemma refines_obtains_subset:
cff477b6d015 Refinement of partitions
paulson <lp15@cam.ac.uk>
parents: 74438
diff changeset
   483
  assumes "refines A P Q" "q \<in> Q"
cff477b6d015 Refinement of partitions
paulson <lp15@cam.ac.uk>
parents: 74438
diff changeset
   484
  shows "partition_on q {p \<in> P. p \<subseteq> q}"
cff477b6d015 Refinement of partitions
paulson <lp15@cam.ac.uk>
parents: 74438
diff changeset
   485
proof -
cff477b6d015 Refinement of partitions
paulson <lp15@cam.ac.uk>
parents: 74438
diff changeset
   486
  have "p \<subseteq> q \<or> disjnt p q" if "p \<in> P" for p
cff477b6d015 Refinement of partitions
paulson <lp15@cam.ac.uk>
parents: 74438
diff changeset
   487
    using that assms unfolding refines_def partition_on_def disjoint_def
cff477b6d015 Refinement of partitions
paulson <lp15@cam.ac.uk>
parents: 74438
diff changeset
   488
    by (metis disjnt_def disjnt_subset1)
cff477b6d015 Refinement of partitions
paulson <lp15@cam.ac.uk>
parents: 74438
diff changeset
   489
  with assms have "q \<subseteq> Union {p \<in> P. p \<subseteq> q}"
cff477b6d015 Refinement of partitions
paulson <lp15@cam.ac.uk>
parents: 74438
diff changeset
   490
    using assms 
cff477b6d015 Refinement of partitions
paulson <lp15@cam.ac.uk>
parents: 74438
diff changeset
   491
   by (clarsimp simp: refines_def disjnt_iff partition_on_def) (metis Union_iff)
cff477b6d015 Refinement of partitions
paulson <lp15@cam.ac.uk>
parents: 74438
diff changeset
   492
  with assms have "q = Union {p \<in> P. p \<subseteq> q}"
cff477b6d015 Refinement of partitions
paulson <lp15@cam.ac.uk>
parents: 74438
diff changeset
   493
    by auto
cff477b6d015 Refinement of partitions
paulson <lp15@cam.ac.uk>
parents: 74438
diff changeset
   494
  then show ?thesis
cff477b6d015 Refinement of partitions
paulson <lp15@cam.ac.uk>
parents: 74438
diff changeset
   495
    using assms by (auto simp: refines_def disjoint_def partition_on_def)
cff477b6d015 Refinement of partitions
paulson <lp15@cam.ac.uk>
parents: 74438
diff changeset
   496
qed 
cff477b6d015 Refinement of partitions
paulson <lp15@cam.ac.uk>
parents: 74438
diff changeset
   497
74590
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   498
subsection \<open>The coarsest common refinement of a set of partitions\<close>
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   499
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   500
definition common_refinement :: "'a set set set \<Rightarrow> 'a set set"
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   501
  where "common_refinement \<P> \<equiv> (\<Union>f \<in> (\<Pi>\<^sub>E P\<in>\<P>. P). {\<Inter> (f ` \<P>)}) - {{}}" 
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   502
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   503
text \<open>With non-extensional function space\<close>
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   504
lemma common_refinement: "common_refinement \<P> = (\<Union>f \<in> (\<Pi> P\<in>\<P>. P). {\<Inter> (f ` \<P>)}) - {{}}" 
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   505
  (is "?lhs = ?rhs")
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   506
proof
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   507
  show "?rhs \<subseteq> ?lhs"
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   508
    apply (clarsimp simp add: common_refinement_def PiE_def Ball_def)
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   509
    by (metis restrict_Pi_cancel image_restrict_eq restrict_extensional)
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   510
qed (auto simp add: common_refinement_def PiE_def)
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   511
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   512
lemma common_refinement_exists: "\<lbrakk>X \<in> common_refinement \<P>; P \<in> \<P>\<rbrakk> \<Longrightarrow> \<exists>R\<in>P. X \<subseteq> R"
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   513
  by (auto simp add: common_refinement)
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   514
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   515
lemma Union_common_refinement: "\<Union> (common_refinement \<P>) = (\<Inter> P\<in>\<P>. \<Union>P)"
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   516
proof
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   517
  show "(\<Inter> P\<in>\<P>. \<Union>P) \<subseteq> \<Union> (common_refinement \<P>)"
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   518
  proof (clarsimp simp: common_refinement)
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   519
    fix x
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   520
    assume "\<forall>P\<in>\<P>. \<exists>X\<in>P. x \<in> X"
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   521
    then obtain F where F: "\<And>P. P\<in>\<P> \<Longrightarrow> F P \<in> P \<and> x \<in> F P"
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   522
      by metis
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   523
    then have "x \<in> \<Inter> (F ` \<P>)"
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   524
      by force
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   525
    with F show "\<exists>X\<in>(\<Union>x\<in>\<Pi> P\<in>\<P>. P. {\<Inter> (x ` \<P>)}) - {{}}. x \<in> X"
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   526
      by (auto simp add: Pi_iff Bex_def)
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   527
  qed
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   528
qed (auto simp: common_refinement_def)
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   529
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   530
lemma partition_on_common_refinement:
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   531
  assumes A: "\<And>P. P \<in> \<P> \<Longrightarrow> partition_on A P" and "\<P> \<noteq> {}"
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   532
  shows "partition_on A (common_refinement \<P>)"
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   533
proof (rule partition_onI)
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   534
  show "\<Union> (common_refinement \<P>) = A"
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   535
    using assms by (simp add: partition_on_def Union_common_refinement)
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   536
  fix P Q
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   537
  assume "P \<in> common_refinement \<P>" and "Q \<in> common_refinement \<P>" and "P \<noteq> Q"
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   538
  then obtain f g where f: "f \<in> (\<Pi>\<^sub>E P\<in>\<P>. P)" and P: "P = \<Inter> (f ` \<P>)" and "P \<noteq> {}"
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   539
                  and   g: "g \<in> (\<Pi>\<^sub>E P\<in>\<P>. P)" and Q: "Q = \<Inter> (g ` \<P>)" and "Q \<noteq> {}"
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   540
    by (auto simp add: common_refinement_def)
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   541
  have "f=g" if "x \<in> P" "x \<in> Q" for x
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   542
  proof (rule extensionalityI [of _ \<P>])
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   543
    fix R
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   544
    assume "R \<in> \<P>"
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   545
    with that P Q f g A [unfolded partition_on_def, OF \<open>R \<in> \<P>\<close>]
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   546
    show "f R = g R"
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   547
      by (metis INT_E Int_iff PiE_iff disjointD emptyE)
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   548
  qed (use PiE_iff f g in auto)
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   549
  then show "disjnt P Q"
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   550
    by (metis P Q \<open>P \<noteq> Q\<close> disjnt_iff) 
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   551
qed (simp add: common_refinement_def)
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   552
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   553
lemma refines_common_refinement:
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   554
  assumes "\<And>P. P \<in> \<P> \<Longrightarrow> partition_on A P" "P \<in> \<P>"
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   555
  shows "refines A (common_refinement \<P>) P"
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   556
  unfolding refines_def
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   557
proof (intro conjI strip)
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   558
  fix X
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   559
  assume "X \<in> common_refinement \<P>"
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   560
  with assms show "\<exists>Y\<in>P. X \<subseteq> Y"
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   561
    by (auto simp: common_refinement_def)
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   562
qed (use assms partition_on_common_refinement in auto)
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   563
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   564
text \<open>The common refinement is itself refined by any other\<close>
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   565
lemma common_refinement_coarsest:
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   566
  assumes "\<And>P. P \<in> \<P> \<Longrightarrow> partition_on A P" "partition_on A R" "\<And>P. P \<in> \<P> \<Longrightarrow> refines A R P" "\<P> \<noteq> {}"
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   567
  shows "refines A R (common_refinement \<P>)"
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   568
  unfolding refines_def
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   569
proof (intro conjI ballI partition_on_common_refinement)
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   570
  fix X
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   571
  assume "X \<in> R"
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   572
  have "\<exists>p \<in> P. X \<subseteq> p" if "P \<in> \<P>" for P
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   573
    by (meson \<open>X \<in> R\<close> assms(3) refines_def that)
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   574
  then obtain F where f: "\<And>P. P \<in> \<P> \<Longrightarrow> F P \<in> P \<and> X \<subseteq> F P"
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   575
    by metis
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   576
  with \<open>partition_on A R\<close> \<open>X \<in> R\<close> \<open>\<P> \<noteq> {}\<close>
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   577
  have "\<Inter> (F ` \<P>) \<in> common_refinement \<P>"
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   578
    apply (simp add: partition_on_def common_refinement Pi_iff Bex_def)
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   579
    by (metis (no_types, lifting) cINF_greatest subset_empty)
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   580
  with f show "\<exists>Y\<in>common_refinement \<P>. X \<subseteq> Y"
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   581
    by (metis \<open>\<P> \<noteq> {}\<close> cINF_greatest)
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   582
qed (use assms in auto)
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   583
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   584
lemma finite_common_refinement:
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   585
  assumes "finite \<P>" "\<And>P. P \<in> \<P> \<Longrightarrow> finite P"
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   586
  shows "finite (common_refinement \<P>)"
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   587
proof -
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   588
  have "finite (\<Pi>\<^sub>E P\<in>\<P>. P)"
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   589
    by (simp add: assms finite_PiE)
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   590
  then show ?thesis
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   591
    by (auto simp: common_refinement_def)
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   592
qed
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   593
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   594
lemma card_common_refinement:
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   595
  assumes "finite \<P>" "\<And>P. P \<in> \<P> \<Longrightarrow> finite P"
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   596
  shows "card (common_refinement \<P>) \<le> (\<Prod>P \<in> \<P>. card P)"
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   597
proof -
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   598
  have "card (common_refinement \<P>) \<le> card (\<Union>f \<in> (\<Pi>\<^sub>E P\<in>\<P>. P). {\<Inter> (f ` \<P>)})"
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   599
    unfolding common_refinement_def by (meson card_Diff1_le)
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   600
  also have "\<dots> \<le> (\<Sum>f\<in>(\<Pi>\<^sub>E P\<in>\<P>. P). card{\<Inter> (f ` \<P>)})"
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   601
    by (metis assms finite_PiE card_UN_le)
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   602
  also have "\<dots> = card(\<Pi>\<^sub>E P\<in>\<P>. P)"
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   603
    by simp
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   604
  also have "\<dots> = (\<Prod>P \<in> \<P>. card P)"
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   605
    by (simp add: assms(1) card_PiE dual_order.eq_iff)
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   606
  finally show ?thesis .
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   607
qed
00ffae972fc0 Added / moved some simple set-theoretic lemmas
paulson <lp15@cam.ac.uk>
parents: 74574
diff changeset
   608
62390
842917225d56 more canonical names
nipkow
parents: 61824
diff changeset
   609
end