src/HOL/Library/Executable_Set.thy
author wenzelm
Tue, 31 Jul 2007 13:30:35 +0200
changeset 24085 cbad32e7ab40
parent 23854 688a8a7bcd4e
child 24423 ae9cd0e92423
permissions -rw-r--r--
added configuration options;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
23854
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
     1
(*  Title:      HOL/Library/Executable_Set.thy
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
     2
    ID:         $Id$
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
     3
    Author:     Stefan Berghofer, TU Muenchen
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
     4
*)
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
     5
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
     6
header {* Implementation of finite sets by lists *}
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
     7
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
     8
theory Executable_Set
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
     9
imports Main
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    10
begin
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    11
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    12
subsection {* Definitional rewrites *}
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    13
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    14
instance set :: (eq) eq ..
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    15
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    16
lemma [code target: Set]:
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    17
  "A = B \<longleftrightarrow> A \<subseteq> B \<and> B \<subseteq> A"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    18
  by blast
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    19
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    20
lemma [code func]:
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    21
  "(A\<Colon>'a\<Colon>eq set) = B \<longleftrightarrow> A \<subseteq> B \<and> B \<subseteq> A"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    22
  by blast
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    23
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    24
lemma [code func]:
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    25
  "(A\<Colon>'a\<Colon>eq set) \<subseteq> B \<longleftrightarrow> (\<forall>x\<in>A. x \<in> B)"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    26
  unfolding subset_def ..
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    27
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    28
lemma [code func]:
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    29
  "(A\<Colon>'a\<Colon>eq set) \<subset> B \<longleftrightarrow> A \<subseteq> B \<and> A \<noteq> B"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    30
  by blast
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    31
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    32
lemma [code]:
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    33
  "a \<in> A \<longleftrightarrow> (\<exists>x\<in>A. x = a)"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    34
  unfolding bex_triv_one_point1 ..
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    35
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    36
definition
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    37
  filter_set :: "('a \<Rightarrow> bool) \<Rightarrow> 'a set \<Rightarrow> 'a set" where
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    38
  "filter_set P xs = {x\<in>xs. P x}"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    39
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    40
lemmas [symmetric, code inline] = filter_set_def
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    41
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    42
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    43
subsection {* Operations on lists *}
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    44
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    45
subsubsection {* Basic definitions *}
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    46
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    47
definition
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    48
  flip :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> 'b \<Rightarrow> 'a \<Rightarrow> 'c" where
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    49
  "flip f a b = f b a"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    50
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    51
definition
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    52
  member :: "'a list \<Rightarrow> 'a \<Rightarrow> bool" where
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    53
  "member xs x \<longleftrightarrow> x \<in> set xs"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    54
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    55
definition
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    56
  insertl :: "'a \<Rightarrow> 'a list \<Rightarrow> 'a list" where
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    57
  "insertl x xs = (if member xs x then xs else x#xs)"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    58
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    59
lemma [code target: List]: "member [] y \<longleftrightarrow> False"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    60
  and [code target: List]: "member (x#xs) y \<longleftrightarrow> y = x \<or> member xs y"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    61
  unfolding member_def by (induct xs) simp_all
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    62
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    63
fun
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    64
  drop_first :: "('a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> 'a list" where
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    65
  "drop_first f [] = []"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    66
| "drop_first f (x#xs) = (if f x then xs else x # drop_first f xs)"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    67
declare drop_first.simps [code del]
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    68
declare drop_first.simps [code target: List]
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    69
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    70
declare remove1.simps [code del]
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    71
lemma [code target: List]:
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    72
  "remove1 x xs = (if member xs x then drop_first (\<lambda>y. y = x) xs else xs)"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    73
proof (cases "member xs x")
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    74
  case False thus ?thesis unfolding member_def by (induct xs) auto
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    75
next
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    76
  case True
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    77
  have "remove1 x xs = drop_first (\<lambda>y. y = x) xs" by (induct xs) simp_all
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    78
  with True show ?thesis by simp
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    79
qed
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    80
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    81
lemma member_nil [simp]:
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    82
  "member [] = (\<lambda>x. False)"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    83
proof
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    84
  fix x
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    85
  show "member [] x = False" unfolding member_def by simp
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    86
qed
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    87
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    88
lemma member_insertl [simp]:
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    89
  "x \<in> set (insertl x xs)"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    90
  unfolding insertl_def member_def mem_iff by simp
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    91
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    92
lemma insertl_member [simp]:
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    93
  fixes xs x
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    94
  assumes member: "member xs x"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    95
  shows "insertl x xs = xs"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    96
  using member unfolding insertl_def by simp
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    97
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    98
lemma insertl_not_member [simp]:
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    99
  fixes xs x
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   100
  assumes member: "\<not> (member xs x)"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   101
  shows "insertl x xs = x # xs"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   102
  using member unfolding insertl_def by simp
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   103
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   104
lemma foldr_remove1_empty [simp]:
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   105
  "foldr remove1 xs [] = []"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   106
  by (induct xs) simp_all
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   107
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   108
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   109
subsubsection {* Derived definitions *}
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   110
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   111
function unionl :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   112
where
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   113
  "unionl [] ys = ys"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   114
| "unionl xs ys = foldr insertl xs ys"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   115
by pat_completeness auto
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   116
termination by lexicographic_order
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   117
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   118
lemmas unionl_def = unionl.simps(2)
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   119
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   120
function intersect :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   121
where
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   122
  "intersect [] ys = []"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   123
| "intersect xs [] = []"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   124
| "intersect xs ys = filter (member xs) ys"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   125
by pat_completeness auto
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   126
termination by lexicographic_order
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   127
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   128
lemmas intersect_def = intersect.simps(3)
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   129
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   130
function subtract :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   131
where
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   132
  "subtract [] ys = ys"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   133
| "subtract xs [] = []"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   134
| "subtract xs ys = foldr remove1 xs ys"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   135
by pat_completeness auto
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   136
termination by lexicographic_order
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   137
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   138
lemmas subtract_def = subtract.simps(3)
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   139
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   140
function map_distinct :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a list \<Rightarrow> 'b list"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   141
where
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   142
  "map_distinct f [] = []"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   143
| "map_distinct f xs = foldr (insertl o f) xs []"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   144
by pat_completeness auto
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   145
termination by lexicographic_order
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   146
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   147
lemmas map_distinct_def = map_distinct.simps(2)
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   148
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   149
function unions :: "'a list list \<Rightarrow> 'a list"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   150
where
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   151
  "unions [] = []"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   152
| "unions xs = foldr unionl xs []"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   153
by pat_completeness auto
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   154
termination by lexicographic_order
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   155
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   156
lemmas unions_def = unions.simps(2)
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   157
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   158
consts intersects :: "'a list list \<Rightarrow> 'a list"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   159
primrec
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   160
  "intersects (x#xs) = foldr intersect xs x"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   161
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   162
definition
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   163
  map_union :: "'a list \<Rightarrow> ('a \<Rightarrow> 'b list) \<Rightarrow> 'b list" where
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   164
  "map_union xs f = unions (map f xs)"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   165
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   166
definition
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   167
  map_inter :: "'a list \<Rightarrow> ('a \<Rightarrow> 'b list) \<Rightarrow> 'b list" where
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   168
  "map_inter xs f = intersects (map f xs)"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   169
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   170
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   171
subsection {* Isomorphism proofs *}
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   172
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   173
lemma iso_member:
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   174
  "member xs x \<longleftrightarrow> x \<in> set xs"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   175
  unfolding member_def mem_iff ..
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   176
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   177
lemma iso_insert:
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   178
  "set (insertl x xs) = insert x (set xs)"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   179
  unfolding insertl_def iso_member by (simp add: Set.insert_absorb)
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   180
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   181
lemma iso_remove1:
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   182
  assumes distnct: "distinct xs"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   183
  shows "set (remove1 x xs) = set xs - {x}"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   184
  using distnct set_remove1_eq by auto
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   185
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   186
lemma iso_union:
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   187
  "set (unionl xs ys) = set xs \<union> set ys"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   188
  unfolding unionl_def
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   189
  by (induct xs arbitrary: ys) (simp_all add: iso_insert)
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   190
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   191
lemma iso_intersect:
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   192
  "set (intersect xs ys) = set xs \<inter> set ys"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   193
  unfolding intersect_def Int_def by (simp add: Int_def iso_member) auto
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   194
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   195
definition
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   196
  subtract' :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" where
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   197
  "subtract' = flip subtract"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   198
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   199
lemma iso_subtract:
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   200
  fixes ys
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   201
  assumes distnct: "distinct ys"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   202
  shows "set (subtract' ys xs) = set ys - set xs"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   203
    and "distinct (subtract' ys xs)"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   204
  unfolding subtract'_def flip_def subtract_def
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   205
  using distnct by (induct xs arbitrary: ys) auto
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   206
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   207
lemma iso_map_distinct:
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   208
  "set (map_distinct f xs) = image f (set xs)"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   209
  unfolding map_distinct_def by (induct xs) (simp_all add: iso_insert)
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   210
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   211
lemma iso_unions:
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   212
  "set (unions xss) = \<Union> set (map set xss)"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   213
  unfolding unions_def
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   214
proof (induct xss)
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   215
  case Nil show ?case by simp
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   216
next
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   217
  case (Cons xs xss) thus ?case by (induct xs) (simp_all add: iso_insert)
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   218
qed
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   219
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   220
lemma iso_intersects:
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   221
  "set (intersects (xs#xss)) = \<Inter> set (map set (xs#xss))"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   222
  by (induct xss) (simp_all add: Int_def iso_member, auto)
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   223
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   224
lemma iso_UNION:
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   225
  "set (map_union xs f) = UNION (set xs) (set o f)"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   226
  unfolding map_union_def iso_unions by simp
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   227
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   228
lemma iso_INTER:
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   229
  "set (map_inter (x#xs) f) = INTER (set (x#xs)) (set o f)"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   230
  unfolding map_inter_def iso_intersects by (induct xs) (simp_all add: iso_member, auto)
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   231
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   232
definition
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   233
  Blall :: "'a list \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> bool" where
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   234
  "Blall = flip list_all"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   235
definition
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   236
  Blex :: "'a list \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> bool" where
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   237
  "Blex = flip list_ex"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   238
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   239
lemma iso_Ball:
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   240
  "Blall xs f = Ball (set xs) f"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   241
  unfolding Blall_def flip_def by (induct xs) simp_all
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   242
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   243
lemma iso_Bex:
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   244
  "Blex xs f = Bex (set xs) f"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   245
  unfolding Blex_def flip_def by (induct xs) simp_all
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   246
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   247
lemma iso_filter:
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   248
  "set (filter P xs) = filter_set P (set xs)"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   249
  unfolding filter_set_def by (induct xs) auto
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   250
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   251
subsection {* code generator setup *}
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   252
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   253
ML {*
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   254
nonfix inter;
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   255
nonfix union;
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   256
nonfix subset;
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   257
*}
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   258
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   259
code_modulename SML
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   260
  Executable_Set List
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   261
  Set List
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   262
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   263
code_modulename OCaml
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   264
  Executable_Set List
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   265
  Set List
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   266
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   267
code_modulename Haskell
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   268
  Executable_Set List
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   269
  Set List
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   270
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   271
definition [code inline]:
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   272
  "empty_list = []"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   273
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   274
lemma [code func]:
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   275
  "insert (x \<Colon> 'a\<Colon>eq) = insert x" ..
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   276
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   277
lemma [code func]:
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   278
  "(xs \<Colon> 'a\<Colon>eq set) \<union> ys = xs \<union> ys" ..
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   279
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   280
lemma [code func]:
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   281
  "(xs \<Colon> 'a\<Colon>eq set) \<inter> ys = xs \<inter> ys" ..
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   282
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   283
lemma [code func]:
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   284
  "(op -) (xs \<Colon> 'a\<Colon>eq set) = (op -) (xs \<Colon> 'a\<Colon>eq set)" ..
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   285
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   286
lemma [code func]:
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   287
  "image (f \<Colon> 'a \<Rightarrow> 'b\<Colon>eq) = image f" ..
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   288
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   289
lemma [code func]:
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   290
  "Union (xss \<Colon> 'a\<Colon>eq set set) = Union xss" ..
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   291
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   292
lemma [code func]:
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   293
  "Inter (xss \<Colon> 'a\<Colon>eq set set) = Inter xss" ..
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   294
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   295
lemma [code func]:
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   296
  "UNION xs (f \<Colon> 'a \<Rightarrow> 'b\<Colon>eq set) = UNION xs f" ..
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   297
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   298
lemma [code func]:
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   299
  "INTER xs (f \<Colon> 'a \<Rightarrow> 'b\<Colon>eq set) = INTER xs f" ..
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   300
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   301
lemma [code func]:
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   302
  "Ball (xs \<Colon> 'a\<Colon>type set) = Ball xs" ..
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   303
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   304
lemma [code func]:
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   305
  "Bex (xs \<Colon> 'a\<Colon>type set) = Bex xs" ..
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   306
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   307
lemma [code func]:
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   308
  "filter_set P (xs \<Colon> 'a\<Colon>type set) = filter_set P xs" ..
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   309
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   310
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   311
code_abstype "'a set" "'a list" where
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   312
  "{}" \<equiv> empty_list
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   313
  insert \<equiv> insertl
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   314
  "op \<union>" \<equiv> unionl
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   315
  "op \<inter>" \<equiv> intersect
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   316
  "op - \<Colon> 'a set \<Rightarrow> 'a set \<Rightarrow> 'a set" \<equiv> subtract'
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   317
  image \<equiv> map_distinct
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   318
  Union \<equiv> unions
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   319
  Inter \<equiv> intersects
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   320
  UNION \<equiv> map_union
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   321
  INTER \<equiv> map_inter
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   322
  Ball \<equiv> Blall
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   323
  Bex \<equiv> Blex
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   324
  filter_set \<equiv> filter
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   325
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   326
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   327
subsubsection {* type serializations *}
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   328
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   329
types_code
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   330
  set ("_ list")
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   331
attach (term_of) {*
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   332
fun term_of_set f T [] = Const ("{}", Type ("set", [T]))
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   333
  | term_of_set f T (x :: xs) = Const ("insert",
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   334
      T --> Type ("set", [T]) --> Type ("set", [T])) $ f x $ term_of_set f T xs;
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   335
*}
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   336
attach (test) {*
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   337
fun gen_set' aG i j = frequency
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   338
  [(i, fn () => aG j :: gen_set' aG (i-1) j), (1, fn () => [])] ()
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   339
and gen_set aG i = gen_set' aG i i;
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   340
*}
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   341
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   342
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   343
subsubsection {* const serializations *}
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   344
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   345
consts_code
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   346
  "{}" ("{*[]*}")
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   347
  insert ("{*insertl*}")
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   348
  "op \<union>" ("{*unionl*}")
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   349
  "op \<inter>" ("{*intersect*}")
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   350
  "op - \<Colon> 'a set \<Rightarrow> 'a set \<Rightarrow> 'a set" ("{* flip subtract *}")
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   351
  image ("{*map_distinct*}")
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   352
  Union ("{*unions*}")
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   353
  Inter ("{*intersects*}")
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   354
  UNION ("{*map_union*}")
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   355
  INTER ("{*map_inter*}")
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   356
  Ball ("{*Blall*}")
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   357
  Bex ("{*Blex*}")
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   358
  filter_set ("{*filter*}")
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   359
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   360
end