src/HOL/Library/More_Set.thy
author bulwahn
Thu, 01 Dec 2011 22:14:35 +0100
changeset 45724 1f5fc44254d7
parent 45012 060f76635bfe
child 45974 2b043ed911ac
permissions -rw-r--r--
the simple exhaustive compilation also indicates if counterexample is potentially spurious;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
31807
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
     1
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
     2
(* Author: Florian Haftmann, TU Muenchen *)
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
     3
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
     4
header {* Relating (finite) sets and lists *}
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
     5
37024
e938a0b5286e renamed List_Set to the now more appropriate More_Set
haftmann
parents: 37023
diff changeset
     6
theory More_Set
37023
efc202e1677e added theory More_List
haftmann
parents: 34977
diff changeset
     7
imports Main More_List
31807
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
     8
begin
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
     9
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    10
subsection {* Various additional set functions *}
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    11
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    12
definition is_empty :: "'a set \<Rightarrow> bool" where
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    13
  "is_empty A \<longleftrightarrow> A = {}"
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    14
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    15
definition remove :: "'a \<Rightarrow> 'a set \<Rightarrow> 'a set" where
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    16
  "remove x A = A - {x}"
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    17
42871
1c0b99f950d9 names of fold_set locales resemble name of characteristic property more closely
haftmann
parents: 42868
diff changeset
    18
lemma comp_fun_idem_remove:
1c0b99f950d9 names of fold_set locales resemble name of characteristic property more closely
haftmann
parents: 42868
diff changeset
    19
  "comp_fun_idem remove"
31807
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    20
proof -
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 39198
diff changeset
    21
  have rem: "remove = (\<lambda>x A. A - {x})" by (simp add: fun_eq_iff remove_def)
42871
1c0b99f950d9 names of fold_set locales resemble name of characteristic property more closely
haftmann
parents: 42868
diff changeset
    22
  show ?thesis by (simp only: comp_fun_idem_remove rem)
31807
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    23
qed
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    24
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    25
lemma minus_fold_remove:
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    26
  assumes "finite A"
37023
efc202e1677e added theory More_List
haftmann
parents: 34977
diff changeset
    27
  shows "B - A = Finite_Set.fold remove B A"
31807
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    28
proof -
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 39198
diff changeset
    29
  have rem: "remove = (\<lambda>x A. A - {x})" by (simp add: fun_eq_iff remove_def)
31807
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    30
  show ?thesis by (simp only: rem assms minus_fold_remove)
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    31
qed
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    32
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    33
definition project :: "('a \<Rightarrow> bool) \<Rightarrow> 'a set \<Rightarrow> 'a set" where
45012
060f76635bfe tuned specification and lemma distribution among theories; tuned proofs
haftmann
parents: 44326
diff changeset
    34
  "project P A = {a \<in> A. P a}"
060f76635bfe tuned specification and lemma distribution among theories; tuned proofs
haftmann
parents: 44326
diff changeset
    35
060f76635bfe tuned specification and lemma distribution among theories; tuned proofs
haftmann
parents: 44326
diff changeset
    36
lemma bounded_Collect_code [code_unfold_post]:
060f76635bfe tuned specification and lemma distribution among theories; tuned proofs
haftmann
parents: 44326
diff changeset
    37
  "{x \<in> A. P x} = project P A"
060f76635bfe tuned specification and lemma distribution among theories; tuned proofs
haftmann
parents: 44326
diff changeset
    38
  by (simp add: project_def)
060f76635bfe tuned specification and lemma distribution among theories; tuned proofs
haftmann
parents: 44326
diff changeset
    39
060f76635bfe tuned specification and lemma distribution among theories; tuned proofs
haftmann
parents: 44326
diff changeset
    40
definition product :: "'a set \<Rightarrow> 'b set \<Rightarrow> ('a \<times> 'b) set" where
060f76635bfe tuned specification and lemma distribution among theories; tuned proofs
haftmann
parents: 44326
diff changeset
    41
  "product A B = Sigma A (\<lambda>_. B)"
060f76635bfe tuned specification and lemma distribution among theories; tuned proofs
haftmann
parents: 44326
diff changeset
    42
060f76635bfe tuned specification and lemma distribution among theories; tuned proofs
haftmann
parents: 44326
diff changeset
    43
hide_const (open) product
060f76635bfe tuned specification and lemma distribution among theories; tuned proofs
haftmann
parents: 44326
diff changeset
    44
060f76635bfe tuned specification and lemma distribution among theories; tuned proofs
haftmann
parents: 44326
diff changeset
    45
lemma [code_unfold_post]:
060f76635bfe tuned specification and lemma distribution among theories; tuned proofs
haftmann
parents: 44326
diff changeset
    46
  "Sigma A (\<lambda>_. B) = More_Set.product A B"
060f76635bfe tuned specification and lemma distribution among theories; tuned proofs
haftmann
parents: 44326
diff changeset
    47
  by (simp add: product_def)
31807
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    48
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    49
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    50
subsection {* Basic set operations *}
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    51
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    52
lemma is_empty_set:
37595
9591362629e3 dropped ancient infix mem; refined code generation operations in List.thy
haftmann
parents: 37024
diff changeset
    53
  "is_empty (set xs) \<longleftrightarrow> List.null xs"
9591362629e3 dropped ancient infix mem; refined code generation operations in List.thy
haftmann
parents: 37024
diff changeset
    54
  by (simp add: is_empty_def null_def)
31807
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    55
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    56
lemma empty_set:
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    57
  "{} = set []"
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    58
  by simp
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    59
32880
b8bee63c7202 sets and cosets
haftmann
parents: 31851
diff changeset
    60
lemma insert_set_compl:
34977
27ceb64d41ea dropped some redundancies
haftmann
parents: 34007
diff changeset
    61
  "insert x (- set xs) = - set (removeAll x xs)"
27ceb64d41ea dropped some redundancies
haftmann
parents: 34007
diff changeset
    62
  by auto
31807
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    63
32880
b8bee63c7202 sets and cosets
haftmann
parents: 31851
diff changeset
    64
lemma remove_set_compl:
34977
27ceb64d41ea dropped some redundancies
haftmann
parents: 34007
diff changeset
    65
  "remove x (- set xs) = - set (List.insert x xs)"
44326
2b088d74beb3 tuned proof
haftmann
parents: 42871
diff changeset
    66
  by (auto simp add: remove_def List.insert_def)
32880
b8bee63c7202 sets and cosets
haftmann
parents: 31851
diff changeset
    67
31807
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    68
lemma image_set:
31846
89c37daebfdd added Inter, Union
haftmann
parents: 31807
diff changeset
    69
  "image f (set xs) = set (map f xs)"
31807
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    70
  by simp
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    71
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    72
lemma project_set:
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    73
  "project P (set xs) = set (filter P xs)"
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    74
  by (auto simp add: project_def)
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    75
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    76
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    77
subsection {* Functorial set operations *}
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    78
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    79
lemma union_set:
37023
efc202e1677e added theory More_List
haftmann
parents: 34977
diff changeset
    80
  "set xs \<union> A = fold Set.insert xs A"
31807
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    81
proof -
42871
1c0b99f950d9 names of fold_set locales resemble name of characteristic property more closely
haftmann
parents: 42868
diff changeset
    82
  interpret comp_fun_idem Set.insert
1c0b99f950d9 names of fold_set locales resemble name of characteristic property more closely
haftmann
parents: 42868
diff changeset
    83
    by (fact comp_fun_idem_insert)
31807
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    84
  show ?thesis by (simp add: union_fold_insert fold_set)
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    85
qed
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    86
37023
efc202e1677e added theory More_List
haftmann
parents: 34977
diff changeset
    87
lemma union_set_foldr:
efc202e1677e added theory More_List
haftmann
parents: 34977
diff changeset
    88
  "set xs \<union> A = foldr Set.insert xs A"
efc202e1677e added theory More_List
haftmann
parents: 34977
diff changeset
    89
proof -
efc202e1677e added theory More_List
haftmann
parents: 34977
diff changeset
    90
  have "\<And>x y :: 'a. insert y \<circ> insert x = insert x \<circ> insert y"
45012
060f76635bfe tuned specification and lemma distribution among theories; tuned proofs
haftmann
parents: 44326
diff changeset
    91
    by auto
37023
efc202e1677e added theory More_List
haftmann
parents: 34977
diff changeset
    92
  then show ?thesis by (simp add: union_set foldr_fold)
efc202e1677e added theory More_List
haftmann
parents: 34977
diff changeset
    93
qed
efc202e1677e added theory More_List
haftmann
parents: 34977
diff changeset
    94
31807
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    95
lemma minus_set:
37023
efc202e1677e added theory More_List
haftmann
parents: 34977
diff changeset
    96
  "A - set xs = fold remove xs A"
31807
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    97
proof -
42871
1c0b99f950d9 names of fold_set locales resemble name of characteristic property more closely
haftmann
parents: 42868
diff changeset
    98
  interpret comp_fun_idem remove
1c0b99f950d9 names of fold_set locales resemble name of characteristic property more closely
haftmann
parents: 42868
diff changeset
    99
    by (fact comp_fun_idem_remove)
31807
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   100
  show ?thesis
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   101
    by (simp add: minus_fold_remove [of _ A] fold_set)
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   102
qed
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   103
37023
efc202e1677e added theory More_List
haftmann
parents: 34977
diff changeset
   104
lemma minus_set_foldr:
efc202e1677e added theory More_List
haftmann
parents: 34977
diff changeset
   105
  "A - set xs = foldr remove xs A"
efc202e1677e added theory More_List
haftmann
parents: 34977
diff changeset
   106
proof -
efc202e1677e added theory More_List
haftmann
parents: 34977
diff changeset
   107
  have "\<And>x y :: 'a. remove y \<circ> remove x = remove x \<circ> remove y"
45012
060f76635bfe tuned specification and lemma distribution among theories; tuned proofs
haftmann
parents: 44326
diff changeset
   108
    by (auto simp add: remove_def)
37023
efc202e1677e added theory More_List
haftmann
parents: 34977
diff changeset
   109
  then show ?thesis by (simp add: minus_set foldr_fold)
efc202e1677e added theory More_List
haftmann
parents: 34977
diff changeset
   110
qed
efc202e1677e added theory More_List
haftmann
parents: 34977
diff changeset
   111
31807
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   112
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   113
subsection {* Derived set operations *}
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   114
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   115
lemma member:
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   116
  "a \<in> A \<longleftrightarrow> (\<exists>x\<in>A. a = x)"
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   117
  by simp
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   118
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   119
lemma subset_eq:
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   120
  "A \<subseteq> B \<longleftrightarrow> (\<forall>x\<in>A. x \<in> B)"
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   121
  by (fact subset_eq)
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   122
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   123
lemma subset:
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   124
  "A \<subset> B \<longleftrightarrow> A \<subseteq> B \<and> \<not> B \<subseteq> A"
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   125
  by (fact less_le_not_le)
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   126
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   127
lemma set_eq:
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   128
  "A = B \<longleftrightarrow> A \<subseteq> B \<and> B \<subseteq> A"
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   129
  by (fact eq_iff)
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   130
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   131
lemma inter:
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   132
  "A \<inter> B = project (\<lambda>x. x \<in> A) B"
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   133
  by (auto simp add: project_def)
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   134
37023
efc202e1677e added theory More_List
haftmann
parents: 34977
diff changeset
   135
45012
060f76635bfe tuned specification and lemma distribution among theories; tuned proofs
haftmann
parents: 44326
diff changeset
   136
subsection {* Theorems on relations *}
060f76635bfe tuned specification and lemma distribution among theories; tuned proofs
haftmann
parents: 44326
diff changeset
   137
060f76635bfe tuned specification and lemma distribution among theories; tuned proofs
haftmann
parents: 44326
diff changeset
   138
lemma product_code:
060f76635bfe tuned specification and lemma distribution among theories; tuned proofs
haftmann
parents: 44326
diff changeset
   139
  "More_Set.product (set xs) (set ys) = set [(x, y). x \<leftarrow> xs, y \<leftarrow> ys]"
060f76635bfe tuned specification and lemma distribution among theories; tuned proofs
haftmann
parents: 44326
diff changeset
   140
  by (auto simp add: product_def)
37023
efc202e1677e added theory More_List
haftmann
parents: 34977
diff changeset
   141
45012
060f76635bfe tuned specification and lemma distribution among theories; tuned proofs
haftmann
parents: 44326
diff changeset
   142
lemma Id_on_set:
060f76635bfe tuned specification and lemma distribution among theories; tuned proofs
haftmann
parents: 44326
diff changeset
   143
  "Id_on (set xs) = set [(x, x). x \<leftarrow> xs]"
060f76635bfe tuned specification and lemma distribution among theories; tuned proofs
haftmann
parents: 44326
diff changeset
   144
  by (auto simp add: Id_on_def)
060f76635bfe tuned specification and lemma distribution among theories; tuned proofs
haftmann
parents: 44326
diff changeset
   145
060f76635bfe tuned specification and lemma distribution among theories; tuned proofs
haftmann
parents: 44326
diff changeset
   146
lemma set_rel_comp:
060f76635bfe tuned specification and lemma distribution among theories; tuned proofs
haftmann
parents: 44326
diff changeset
   147
  "set xys O set yzs = set ([(fst xy, snd yz). xy \<leftarrow> xys, yz \<leftarrow> yzs, snd xy = fst yz])"
060f76635bfe tuned specification and lemma distribution among theories; tuned proofs
haftmann
parents: 44326
diff changeset
   148
  by (auto simp add: Bex_def)
060f76635bfe tuned specification and lemma distribution among theories; tuned proofs
haftmann
parents: 44326
diff changeset
   149
060f76635bfe tuned specification and lemma distribution among theories; tuned proofs
haftmann
parents: 44326
diff changeset
   150
lemma wf_set:
060f76635bfe tuned specification and lemma distribution among theories; tuned proofs
haftmann
parents: 44326
diff changeset
   151
  "wf (set xs) = acyclic (set xs)"
060f76635bfe tuned specification and lemma distribution among theories; tuned proofs
haftmann
parents: 44326
diff changeset
   152
  by (simp add: wf_iff_acyclic_if_finite)
37023
efc202e1677e added theory More_List
haftmann
parents: 34977
diff changeset
   153
37024
e938a0b5286e renamed List_Set to the now more appropriate More_Set
haftmann
parents: 37023
diff changeset
   154
end