src/HOL/Library/Code_Set.thy
author wenzelm
Sun, 28 Jun 2009 22:51:29 +0200
changeset 31845 cc7ddda02436
parent 31807 039893a9a77d
child 31846 89c37daebfdd
permissions -rw-r--r--
more precise dependencies for "jars"; allow spaces in SCALA_HOME;
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 {* Executable finite sets *}
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 Code_Set
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
     7
imports List_Set
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
lemma foldl_apply_inv:
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    11
  assumes "\<And>x. g (h x) = x"
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    12
  shows "foldl f (g s) xs = g (foldl (\<lambda>s x. h (f (g s) x)) s xs)"
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    13
  by (rule sym, induct xs arbitrary: s) (simp_all add: assms)
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
subsection {* Lifting *}
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    16
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    17
datatype 'a fset = Fset "'a set"
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
primrec member :: "'a fset \<Rightarrow> 'a set" where
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    20
  "member (Fset A) = A"
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    21
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    22
lemma Fset_member [simp]:
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    23
  "Fset (member A) = A"
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    24
  by (cases A) simp
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    25
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    26
definition Set :: "'a list \<Rightarrow> 'a fset" where
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    27
  "Set xs = Fset (set xs)"
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    28
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    29
lemma member_Set [simp]:
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    30
  "member (Set xs) = set xs"
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    31
  by (simp add: Set_def)
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
code_datatype Set
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    34
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
subsection {* Basic operations *}
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    37
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    38
definition is_empty :: "'a fset \<Rightarrow> bool" where
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    39
  "is_empty A \<longleftrightarrow> List_Set.is_empty (member A)"
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    40
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    41
lemma is_empty_Set [code]:
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    42
  "is_empty (Set xs) \<longleftrightarrow> null xs"
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    43
  by (simp add: is_empty_def is_empty_set)
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
definition empty :: "'a fset" where
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    46
  "empty = Fset {}"
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 empty_Set [code]:
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    49
  "empty = Set []"
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    50
  by (simp add: empty_def Set_def)
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
definition insert :: "'a \<Rightarrow> 'a fset \<Rightarrow> 'a fset" where
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    53
  "insert x A = Fset (Set.insert x (member A))"
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    54
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    55
lemma insert_Set [code]:
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    56
  "insert x (Set xs) = Set (List_Set.insert x xs)"
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    57
  by (simp add: insert_def Set_def insert_set)
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
definition remove :: "'a \<Rightarrow> 'a fset \<Rightarrow> 'a fset" where
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    60
  "remove x A = Fset (List_Set.remove x (member A))"
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    61
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    62
lemma remove_Set [code]:
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    63
  "remove x (Set xs) = Set (remove_all x xs)"
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    64
  by (simp add: remove_def Set_def remove_set)
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
definition map :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a fset \<Rightarrow> 'b fset" where
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    67
  "map f A = Fset (image f (member A))"
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    68
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    69
lemma map_Set [code]:
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    70
  "map f (Set xs) = Set (remdups (List.map f xs))"
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    71
  by (simp add: map_def Set_def)
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    72
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    73
definition project :: "('a \<Rightarrow> bool) \<Rightarrow> 'a fset \<Rightarrow> 'a fset" where
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    74
  "project P A = Fset (List_Set.project P (member A))"
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 [code]:
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 (simp add: project_def Set_def project_set)
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
definition forall :: "('a \<Rightarrow> bool) \<Rightarrow> 'a fset \<Rightarrow> bool" where
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    81
  "forall P A \<longleftrightarrow> Ball (member A) P"
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 forall_Set [code]:
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    84
  "forall P (Set xs) \<longleftrightarrow> list_all P xs"
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    85
  by (simp add: forall_def Set_def ball_set)
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    86
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    87
definition exists :: "('a \<Rightarrow> bool) \<Rightarrow> 'a fset \<Rightarrow> bool" where
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    88
  "exists P A \<longleftrightarrow> Bex (member A) P"
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    89
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    90
lemma exists_Set [code]:
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    91
  "exists P (Set xs) \<longleftrightarrow> list_ex P xs"
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    92
  by (simp add: exists_def Set_def bex_set)
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    93
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    94
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    95
subsection {* Functorial operations *}
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    96
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    97
definition union :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> 'a fset" where
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    98
  "union A B = Fset (member A \<union> member B)"
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 union_insert [code]:
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   101
  "union (Set xs) A = foldl (\<lambda>A x. insert x A) A xs"
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 "foldl (\<lambda>A x. Set.insert x A) (member A) xs =
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   104
    member (foldl (\<lambda>A x. Fset (Set.insert x (member A))) A xs)"
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   105
    by (rule foldl_apply_inv) simp
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   106
  then show ?thesis by (simp add: union_def union_set insert_def)
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   107
qed
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   108
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   109
definition subtract :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> 'a fset" where
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   110
  "subtract A B = Fset (member B - member A)"
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   111
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   112
lemma subtract_remove [code]:
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   113
  "subtract (Set xs) A = foldl (\<lambda>A x. remove x A) A xs"
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   114
proof -
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   115
  have "foldl (\<lambda>A x. List_Set.remove x A) (member A) xs =
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   116
    member (foldl (\<lambda>A x. Fset (List_Set.remove x (member A))) A xs)"
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   117
    by (rule foldl_apply_inv) simp
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   118
  then show ?thesis by (simp add: subtract_def minus_set remove_def)
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   119
qed
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   120
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   121
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   122
subsection {* Derived operations *}
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   123
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   124
lemma member_exists [code]:
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   125
  "member A y \<longleftrightarrow> exists (\<lambda>x. y = x) A"
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   126
  by (simp add: exists_def mem_def)
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   127
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   128
definition subfset_eq :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> bool" where
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   129
  "subfset_eq A B \<longleftrightarrow> member A \<subseteq> member B"
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 subfset_eq_forall [code]:
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   132
  "subfset_eq A B \<longleftrightarrow> forall (\<lambda>x. member B x) A"
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   133
  by (simp add: subfset_eq_def subset_eq forall_def mem_def)
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   134
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   135
definition subfset :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> bool" where
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   136
  "subfset A B \<longleftrightarrow> member A \<subset> member B"
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   137
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   138
lemma subfset_subfset_eq [code]:
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   139
  "subfset A B \<longleftrightarrow> subfset_eq A B \<and> \<not> subfset_eq B A"
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   140
  by (simp add: subfset_def subfset_eq_def subset)
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   141
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   142
lemma eq_fset_subfset_eq [code]:
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   143
  "eq_class.eq A B \<longleftrightarrow> subfset_eq A B \<and> subfset_eq B A"
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   144
  by (cases A, cases B) (simp add: eq subfset_eq_def set_eq)
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   145
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   146
definition inter :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> 'a fset" where
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   147
  "inter A B = Fset (List_Set.project (member A) (member B))"
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   148
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   149
lemma inter_project [code]:
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   150
  "inter A B = project (member A) B"
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   151
  by (simp add: inter_def project_def inter)
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   152
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   153
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   154
subsection {* Misc operations *}
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   155
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   156
lemma size_fset [code]:
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   157
  "fset_size f A = 0"
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   158
  "size A = 0"
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   159
  by (cases A, simp) (cases A, simp)
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   160
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   161
lemma fset_case_code [code]:
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   162
  "fset_case f A = f (member A)"
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   163
  by (cases A) simp
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   164
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   165
lemma fset_rec_code [code]:
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   166
  "fset_rec f A = f (member A)"
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   167
  by (cases A) simp
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   168
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   169
end