src/HOL/Library/Executable_Set.thy
author wenzelm
Fri, 16 Apr 2010 21:28:09 +0200
changeset 36176 3fe7e97ccca8
parent 34980 6676fd863e02
child 37023 efc202e1677e
permissions -rw-r--r--
replaced generic 'hide' command by more conventional 'hide_class', 'hide_type', 'hide_const', 'hide_fact' -- frees some popular keywords;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
34020
2573c794034c merged Crude_Executable_Set into Executable_Set
haftmann
parents: 34008
diff changeset
     1
(*  Title:      HOL/Library/Executable_Set.thy
2573c794034c merged Crude_Executable_Set into Executable_Set
haftmann
parents: 34008
diff changeset
     2
    Author:     Stefan Berghofer, TU Muenchen
33947
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
     3
    Author:     Florian Haftmann, TU Muenchen
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
     4
*)
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
     5
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
     6
header {* A crude implementation of finite sets by lists -- avoid using this at any cost! *}
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
     7
34020
2573c794034c merged Crude_Executable_Set into Executable_Set
haftmann
parents: 34008
diff changeset
     8
theory Executable_Set
33947
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
     9
imports List_Set
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    10
begin
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    11
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    12
declare mem_def [code del]
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    13
declare Collect_def [code del]
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    14
declare insert_code [code del]
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    15
declare vimage_code [code del]
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    16
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    17
subsection {* Set representation *}
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    18
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    19
setup {*
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    20
  Code.add_type_cmd "set"
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    21
*}
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    22
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    23
definition Set :: "'a list \<Rightarrow> 'a set" where
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    24
  [simp]: "Set = set"
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    25
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    26
definition Coset :: "'a list \<Rightarrow> 'a set" where
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    27
  [simp]: "Coset xs = - set xs"
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    28
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    29
setup {*
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    30
  Code.add_signature_cmd ("Set", "'a list \<Rightarrow> 'a set")
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    31
  #> Code.add_signature_cmd ("Coset", "'a list \<Rightarrow> 'a set")
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    32
  #> Code.add_signature_cmd ("set", "'a list \<Rightarrow> 'a set")
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    33
  #> Code.add_signature_cmd ("op \<in>", "'a \<Rightarrow> 'a set \<Rightarrow> bool")
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    34
*}
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    35
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    36
code_datatype Set Coset
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    37
34020
2573c794034c merged Crude_Executable_Set into Executable_Set
haftmann
parents: 34008
diff changeset
    38
consts_code
2573c794034c merged Crude_Executable_Set into Executable_Set
haftmann
parents: 34008
diff changeset
    39
  Coset ("\<module>Coset")
2573c794034c merged Crude_Executable_Set into Executable_Set
haftmann
parents: 34008
diff changeset
    40
  Set ("\<module>Set")
2573c794034c merged Crude_Executable_Set into Executable_Set
haftmann
parents: 34008
diff changeset
    41
attach {*
2573c794034c merged Crude_Executable_Set into Executable_Set
haftmann
parents: 34008
diff changeset
    42
  datatype 'a set = Set of 'a list | Coset of 'a list;
2573c794034c merged Crude_Executable_Set into Executable_Set
haftmann
parents: 34008
diff changeset
    43
*} -- {* This assumes that there won't be a @{text Coset} without a @{text Set} *}
2573c794034c merged Crude_Executable_Set into Executable_Set
haftmann
parents: 34008
diff changeset
    44
33947
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    45
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    46
subsection {* Basic operations *}
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    47
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    48
lemma [code]:
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    49
  "set xs = Set (remdups xs)"
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    50
  by simp
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    51
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    52
lemma [code]:
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    53
  "x \<in> Set xs \<longleftrightarrow> member x xs"
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    54
  "x \<in> Coset xs \<longleftrightarrow> \<not> member x xs"
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    55
  by (simp_all add: mem_iff)
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    56
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    57
definition is_empty :: "'a set \<Rightarrow> bool" where
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    58
  [simp]: "is_empty A \<longleftrightarrow> A = {}"
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    59
34020
2573c794034c merged Crude_Executable_Set into Executable_Set
haftmann
parents: 34008
diff changeset
    60
lemma [code_unfold]:
33947
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    61
  "A = {} \<longleftrightarrow> is_empty A"
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    62
  by simp
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    63
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    64
definition empty :: "'a set" where
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    65
  [simp]: "empty = {}"
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    66
34020
2573c794034c merged Crude_Executable_Set into Executable_Set
haftmann
parents: 34008
diff changeset
    67
lemma [code_unfold]:
33947
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    68
  "{} = empty"
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    69
  by simp
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    70
34020
2573c794034c merged Crude_Executable_Set into Executable_Set
haftmann
parents: 34008
diff changeset
    71
lemma [code_unfold, code_inline del]:
2573c794034c merged Crude_Executable_Set into Executable_Set
haftmann
parents: 34008
diff changeset
    72
  "empty = Set []"
2573c794034c merged Crude_Executable_Set into Executable_Set
haftmann
parents: 34008
diff changeset
    73
  by simp -- {* Otherwise @{text \<eta>}-expansion produces funny things. *}
2573c794034c merged Crude_Executable_Set into Executable_Set
haftmann
parents: 34008
diff changeset
    74
33947
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    75
setup {*
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    76
  Code.add_signature_cmd ("is_empty", "'a set \<Rightarrow> bool")
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    77
  #> Code.add_signature_cmd ("empty", "'a set")
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    78
  #> Code.add_signature_cmd ("insert", "'a \<Rightarrow> 'a set \<Rightarrow> 'a set")
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    79
  #> Code.add_signature_cmd ("List_Set.remove", "'a \<Rightarrow> 'a set \<Rightarrow> 'a set")
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    80
  #> Code.add_signature_cmd ("image", "('a \<Rightarrow> 'b) \<Rightarrow> 'a set \<Rightarrow> 'b set")
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    81
  #> Code.add_signature_cmd ("List_Set.project", "('a \<Rightarrow> bool) \<Rightarrow> 'a set \<Rightarrow> 'a set")
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    82
  #> Code.add_signature_cmd ("Ball", "'a set \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> bool")
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    83
  #> Code.add_signature_cmd ("Bex", "'a set \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> bool")
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    84
  #> Code.add_signature_cmd ("card", "'a set \<Rightarrow> nat")
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    85
*}
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    86
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    87
lemma is_empty_Set [code]:
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    88
  "is_empty (Set xs) \<longleftrightarrow> null xs"
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    89
  by (simp add: empty_null)
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    90
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    91
lemma empty_Set [code]:
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    92
  "empty = Set []"
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    93
  by simp
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    94
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    95
lemma insert_Set [code]:
34980
6676fd863e02 adjusted to changes in List_Set.thy
haftmann
parents: 34022
diff changeset
    96
  "insert x (Set xs) = Set (List.insert x xs)"
6676fd863e02 adjusted to changes in List_Set.thy
haftmann
parents: 34022
diff changeset
    97
  "insert x (Coset xs) = Coset (removeAll x xs)"
6676fd863e02 adjusted to changes in List_Set.thy
haftmann
parents: 34022
diff changeset
    98
  by (simp_all add: set_insert)
33947
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    99
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   100
lemma remove_Set [code]:
34980
6676fd863e02 adjusted to changes in List_Set.thy
haftmann
parents: 34022
diff changeset
   101
  "remove x (Set xs) = Set (removeAll x xs)"
6676fd863e02 adjusted to changes in List_Set.thy
haftmann
parents: 34022
diff changeset
   102
  "remove x (Coset xs) = Coset (List.insert x xs)"
6676fd863e02 adjusted to changes in List_Set.thy
haftmann
parents: 34022
diff changeset
   103
  by (auto simp add: set_insert remove_def)
33947
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   104
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   105
lemma image_Set [code]:
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   106
  "image f (Set xs) = Set (remdups (map f xs))"
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   107
  by simp
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   108
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   109
lemma project_Set [code]:
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   110
  "project P (Set xs) = Set (filter P xs)"
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   111
  by (simp add: project_set)
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   112
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   113
lemma Ball_Set [code]:
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   114
  "Ball (Set xs) P \<longleftrightarrow> list_all P xs"
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   115
  by (simp add: ball_set)
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   116
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   117
lemma Bex_Set [code]:
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   118
  "Bex (Set xs) P \<longleftrightarrow> list_ex P xs"
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   119
  by (simp add: bex_set)
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   120
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   121
lemma card_Set [code]:
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   122
  "card (Set xs) = length (remdups xs)"
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   123
proof -
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   124
  have "card (set (remdups xs)) = length (remdups xs)"
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   125
    by (rule distinct_card) simp
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   126
  then show ?thesis by simp
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   127
qed
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   128
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   129
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   130
subsection {* Derived operations *}
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   131
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   132
definition set_eq :: "'a set \<Rightarrow> 'a set \<Rightarrow> bool" where
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   133
  [simp]: "set_eq = op ="
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   134
34020
2573c794034c merged Crude_Executable_Set into Executable_Set
haftmann
parents: 34008
diff changeset
   135
lemma [code_unfold]:
33947
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   136
  "op = = set_eq"
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   137
  by simp
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   138
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   139
definition subset_eq :: "'a set \<Rightarrow> 'a set \<Rightarrow> bool" where
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   140
  [simp]: "subset_eq = op \<subseteq>"
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   141
34020
2573c794034c merged Crude_Executable_Set into Executable_Set
haftmann
parents: 34008
diff changeset
   142
lemma [code_unfold]:
33947
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   143
  "op \<subseteq> = subset_eq"
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   144
  by simp
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   145
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   146
definition subset :: "'a set \<Rightarrow> 'a set \<Rightarrow> bool" where
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   147
  [simp]: "subset = op \<subset>"
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   148
34020
2573c794034c merged Crude_Executable_Set into Executable_Set
haftmann
parents: 34008
diff changeset
   149
lemma [code_unfold]:
33947
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   150
  "op \<subset> = subset"
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   151
  by simp
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   152
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   153
setup {*
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   154
  Code.add_signature_cmd ("set_eq", "'a set \<Rightarrow> 'a set \<Rightarrow> bool")
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   155
  #> Code.add_signature_cmd ("subset_eq", "'a set \<Rightarrow> 'a set \<Rightarrow> bool")
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   156
  #> Code.add_signature_cmd ("subset", "'a set \<Rightarrow> 'a set \<Rightarrow> bool")
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   157
*}
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   158
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   159
lemma set_eq_subset_eq [code]:
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   160
  "set_eq A B \<longleftrightarrow> subset_eq A B \<and> subset_eq B A"
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   161
  by auto
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   162
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   163
lemma subset_eq_forall [code]:
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   164
  "subset_eq A B \<longleftrightarrow> (\<forall>x\<in>A. x \<in> B)"
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   165
  by (simp add: subset_eq)
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   166
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   167
lemma subset_subset_eq [code]:
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   168
  "subset A B \<longleftrightarrow> subset_eq A B \<and> \<not> subset_eq B A"
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   169
  by (simp add: subset)
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   170
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   171
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   172
subsection {* Functorial operations *}
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   173
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   174
definition inter :: "'a set \<Rightarrow> 'a set \<Rightarrow> 'a set" where
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   175
  [simp]: "inter = op \<inter>"
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   176
34020
2573c794034c merged Crude_Executable_Set into Executable_Set
haftmann
parents: 34008
diff changeset
   177
lemma [code_unfold]:
33947
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   178
  "op \<inter> = inter"
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   179
  by simp
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   180
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   181
definition subtract :: "'a set \<Rightarrow> 'a set \<Rightarrow> 'a set" where
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   182
  [simp]: "subtract A B = B - A"
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   183
34020
2573c794034c merged Crude_Executable_Set into Executable_Set
haftmann
parents: 34008
diff changeset
   184
lemma [code_unfold]:
33947
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   185
  "B - A = subtract A B"
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   186
  by simp
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   187
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   188
definition union :: "'a set \<Rightarrow> 'a set \<Rightarrow> 'a set" where
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   189
  [simp]: "union = op \<union>"
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   190
34020
2573c794034c merged Crude_Executable_Set into Executable_Set
haftmann
parents: 34008
diff changeset
   191
lemma [code_unfold]:
33947
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   192
  "op \<union> = union"
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   193
  by simp
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   194
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   195
definition Inf :: "'a::complete_lattice set \<Rightarrow> 'a" where
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   196
  [simp]: "Inf = Complete_Lattice.Inf"
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   197
34020
2573c794034c merged Crude_Executable_Set into Executable_Set
haftmann
parents: 34008
diff changeset
   198
lemma [code_unfold]:
33947
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   199
  "Complete_Lattice.Inf = Inf"
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   200
  by simp
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   201
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   202
definition Sup :: "'a::complete_lattice set \<Rightarrow> 'a" where
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   203
  [simp]: "Sup = Complete_Lattice.Sup"
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   204
34020
2573c794034c merged Crude_Executable_Set into Executable_Set
haftmann
parents: 34008
diff changeset
   205
lemma [code_unfold]:
33947
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   206
  "Complete_Lattice.Sup = Sup"
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   207
  by simp
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   208
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   209
definition Inter :: "'a set set \<Rightarrow> 'a set" where
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   210
  [simp]: "Inter = Inf"
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   211
34020
2573c794034c merged Crude_Executable_Set into Executable_Set
haftmann
parents: 34008
diff changeset
   212
lemma [code_unfold]:
33947
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   213
  "Inf = Inter"
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   214
  by simp
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   215
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   216
definition Union :: "'a set set \<Rightarrow> 'a set" where
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   217
  [simp]: "Union = Sup"
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   218
34020
2573c794034c merged Crude_Executable_Set into Executable_Set
haftmann
parents: 34008
diff changeset
   219
lemma [code_unfold]:
33947
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   220
  "Sup = Union"
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   221
  by simp
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   222
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   223
setup {*
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   224
  Code.add_signature_cmd ("inter", "'a set \<Rightarrow> 'a set \<Rightarrow> 'a set")
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   225
  #> Code.add_signature_cmd ("subtract", "'a set \<Rightarrow> 'a set \<Rightarrow> 'a set")
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   226
  #> Code.add_signature_cmd ("union", "'a set \<Rightarrow> 'a set \<Rightarrow> 'a set")
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   227
  #> Code.add_signature_cmd ("Inf", "'a set \<Rightarrow> 'a")
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   228
  #> Code.add_signature_cmd ("Sup", "'a set \<Rightarrow> 'a")
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   229
  #> Code.add_signature_cmd ("Inter", "'a set set \<Rightarrow> 'a set")
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   230
  #> Code.add_signature_cmd ("Union", "'a set set \<Rightarrow> 'a set")
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   231
*}
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   232
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   233
lemma inter_project [code]:
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   234
  "inter A (Set xs) = Set (List.filter (\<lambda>x. x \<in> A) xs)"
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   235
  "inter A (Coset xs) = foldl (\<lambda>A x. remove x A) A xs"
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   236
  by (simp add: inter project_def, simp add: Diff_eq [symmetric] minus_set)
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   237
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   238
lemma subtract_remove [code]:
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   239
  "subtract (Set xs) A = foldl (\<lambda>A x. remove x A) A xs"
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   240
  "subtract (Coset xs) A = Set (List.filter (\<lambda>x. x \<in> A) xs)"
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   241
  by (auto simp add: minus_set)
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   242
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   243
lemma union_insert [code]:
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   244
  "union (Set xs) A = foldl (\<lambda>A x. insert x A) A xs"
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   245
  "union (Coset xs) A = Coset (List.filter (\<lambda>x. x \<notin> A) xs)"
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   246
  by (auto simp add: union_set)
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   247
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   248
lemma Inf_inf [code]:
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   249
  "Inf (Set xs) = foldl inf (top :: 'a::complete_lattice) xs"
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   250
  "Inf (Coset []) = (bot :: 'a::complete_lattice)"
34008
1ce58e8c02ee tuned proofs
haftmann
parents: 33947
diff changeset
   251
  by (simp_all add: Inf_UNIV Inf_set_fold)
33947
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   252
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   253
lemma Sup_sup [code]:
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   254
  "Sup (Set xs) = foldl sup (bot :: 'a::complete_lattice) xs"
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   255
  "Sup (Coset []) = (top :: 'a::complete_lattice)"
34008
1ce58e8c02ee tuned proofs
haftmann
parents: 33947
diff changeset
   256
  by (simp_all add: Sup_UNIV Sup_set_fold)
33947
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   257
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   258
lemma Inter_inter [code]:
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   259
  "Inter (Set xs) = foldl inter (Coset []) xs"
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   260
  "Inter (Coset []) = empty"
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   261
  unfolding Inter_def Inf_inf by simp_all
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   262
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   263
lemma Union_union [code]:
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   264
  "Union (Set xs) = foldl union empty xs"
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   265
  "Union (Coset []) = Coset []"
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   266
  unfolding Union_def Sup_sup by simp_all
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   267
36176
3fe7e97ccca8 replaced generic 'hide' command by more conventional 'hide_class', 'hide_type', 'hide_const', 'hide_fact' -- frees some popular keywords;
wenzelm
parents: 34980
diff changeset
   268
hide_const (open) is_empty empty remove
33947
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   269
  set_eq subset_eq subset inter union subtract Inf Sup Inter Union
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   270
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   271
end