src/HOL/Library/List_Set.thy
author haftmann
Thu, 25 Jun 2009 17:07:18 +0200
changeset 31807 039893a9a77d
child 31846 89c37daebfdd
permissions -rw-r--r--
added List_Set and Code_Set theories
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
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
     6
theory List_Set
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
     7
imports Main
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 list 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 insert :: "'a \<Rightarrow> 'a list \<Rightarrow> 'a list" where
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    13
  "insert x xs = (if x \<in> set xs then xs else x # xs)"
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_all :: "'a \<Rightarrow> 'a list \<Rightarrow> 'a list" where
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    16
  "remove_all x xs = filter (Not o op = x) xs"
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
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    19
subsection {* Various additional set functions *}
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    20
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    21
definition is_empty :: "'a set \<Rightarrow> bool" where
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    22
  "is_empty A \<longleftrightarrow> A = {}"
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    23
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    24
definition remove :: "'a \<Rightarrow> 'a set \<Rightarrow> 'a set" where
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    25
  "remove x A = A - {x}"
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    26
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    27
lemma fun_left_comm_idem_remove:
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    28
  "fun_left_comm_idem remove"
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    29
proof -
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    30
  have rem: "remove = (\<lambda>x A. A - {x})" by (simp add: expand_fun_eq remove_def)
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    31
  show ?thesis by (simp only: fun_left_comm_idem_remove rem)
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    32
qed
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    33
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    34
lemma minus_fold_remove:
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    35
  assumes "finite A"
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    36
  shows "B - A = fold remove B A"
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    37
proof -
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    38
  have rem: "remove = (\<lambda>x A. A - {x})" by (simp add: expand_fun_eq remove_def)
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    39
  show ?thesis by (simp only: rem assms minus_fold_remove)
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    40
qed
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    41
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    42
definition project :: "('a \<Rightarrow> bool) \<Rightarrow> 'a set \<Rightarrow> 'a set" where
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    43
  "project P A = {a\<in>A. P a}"
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    44
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    45
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    46
subsection {* Basic set operations *}
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    47
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    48
lemma is_empty_set:
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    49
  "is_empty (set xs) \<longleftrightarrow> null xs"
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    50
  by (simp add: is_empty_def null_empty)
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 ball_set:
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    53
  "(\<forall>x\<in>set xs. P x) \<longleftrightarrow> list_all P xs"
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    54
  by (rule list_ball_code)
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 bex_set:
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    57
  "(\<exists>x\<in>set xs. P x) \<longleftrightarrow> list_ex P xs"
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    58
  by (rule list_bex_code)
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    59
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    60
lemma empty_set:
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    61
  "{} = set []"
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    62
  by simp
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
lemma insert_set:
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    65
  "Set.insert x (set xs) = set (insert x xs)"
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    66
  by (auto simp add: insert_def)
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    67
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    68
lemma remove_set:
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    69
  "remove x (set xs) = set (remove_all x xs)"
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    70
  by (auto simp add: remove_def remove_all_def)
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 image_set:
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    73
  "image f (set xs) = set (remdups (map f xs))"
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    74
  by simp
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
lemma project_set:
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    77
  "project P (set xs) = set (filter P xs)"
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    78
  by (auto simp add: project_def)
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    79
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    80
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    81
subsection {* Functorial set operations *}
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    82
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    83
lemma union_set:
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    84
  "set xs \<union> A = foldl (\<lambda>A x. Set.insert x A) A xs"
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    85
proof -
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    86
  interpret fun_left_comm_idem Set.insert
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    87
    by (fact fun_left_comm_idem_insert)
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    88
  show ?thesis by (simp add: union_fold_insert 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
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    91
lemma minus_set:
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    92
  "A - set xs = foldl (\<lambda>A x. remove x A) A xs"
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    93
proof -
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    94
  interpret fun_left_comm_idem remove
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    95
    by (fact fun_left_comm_idem_remove)
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    96
  show ?thesis
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    97
    by (simp add: minus_fold_remove [of _ A] fold_set)
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    98
qed
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
lemma Inter_set:
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   101
  "Inter (set (A # As)) = foldl (op \<inter>) A As"
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   102
proof -
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   103
  have "finite (set (A # As))" by simp
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   104
  moreover have "fold (op \<inter>) UNIV (set (A # As)) = foldl (\<lambda>y x. x \<inter> y) UNIV (A # As)"
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   105
    by (rule fun_left_comm_idem.fold_set, unfold_locales, auto)
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   106
  ultimately have "Inter (set (A # As)) = foldl (op \<inter>) UNIV (A # As)"
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   107
    by (simp only: Inter_fold_inter Int_commute)
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   108
  then show ?thesis by simp
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   109
qed
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   110
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   111
lemma Union_set:
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   112
  "Union (set As) = foldl (op \<union>) {} As"
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   113
proof -
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   114
  have "fold (op \<union>) {} (set As) = foldl (\<lambda>y x. x \<union> y) {} As"
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   115
    by (rule fun_left_comm_idem.fold_set, unfold_locales, auto)
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   116
  then show ?thesis
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   117
    by (simp only: Union_fold_union finite_set Un_commute)
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   118
qed
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   119
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   120
lemma INTER_set:
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   121
  "INTER (set (A # As)) f = foldl (\<lambda>B A. f A \<inter> B) (f A) As"
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   122
proof -
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   123
  have "finite (set (A # As))" by simp
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   124
  moreover have "fold (\<lambda>A. op \<inter> (f A)) UNIV (set (A # As)) = foldl (\<lambda>B A. f A \<inter> B) UNIV (A # As)"
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   125
    by (rule fun_left_comm_idem.fold_set, unfold_locales, auto)
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   126
  ultimately have "INTER (set (A # As)) f = foldl (\<lambda>B A. f A \<inter> B) UNIV (A # As)"
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   127
    by (simp only: INTER_fold_inter) 
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   128
  then show ?thesis by simp
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   129
qed
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 UNION_set:
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   132
  "UNION (set As) f = foldl (\<lambda>B A. f A \<union> B) {} As"
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   133
proof -
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   134
  have "fold (\<lambda>A. op \<union> (f A)) {} (set As) = foldl (\<lambda>B A. f A \<union> B) {} As"
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   135
    by (rule fun_left_comm_idem.fold_set, unfold_locales, auto)
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   136
  then show ?thesis
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   137
    by (simp only: UNION_fold_union finite_set)
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   138
qed
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   139
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   140
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   141
subsection {* Derived set operations *}
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   142
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   143
lemma member:
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   144
  "a \<in> A \<longleftrightarrow> (\<exists>x\<in>A. a = x)"
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   145
  by simp
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   146
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   147
lemma subset_eq:
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   148
  "A \<subseteq> B \<longleftrightarrow> (\<forall>x\<in>A. x \<in> B)"
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   149
  by (fact subset_eq)
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   150
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   151
lemma subset:
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   152
  "A \<subset> B \<longleftrightarrow> A \<subseteq> B \<and> \<not> B \<subseteq> A"
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   153
  by (fact less_le_not_le)
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   154
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   155
lemma set_eq:
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   156
  "A = B \<longleftrightarrow> A \<subseteq> B \<and> B \<subseteq> A"
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   157
  by (fact eq_iff)
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   158
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   159
lemma inter:
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   160
  "A \<inter> B = project (\<lambda>x. x \<in> A) B"
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   161
  by (auto simp add: project_def)
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   162
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   163
end