src/HOL/Library/More_Set.thy
author bulwahn
Mon, 22 Nov 2010 11:34:56 +0100
changeset 40650 d40b347d5b0b
parent 39302 d7728f65b353
child 42868 1608daf27c1f
permissions -rw-r--r--
adding Enum to HOL-Main image and removing it from HOL-Library
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
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    18
lemma fun_left_comm_idem_remove:
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    19
  "fun_left_comm_idem remove"
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)
31807
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    22
  show ?thesis by (simp only: fun_left_comm_idem_remove rem)
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
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    34
  "project P A = {a\<in>A. P a}"
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    35
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    36
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    37
subsection {* Basic set operations *}
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    38
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    39
lemma is_empty_set:
37595
9591362629e3 dropped ancient infix mem; refined code generation operations in List.thy
haftmann
parents: 37024
diff changeset
    40
  "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
    41
  by (simp add: is_empty_def null_def)
31807
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    42
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    43
lemma empty_set:
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    44
  "{} = set []"
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    45
  by simp
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    46
32880
b8bee63c7202 sets and cosets
haftmann
parents: 31851
diff changeset
    47
lemma insert_set_compl:
34977
27ceb64d41ea dropped some redundancies
haftmann
parents: 34007
diff changeset
    48
  "insert x (- set xs) = - set (removeAll x xs)"
27ceb64d41ea dropped some redundancies
haftmann
parents: 34007
diff changeset
    49
  by auto
31807
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    50
32880
b8bee63c7202 sets and cosets
haftmann
parents: 31851
diff changeset
    51
lemma remove_set_compl:
34977
27ceb64d41ea dropped some redundancies
haftmann
parents: 34007
diff changeset
    52
  "remove x (- set xs) = - set (List.insert x xs)"
27ceb64d41ea dropped some redundancies
haftmann
parents: 34007
diff changeset
    53
  by (auto simp del: mem_def simp add: remove_def List.insert_def)
32880
b8bee63c7202 sets and cosets
haftmann
parents: 31851
diff changeset
    54
31807
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    55
lemma image_set:
31846
89c37daebfdd added Inter, Union
haftmann
parents: 31807
diff changeset
    56
  "image f (set xs) = set (map f xs)"
31807
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    57
  by simp
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    58
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    59
lemma project_set:
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    60
  "project P (set xs) = set (filter P xs)"
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    61
  by (auto simp add: project_def)
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    62
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    63
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    64
subsection {* Functorial set operations *}
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    65
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    66
lemma union_set:
37023
efc202e1677e added theory More_List
haftmann
parents: 34977
diff changeset
    67
  "set xs \<union> A = fold Set.insert xs A"
31807
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    68
proof -
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    69
  interpret fun_left_comm_idem Set.insert
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    70
    by (fact fun_left_comm_idem_insert)
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    71
  show ?thesis by (simp add: union_fold_insert fold_set)
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    72
qed
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    73
37023
efc202e1677e added theory More_List
haftmann
parents: 34977
diff changeset
    74
lemma union_set_foldr:
efc202e1677e added theory More_List
haftmann
parents: 34977
diff changeset
    75
  "set xs \<union> A = foldr Set.insert xs A"
efc202e1677e added theory More_List
haftmann
parents: 34977
diff changeset
    76
proof -
efc202e1677e added theory More_List
haftmann
parents: 34977
diff changeset
    77
  have "\<And>x y :: 'a. insert y \<circ> insert x = insert x \<circ> insert y"
efc202e1677e added theory More_List
haftmann
parents: 34977
diff changeset
    78
    by (auto intro: ext)
efc202e1677e added theory More_List
haftmann
parents: 34977
diff changeset
    79
  then show ?thesis by (simp add: union_set foldr_fold)
efc202e1677e added theory More_List
haftmann
parents: 34977
diff changeset
    80
qed
efc202e1677e added theory More_List
haftmann
parents: 34977
diff changeset
    81
31807
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    82
lemma minus_set:
37023
efc202e1677e added theory More_List
haftmann
parents: 34977
diff changeset
    83
  "A - set xs = fold remove xs A"
31807
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    84
proof -
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    85
  interpret fun_left_comm_idem remove
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    86
    by (fact fun_left_comm_idem_remove)
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    87
  show ?thesis
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    88
    by (simp add: minus_fold_remove [of _ A] fold_set)
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    89
qed
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    90
37023
efc202e1677e added theory More_List
haftmann
parents: 34977
diff changeset
    91
lemma minus_set_foldr:
efc202e1677e added theory More_List
haftmann
parents: 34977
diff changeset
    92
  "A - set xs = foldr remove xs A"
efc202e1677e added theory More_List
haftmann
parents: 34977
diff changeset
    93
proof -
efc202e1677e added theory More_List
haftmann
parents: 34977
diff changeset
    94
  have "\<And>x y :: 'a. remove y \<circ> remove x = remove x \<circ> remove y"
efc202e1677e added theory More_List
haftmann
parents: 34977
diff changeset
    95
    by (auto simp add: remove_def intro: ext)
efc202e1677e added theory More_List
haftmann
parents: 34977
diff changeset
    96
  then show ?thesis by (simp add: minus_set foldr_fold)
efc202e1677e added theory More_List
haftmann
parents: 34977
diff changeset
    97
qed
efc202e1677e added theory More_List
haftmann
parents: 34977
diff changeset
    98
31807
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    99
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   100
subsection {* Derived set operations *}
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   101
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   102
lemma member:
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   103
  "a \<in> A \<longleftrightarrow> (\<exists>x\<in>A. a = x)"
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   104
  by simp
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   105
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   106
lemma subset_eq:
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   107
  "A \<subseteq> B \<longleftrightarrow> (\<forall>x\<in>A. x \<in> B)"
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   108
  by (fact subset_eq)
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   109
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   110
lemma subset:
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   111
  "A \<subset> B \<longleftrightarrow> A \<subseteq> B \<and> \<not> B \<subseteq> A"
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   112
  by (fact less_le_not_le)
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   113
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   114
lemma set_eq:
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   115
  "A = B \<longleftrightarrow> A \<subseteq> B \<and> B \<subseteq> A"
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   116
  by (fact eq_iff)
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   117
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   118
lemma inter:
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   119
  "A \<inter> B = project (\<lambda>x. x \<in> A) B"
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   120
  by (auto simp add: project_def)
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   121
37023
efc202e1677e added theory More_List
haftmann
parents: 34977
diff changeset
   122
efc202e1677e added theory More_List
haftmann
parents: 34977
diff changeset
   123
subsection {* Various lemmas *}
efc202e1677e added theory More_List
haftmann
parents: 34977
diff changeset
   124
efc202e1677e added theory More_List
haftmann
parents: 34977
diff changeset
   125
lemma not_set_compl:
efc202e1677e added theory More_List
haftmann
parents: 34977
diff changeset
   126
  "Not \<circ> set xs = - set xs"
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 39198
diff changeset
   127
  by (simp add: fun_Compl_def bool_Compl_def comp_def fun_eq_iff)
37023
efc202e1677e added theory More_List
haftmann
parents: 34977
diff changeset
   128
37024
e938a0b5286e renamed List_Set to the now more appropriate More_Set
haftmann
parents: 37023
diff changeset
   129
end