src/HOL/Library/Executable_Set.thy
author nipkow
Sun, 10 May 2009 14:21:41 +0200
changeset 31084 f4db921165ce
parent 30664 167da873c3b3
child 31847 7de0e20ca24d
permissions -rw-r--r--
fixed HOLCF proofs
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
    Author:     Stefan Berghofer, TU Muenchen
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
     3
*)
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
     4
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
     5
header {* Implementation of finite sets by lists *}
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
     6
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
     7
theory Executable_Set
30664
167da873c3b3 Main is (Complex_Main) base entry point in library theories
haftmann
parents: 30304
diff changeset
     8
imports Main
23854
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
     9
begin
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    10
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    11
subsection {* Definitional rewrites *}
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    12
28522
eacb54d9e78d only one theorem table for both code generators
haftmann
parents: 27487
diff changeset
    13
definition subset :: "'a set \<Rightarrow> 'a set \<Rightarrow> bool" where
eacb54d9e78d only one theorem table for both code generators
haftmann
parents: 27487
diff changeset
    14
  "subset = op \<le>"
eacb54d9e78d only one theorem table for both code generators
haftmann
parents: 27487
diff changeset
    15
eacb54d9e78d only one theorem table for both code generators
haftmann
parents: 27487
diff changeset
    16
declare subset_def [symmetric, code unfold]
eacb54d9e78d only one theorem table for both code generators
haftmann
parents: 27487
diff changeset
    17
28939
08004ce1b167 added code equation for subset
haftmann
parents: 28522
diff changeset
    18
lemma [code]: "subset A B \<longleftrightarrow> (\<forall>x\<in>A. x \<in> B)"
28522
eacb54d9e78d only one theorem table for both code generators
haftmann
parents: 27487
diff changeset
    19
  unfolding subset_def subset_eq ..
23854
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    20
28522
eacb54d9e78d only one theorem table for both code generators
haftmann
parents: 27487
diff changeset
    21
definition is_empty :: "'a set \<Rightarrow> bool" where
eacb54d9e78d only one theorem table for both code generators
haftmann
parents: 27487
diff changeset
    22
  "is_empty A \<longleftrightarrow> A = {}"
eacb54d9e78d only one theorem table for both code generators
haftmann
parents: 27487
diff changeset
    23
eacb54d9e78d only one theorem table for both code generators
haftmann
parents: 27487
diff changeset
    24
definition eq_set :: "'a set \<Rightarrow> 'a set \<Rightarrow> bool" where
eacb54d9e78d only one theorem table for both code generators
haftmann
parents: 27487
diff changeset
    25
  [code del]: "eq_set = op ="
eacb54d9e78d only one theorem table for both code generators
haftmann
parents: 27487
diff changeset
    26
eacb54d9e78d only one theorem table for both code generators
haftmann
parents: 27487
diff changeset
    27
lemma [code]: "eq_set A B \<longleftrightarrow> A \<subseteq> B \<and> B \<subseteq> A"
eacb54d9e78d only one theorem table for both code generators
haftmann
parents: 27487
diff changeset
    28
  unfolding eq_set_def by auto
26816
e82229ee8f43 - Declared subset_eq as code lemma
berghofe
parents: 26312
diff changeset
    29
29107
e70b9c2bee14 code for {x:A. P(x)} and for fold
nipkow
parents: 28939
diff changeset
    30
(* FIXME allow for Stefan's code generator:
e70b9c2bee14 code for {x:A. P(x)} and for fold
nipkow
parents: 28939
diff changeset
    31
declare set_eq_subset[code unfold]
e70b9c2bee14 code for {x:A. P(x)} and for fold
nipkow
parents: 28939
diff changeset
    32
*)
e70b9c2bee14 code for {x:A. P(x)} and for fold
nipkow
parents: 28939
diff changeset
    33
23854
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    34
lemma [code]:
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    35
  "a \<in> A \<longleftrightarrow> (\<exists>x\<in>A. x = a)"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    36
  unfolding bex_triv_one_point1 ..
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    37
28939
08004ce1b167 added code equation for subset
haftmann
parents: 28522
diff changeset
    38
definition filter_set :: "('a \<Rightarrow> bool) \<Rightarrow> 'a set \<Rightarrow> 'a set" where
23854
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    39
  "filter_set P xs = {x\<in>xs. P x}"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    40
29107
e70b9c2bee14 code for {x:A. P(x)} and for fold
nipkow
parents: 28939
diff changeset
    41
declare filter_set_def[symmetric, code unfold] 
e70b9c2bee14 code for {x:A. P(x)} and for fold
nipkow
parents: 28939
diff changeset
    42
23854
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    43
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    44
subsection {* Operations on lists *}
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    45
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    46
subsubsection {* Basic definitions *}
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    47
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    48
definition
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    49
  flip :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> 'b \<Rightarrow> 'a \<Rightarrow> 'c" where
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    50
  "flip f a b = f b a"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    51
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    52
definition
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    53
  member :: "'a list \<Rightarrow> 'a \<Rightarrow> bool" where
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    54
  "member xs x \<longleftrightarrow> x \<in> set xs"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    55
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    56
definition
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    57
  insertl :: "'a \<Rightarrow> 'a list \<Rightarrow> 'a list" where
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    58
  "insertl x xs = (if member xs x then xs else x#xs)"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    59
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    60
lemma [code target: List]: "member [] y \<longleftrightarrow> False"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    61
  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
    62
  unfolding member_def by (induct xs) simp_all
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    63
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    64
fun
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    65
  drop_first :: "('a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> 'a list" where
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    66
  "drop_first f [] = []"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    67
| "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
    68
declare drop_first.simps [code del]
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    69
declare drop_first.simps [code target: List]
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    70
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    71
declare remove1.simps [code del]
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    72
lemma [code target: List]:
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    73
  "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
    74
proof (cases "member xs x")
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    75
  case False thus ?thesis unfolding member_def by (induct xs) auto
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    76
next
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    77
  case True
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    78
  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
    79
  with True show ?thesis by simp
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    80
qed
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    81
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    82
lemma member_nil [simp]:
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    83
  "member [] = (\<lambda>x. False)"
26816
e82229ee8f43 - Declared subset_eq as code lemma
berghofe
parents: 26312
diff changeset
    84
proof (rule ext)
23854
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    85
  fix x
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    86
  show "member [] x = False" unfolding member_def by simp
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    87
qed
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    88
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    89
lemma member_insertl [simp]:
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    90
  "x \<in> set (insertl x xs)"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    91
  unfolding insertl_def member_def mem_iff by simp
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    92
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    93
lemma insertl_member [simp]:
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    94
  fixes xs x
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    95
  assumes member: "member xs x"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    96
  shows "insertl x xs = xs"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    97
  using member unfolding insertl_def by simp
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    98
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    99
lemma insertl_not_member [simp]:
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   100
  fixes xs x
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   101
  assumes member: "\<not> (member xs x)"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   102
  shows "insertl x xs = x # xs"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   103
  using member unfolding insertl_def by simp
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   104
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   105
lemma foldr_remove1_empty [simp]:
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   106
  "foldr remove1 xs [] = []"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   107
  by (induct xs) simp_all
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   108
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   109
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   110
subsubsection {* Derived definitions *}
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   111
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   112
function unionl :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   113
where
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   114
  "unionl [] ys = ys"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   115
| "unionl xs ys = foldr insertl xs ys"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   116
by pat_completeness auto
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   117
termination by lexicographic_order
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   118
26312
e9a65675e5e8 avoid rebinding of existing facts;
wenzelm
parents: 25885
diff changeset
   119
lemmas unionl_eq = unionl.simps(2)
23854
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   120
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   121
function intersect :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   122
where
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   123
  "intersect [] ys = []"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   124
| "intersect xs [] = []"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   125
| "intersect xs ys = filter (member xs) ys"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   126
by pat_completeness auto
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   127
termination by lexicographic_order
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   128
26312
e9a65675e5e8 avoid rebinding of existing facts;
wenzelm
parents: 25885
diff changeset
   129
lemmas intersect_eq = intersect.simps(3)
23854
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   130
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   131
function subtract :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   132
where
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   133
  "subtract [] ys = ys"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   134
| "subtract xs [] = []"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   135
| "subtract xs ys = foldr remove1 xs ys"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   136
by pat_completeness auto
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   137
termination by lexicographic_order
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   138
26312
e9a65675e5e8 avoid rebinding of existing facts;
wenzelm
parents: 25885
diff changeset
   139
lemmas subtract_eq = subtract.simps(3)
23854
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   140
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   141
function map_distinct :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a list \<Rightarrow> 'b list"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   142
where
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   143
  "map_distinct f [] = []"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   144
| "map_distinct f xs = foldr (insertl o f) xs []"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   145
by pat_completeness auto
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   146
termination by lexicographic_order
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   147
26312
e9a65675e5e8 avoid rebinding of existing facts;
wenzelm
parents: 25885
diff changeset
   148
lemmas map_distinct_eq = map_distinct.simps(2)
23854
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   149
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   150
function unions :: "'a list list \<Rightarrow> 'a list"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   151
where
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   152
  "unions [] = []"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   153
| "unions xs = foldr unionl xs []"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   154
by pat_completeness auto
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   155
termination by lexicographic_order
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   156
26312
e9a65675e5e8 avoid rebinding of existing facts;
wenzelm
parents: 25885
diff changeset
   157
lemmas unions_eq = unions.simps(2)
23854
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   158
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   159
consts intersects :: "'a list list \<Rightarrow> 'a list"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   160
primrec
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   161
  "intersects (x#xs) = foldr intersect xs x"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   162
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   163
definition
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   164
  map_union :: "'a list \<Rightarrow> ('a \<Rightarrow> 'b list) \<Rightarrow> 'b list" where
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   165
  "map_union xs f = unions (map f xs)"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   166
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   167
definition
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   168
  map_inter :: "'a list \<Rightarrow> ('a \<Rightarrow> 'b list) \<Rightarrow> 'b list" where
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   169
  "map_inter xs f = intersects (map f xs)"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   170
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   171
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   172
subsection {* Isomorphism proofs *}
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   173
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   174
lemma iso_member:
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   175
  "member xs x \<longleftrightarrow> x \<in> set xs"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   176
  unfolding member_def mem_iff ..
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   177
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   178
lemma iso_insert:
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   179
  "set (insertl x xs) = insert x (set xs)"
30304
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29110
diff changeset
   180
  unfolding insertl_def iso_member by (simp add: insert_absorb)
23854
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   181
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   182
lemma iso_remove1:
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   183
  assumes distnct: "distinct xs"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   184
  shows "set (remove1 x xs) = set xs - {x}"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   185
  using distnct set_remove1_eq by auto
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   186
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   187
lemma iso_union:
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   188
  "set (unionl xs ys) = set xs \<union> set ys"
26312
e9a65675e5e8 avoid rebinding of existing facts;
wenzelm
parents: 25885
diff changeset
   189
  unfolding unionl_eq
23854
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   190
  by (induct xs arbitrary: ys) (simp_all add: iso_insert)
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   191
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   192
lemma iso_intersect:
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   193
  "set (intersect xs ys) = set xs \<inter> set ys"
26312
e9a65675e5e8 avoid rebinding of existing facts;
wenzelm
parents: 25885
diff changeset
   194
  unfolding intersect_eq Int_def by (simp add: Int_def iso_member) auto
23854
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   195
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   196
definition
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   197
  subtract' :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" where
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   198
  "subtract' = flip subtract"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   199
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   200
lemma iso_subtract:
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   201
  fixes ys
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   202
  assumes distnct: "distinct ys"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   203
  shows "set (subtract' ys xs) = set ys - set xs"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   204
    and "distinct (subtract' ys xs)"
26312
e9a65675e5e8 avoid rebinding of existing facts;
wenzelm
parents: 25885
diff changeset
   205
  unfolding subtract'_def flip_def subtract_eq
23854
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   206
  using distnct by (induct xs arbitrary: ys) auto
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   207
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   208
lemma iso_map_distinct:
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   209
  "set (map_distinct f xs) = image f (set xs)"
26312
e9a65675e5e8 avoid rebinding of existing facts;
wenzelm
parents: 25885
diff changeset
   210
  unfolding map_distinct_eq by (induct xs) (simp_all add: iso_insert)
23854
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   211
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   212
lemma iso_unions:
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   213
  "set (unions xss) = \<Union> set (map set xss)"
26312
e9a65675e5e8 avoid rebinding of existing facts;
wenzelm
parents: 25885
diff changeset
   214
  unfolding unions_eq
23854
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   215
proof (induct xss)
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   216
  case Nil show ?case by simp
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   217
next
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   218
  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
   219
qed
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   220
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   221
lemma iso_intersects:
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   222
  "set (intersects (xs#xss)) = \<Inter> set (map set (xs#xss))"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   223
  by (induct xss) (simp_all add: Int_def iso_member, auto)
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   224
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   225
lemma iso_UNION:
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   226
  "set (map_union xs f) = UNION (set xs) (set o f)"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   227
  unfolding map_union_def iso_unions by simp
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   228
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   229
lemma iso_INTER:
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   230
  "set (map_inter (x#xs) f) = INTER (set (x#xs)) (set o f)"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   231
  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
   232
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   233
definition
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   234
  Blall :: "'a list \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> bool" where
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   235
  "Blall = flip list_all"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   236
definition
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   237
  Blex :: "'a list \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> bool" where
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   238
  "Blex = flip list_ex"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   239
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   240
lemma iso_Ball:
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   241
  "Blall xs f = Ball (set xs) f"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   242
  unfolding Blall_def flip_def by (induct xs) simp_all
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   243
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   244
lemma iso_Bex:
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   245
  "Blex xs f = Bex (set xs) f"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   246
  unfolding Blex_def flip_def by (induct xs) simp_all
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   247
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   248
lemma iso_filter:
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   249
  "set (filter P xs) = filter_set P (set xs)"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   250
  unfolding filter_set_def by (induct xs) auto
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   251
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   252
subsection {* code generator setup *}
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   253
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   254
ML {*
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   255
nonfix inter;
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   256
nonfix union;
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   257
nonfix subset;
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   258
*}
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   259
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   260
subsubsection {* const serializations *}
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   261
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   262
consts_code
30304
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29110
diff changeset
   263
  "Set.empty" ("{*[]*}")
23854
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   264
  insert ("{*insertl*}")
28522
eacb54d9e78d only one theorem table for both code generators
haftmann
parents: 27487
diff changeset
   265
  is_empty ("{*null*}")
23854
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   266
  "op \<union>" ("{*unionl*}")
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   267
  "op \<inter>" ("{*intersect*}")
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   268
  "op - \<Colon> 'a set \<Rightarrow> 'a set \<Rightarrow> 'a set" ("{* flip subtract *}")
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   269
  image ("{*map_distinct*}")
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   270
  Union ("{*unions*}")
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   271
  Inter ("{*intersects*}")
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   272
  UNION ("{*map_union*}")
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   273
  INTER ("{*map_inter*}")
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   274
  Ball ("{*Blall*}")
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   275
  Bex ("{*Blex*}")
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   276
  filter_set ("{*filter*}")
29110
476c46e99ada flipped fold implementation
nipkow
parents: 29107
diff changeset
   277
  fold ("{* foldl o flip *}")
23854
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   278
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
   279
end