src/HOL/Library/Executable_Set.thy
author bulwahn
Fri, 11 Mar 2011 15:21:13 +0100
changeset 41930 1e008cc4883a
parent 40672 abd4e7358847
child 43363 eaf8b7f22d39
permissions -rw-r--r--
renaming lazysmallcheck ML file to Quickcheck_Narrowing
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
37024
e938a0b5286e renamed List_Set to the now more appropriate More_Set
haftmann
parents: 37023
diff changeset
     9
imports More_Set
33947
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    10
begin
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    11
39141
5ec8e4404c33 added more explicit warning
haftmann
parents: 38304
diff changeset
    12
text {*
5ec8e4404c33 added more explicit warning
haftmann
parents: 38304
diff changeset
    13
  This is just an ad-hoc hack which will rarely give you what you want.
5ec8e4404c33 added more explicit warning
haftmann
parents: 38304
diff changeset
    14
  For the moment, whenever you need executable sets, consider using
40672
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 39141
diff changeset
    15
  type @{text fset} from theory @{text Cset}.
39141
5ec8e4404c33 added more explicit warning
haftmann
parents: 38304
diff changeset
    16
*}
5ec8e4404c33 added more explicit warning
haftmann
parents: 38304
diff changeset
    17
33947
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    18
declare mem_def [code del]
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    19
declare Collect_def [code del]
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    20
declare insert_code [code del]
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    21
declare vimage_code [code del]
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    22
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    23
subsection {* Set representation *}
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    24
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    25
setup {*
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    26
  Code.add_type_cmd "set"
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    27
*}
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    28
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    29
definition Set :: "'a list \<Rightarrow> 'a set" where
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    30
  [simp]: "Set = set"
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    31
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    32
definition Coset :: "'a list \<Rightarrow> 'a set" where
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    33
  [simp]: "Coset xs = - set xs"
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    34
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    35
setup {*
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    36
  Code.add_signature_cmd ("Set", "'a list \<Rightarrow> 'a set")
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    37
  #> Code.add_signature_cmd ("Coset", "'a list \<Rightarrow> 'a set")
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    38
  #> Code.add_signature_cmd ("set", "'a list \<Rightarrow> 'a set")
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    39
  #> Code.add_signature_cmd ("op \<in>", "'a \<Rightarrow> 'a set \<Rightarrow> bool")
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    40
*}
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    41
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    42
code_datatype Set Coset
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    43
34020
2573c794034c merged Crude_Executable_Set into Executable_Set
haftmann
parents: 34008
diff changeset
    44
consts_code
2573c794034c merged Crude_Executable_Set into Executable_Set
haftmann
parents: 34008
diff changeset
    45
  Coset ("\<module>Coset")
2573c794034c merged Crude_Executable_Set into Executable_Set
haftmann
parents: 34008
diff changeset
    46
  Set ("\<module>Set")
2573c794034c merged Crude_Executable_Set into Executable_Set
haftmann
parents: 34008
diff changeset
    47
attach {*
2573c794034c merged Crude_Executable_Set into Executable_Set
haftmann
parents: 34008
diff changeset
    48
  datatype 'a set = Set of 'a list | Coset of 'a list;
2573c794034c merged Crude_Executable_Set into Executable_Set
haftmann
parents: 34008
diff changeset
    49
*} -- {* 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
    50
33947
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    51
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    52
subsection {* Basic operations *}
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    53
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    54
lemma [code]:
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    55
  "set xs = Set (remdups xs)"
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    56
  by simp
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    57
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    58
lemma [code]:
37595
9591362629e3 dropped ancient infix mem; refined code generation operations in List.thy
haftmann
parents: 37024
diff changeset
    59
  "x \<in> Set xs \<longleftrightarrow> List.member xs x"
9591362629e3 dropped ancient infix mem; refined code generation operations in List.thy
haftmann
parents: 37024
diff changeset
    60
  "x \<in> Coset xs \<longleftrightarrow> \<not> List.member xs x"
9591362629e3 dropped ancient infix mem; refined code generation operations in List.thy
haftmann
parents: 37024
diff changeset
    61
  by (simp_all add: member_def)
33947
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    62
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    63
definition is_empty :: "'a set \<Rightarrow> bool" where
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    64
  [simp]: "is_empty A \<longleftrightarrow> A = {}"
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    65
34020
2573c794034c merged Crude_Executable_Set into Executable_Set
haftmann
parents: 34008
diff changeset
    66
lemma [code_unfold]:
33947
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    67
  "A = {} \<longleftrightarrow> is_empty A"
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    68
  by simp
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    69
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    70
definition empty :: "'a set" where
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    71
  [simp]: "empty = {}"
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    72
34020
2573c794034c merged Crude_Executable_Set into Executable_Set
haftmann
parents: 34008
diff changeset
    73
lemma [code_unfold]:
33947
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    74
  "{} = empty"
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    75
  by simp
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    76
34020
2573c794034c merged Crude_Executable_Set into Executable_Set
haftmann
parents: 34008
diff changeset
    77
lemma [code_unfold, code_inline del]:
2573c794034c merged Crude_Executable_Set into Executable_Set
haftmann
parents: 34008
diff changeset
    78
  "empty = Set []"
2573c794034c merged Crude_Executable_Set into Executable_Set
haftmann
parents: 34008
diff changeset
    79
  by simp -- {* Otherwise @{text \<eta>}-expansion produces funny things. *}
2573c794034c merged Crude_Executable_Set into Executable_Set
haftmann
parents: 34008
diff changeset
    80
33947
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    81
setup {*
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    82
  Code.add_signature_cmd ("is_empty", "'a set \<Rightarrow> bool")
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    83
  #> Code.add_signature_cmd ("empty", "'a set")
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    84
  #> Code.add_signature_cmd ("insert", "'a \<Rightarrow> 'a set \<Rightarrow> 'a set")
37024
e938a0b5286e renamed List_Set to the now more appropriate More_Set
haftmann
parents: 37023
diff changeset
    85
  #> Code.add_signature_cmd ("More_Set.remove", "'a \<Rightarrow> 'a set \<Rightarrow> 'a set")
33947
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    86
  #> Code.add_signature_cmd ("image", "('a \<Rightarrow> 'b) \<Rightarrow> 'a set \<Rightarrow> 'b set")
37024
e938a0b5286e renamed List_Set to the now more appropriate More_Set
haftmann
parents: 37023
diff changeset
    87
  #> Code.add_signature_cmd ("More_Set.project", "('a \<Rightarrow> bool) \<Rightarrow> 'a set \<Rightarrow> 'a set")
33947
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    88
  #> Code.add_signature_cmd ("Ball", "'a set \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> bool")
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    89
  #> Code.add_signature_cmd ("Bex", "'a set \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> bool")
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    90
  #> Code.add_signature_cmd ("card", "'a set \<Rightarrow> nat")
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    91
*}
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    92
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    93
lemma is_empty_Set [code]:
37595
9591362629e3 dropped ancient infix mem; refined code generation operations in List.thy
haftmann
parents: 37024
diff changeset
    94
  "is_empty (Set xs) \<longleftrightarrow> List.null xs"
9591362629e3 dropped ancient infix mem; refined code generation operations in List.thy
haftmann
parents: 37024
diff changeset
    95
  by (simp add: null_def)
33947
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    96
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    97
lemma empty_Set [code]:
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    98
  "empty = Set []"
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
    99
  by simp
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   100
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   101
lemma insert_Set [code]:
34980
6676fd863e02 adjusted to changes in List_Set.thy
haftmann
parents: 34022
diff changeset
   102
  "insert x (Set xs) = Set (List.insert x xs)"
6676fd863e02 adjusted to changes in List_Set.thy
haftmann
parents: 34022
diff changeset
   103
  "insert x (Coset xs) = Coset (removeAll x xs)"
6676fd863e02 adjusted to changes in List_Set.thy
haftmann
parents: 34022
diff changeset
   104
  by (simp_all add: set_insert)
33947
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   105
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   106
lemma remove_Set [code]:
34980
6676fd863e02 adjusted to changes in List_Set.thy
haftmann
parents: 34022
diff changeset
   107
  "remove x (Set xs) = Set (removeAll x xs)"
6676fd863e02 adjusted to changes in List_Set.thy
haftmann
parents: 34022
diff changeset
   108
  "remove x (Coset xs) = Coset (List.insert x xs)"
6676fd863e02 adjusted to changes in List_Set.thy
haftmann
parents: 34022
diff changeset
   109
  by (auto simp add: set_insert remove_def)
33947
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   110
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   111
lemma image_Set [code]:
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   112
  "image f (Set xs) = Set (remdups (map f xs))"
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   113
  by simp
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   114
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   115
lemma project_Set [code]:
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   116
  "project P (Set xs) = Set (filter P xs)"
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   117
  by (simp add: project_set)
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   118
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   119
lemma Ball_Set [code]:
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   120
  "Ball (Set xs) P \<longleftrightarrow> list_all P xs"
37595
9591362629e3 dropped ancient infix mem; refined code generation operations in List.thy
haftmann
parents: 37024
diff changeset
   121
  by (simp add: list_all_iff)
33947
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   122
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   123
lemma Bex_Set [code]:
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   124
  "Bex (Set xs) P \<longleftrightarrow> list_ex P xs"
37595
9591362629e3 dropped ancient infix mem; refined code generation operations in List.thy
haftmann
parents: 37024
diff changeset
   125
  by (simp add: list_ex_iff)
33947
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   126
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   127
lemma card_Set [code]:
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   128
  "card (Set xs) = length (remdups xs)"
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   129
proof -
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   130
  have "card (set (remdups xs)) = length (remdups xs)"
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   131
    by (rule distinct_card) simp
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   132
  then show ?thesis by simp
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   133
qed
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   134
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   135
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   136
subsection {* Derived operations *}
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   137
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   138
definition set_eq :: "'a set \<Rightarrow> 'a set \<Rightarrow> bool" where
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   139
  [simp]: "set_eq = op ="
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   140
34020
2573c794034c merged Crude_Executable_Set into Executable_Set
haftmann
parents: 34008
diff changeset
   141
lemma [code_unfold]:
33947
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   142
  "op = = set_eq"
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   143
  by simp
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   144
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   145
definition subset_eq :: "'a set \<Rightarrow> 'a set \<Rightarrow> bool" where
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   146
  [simp]: "subset_eq = op \<subseteq>"
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   147
34020
2573c794034c merged Crude_Executable_Set into Executable_Set
haftmann
parents: 34008
diff changeset
   148
lemma [code_unfold]:
33947
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   149
  "op \<subseteq> = subset_eq"
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   150
  by simp
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   151
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   152
definition subset :: "'a set \<Rightarrow> 'a set \<Rightarrow> bool" where
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   153
  [simp]: "subset = op \<subset>"
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   154
34020
2573c794034c merged Crude_Executable_Set into Executable_Set
haftmann
parents: 34008
diff changeset
   155
lemma [code_unfold]:
33947
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   156
  "op \<subset> = subset"
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   157
  by simp
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   158
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   159
setup {*
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   160
  Code.add_signature_cmd ("set_eq", "'a set \<Rightarrow> 'a set \<Rightarrow> bool")
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   161
  #> Code.add_signature_cmd ("subset_eq", "'a set \<Rightarrow> 'a set \<Rightarrow> bool")
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   162
  #> Code.add_signature_cmd ("subset", "'a set \<Rightarrow> 'a set \<Rightarrow> bool")
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   163
*}
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   164
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   165
lemma set_eq_subset_eq [code]:
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   166
  "set_eq A B \<longleftrightarrow> subset_eq A B \<and> subset_eq B A"
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   167
  by auto
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   168
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   169
lemma subset_eq_forall [code]:
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   170
  "subset_eq A B \<longleftrightarrow> (\<forall>x\<in>A. x \<in> B)"
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   171
  by (simp add: subset_eq)
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   172
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   173
lemma subset_subset_eq [code]:
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   174
  "subset A B \<longleftrightarrow> subset_eq A B \<and> \<not> subset_eq B A"
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   175
  by (simp add: subset)
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   176
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   177
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   178
subsection {* Functorial operations *}
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   179
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   180
definition inter :: "'a set \<Rightarrow> 'a set \<Rightarrow> 'a set" where
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   181
  [simp]: "inter = op \<inter>"
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   182
34020
2573c794034c merged Crude_Executable_Set into Executable_Set
haftmann
parents: 34008
diff changeset
   183
lemma [code_unfold]:
33947
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   184
  "op \<inter> = inter"
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   185
  by simp
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   186
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   187
definition subtract :: "'a set \<Rightarrow> 'a set \<Rightarrow> 'a set" where
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   188
  [simp]: "subtract A B = B - A"
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   189
34020
2573c794034c merged Crude_Executable_Set into Executable_Set
haftmann
parents: 34008
diff changeset
   190
lemma [code_unfold]:
33947
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   191
  "B - A = subtract A B"
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   192
  by simp
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   193
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   194
definition union :: "'a set \<Rightarrow> 'a set \<Rightarrow> 'a set" where
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   195
  [simp]: "union = op \<union>"
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   196
34020
2573c794034c merged Crude_Executable_Set into Executable_Set
haftmann
parents: 34008
diff changeset
   197
lemma [code_unfold]:
33947
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   198
  "op \<union> = union"
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   199
  by simp
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   200
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   201
definition Inf :: "'a::complete_lattice set \<Rightarrow> 'a" where
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   202
  [simp]: "Inf = Complete_Lattice.Inf"
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   203
34020
2573c794034c merged Crude_Executable_Set into Executable_Set
haftmann
parents: 34008
diff changeset
   204
lemma [code_unfold]:
33947
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   205
  "Complete_Lattice.Inf = Inf"
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   206
  by simp
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   207
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   208
definition Sup :: "'a::complete_lattice set \<Rightarrow> 'a" where
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   209
  [simp]: "Sup = Complete_Lattice.Sup"
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   210
34020
2573c794034c merged Crude_Executable_Set into Executable_Set
haftmann
parents: 34008
diff changeset
   211
lemma [code_unfold]:
33947
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   212
  "Complete_Lattice.Sup = Sup"
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   213
  by simp
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   214
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   215
definition Inter :: "'a set set \<Rightarrow> 'a set" where
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   216
  [simp]: "Inter = Inf"
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   217
34020
2573c794034c merged Crude_Executable_Set into Executable_Set
haftmann
parents: 34008
diff changeset
   218
lemma [code_unfold]:
33947
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   219
  "Inf = Inter"
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   220
  by simp
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   221
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   222
definition Union :: "'a set set \<Rightarrow> 'a set" where
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   223
  [simp]: "Union = Sup"
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   224
34020
2573c794034c merged Crude_Executable_Set into Executable_Set
haftmann
parents: 34008
diff changeset
   225
lemma [code_unfold]:
33947
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   226
  "Sup = Union"
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   227
  by simp
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   228
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   229
setup {*
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   230
  Code.add_signature_cmd ("inter", "'a set \<Rightarrow> 'a set \<Rightarrow> 'a set")
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   231
  #> Code.add_signature_cmd ("subtract", "'a set \<Rightarrow> 'a set \<Rightarrow> 'a set")
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   232
  #> Code.add_signature_cmd ("union", "'a set \<Rightarrow> 'a set \<Rightarrow> 'a set")
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   233
  #> Code.add_signature_cmd ("Inf", "'a set \<Rightarrow> 'a")
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   234
  #> Code.add_signature_cmd ("Sup", "'a set \<Rightarrow> 'a")
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   235
  #> Code.add_signature_cmd ("Inter", "'a set set \<Rightarrow> 'a set")
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   236
  #> Code.add_signature_cmd ("Union", "'a set set \<Rightarrow> 'a set")
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   237
*}
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   238
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   239
lemma inter_project [code]:
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   240
  "inter A (Set xs) = Set (List.filter (\<lambda>x. x \<in> A) xs)"
37023
efc202e1677e added theory More_List
haftmann
parents: 36176
diff changeset
   241
  "inter A (Coset xs) = foldr remove xs A"
efc202e1677e added theory More_List
haftmann
parents: 36176
diff changeset
   242
  by (simp add: inter project_def) (simp add: Diff_eq [symmetric] minus_set_foldr)
33947
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   243
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   244
lemma subtract_remove [code]:
37023
efc202e1677e added theory More_List
haftmann
parents: 36176
diff changeset
   245
  "subtract (Set xs) A = foldr remove xs A"
33947
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   246
  "subtract (Coset xs) A = Set (List.filter (\<lambda>x. x \<in> A) xs)"
37023
efc202e1677e added theory More_List
haftmann
parents: 36176
diff changeset
   247
  by (auto simp add: minus_set_foldr)
33947
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   248
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   249
lemma union_insert [code]:
37023
efc202e1677e added theory More_List
haftmann
parents: 36176
diff changeset
   250
  "union (Set xs) A = foldr insert xs A"
33947
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   251
  "union (Coset xs) A = Coset (List.filter (\<lambda>x. x \<notin> A) xs)"
37023
efc202e1677e added theory More_List
haftmann
parents: 36176
diff changeset
   252
  by (auto simp add: union_set_foldr)
33947
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   253
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   254
lemma Inf_inf [code]:
37023
efc202e1677e added theory More_List
haftmann
parents: 36176
diff changeset
   255
  "Inf (Set xs) = foldr inf xs (top :: 'a::complete_lattice)"
33947
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   256
  "Inf (Coset []) = (bot :: 'a::complete_lattice)"
37023
efc202e1677e added theory More_List
haftmann
parents: 36176
diff changeset
   257
  by (simp_all add: Inf_UNIV Inf_set_foldr)
33947
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   258
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   259
lemma Sup_sup [code]:
37023
efc202e1677e added theory More_List
haftmann
parents: 36176
diff changeset
   260
  "Sup (Set xs) = foldr sup xs (bot :: 'a::complete_lattice)"
33947
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   261
  "Sup (Coset []) = (top :: 'a::complete_lattice)"
37023
efc202e1677e added theory More_List
haftmann
parents: 36176
diff changeset
   262
  by (simp_all add: Sup_UNIV Sup_set_foldr)
33947
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   263
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   264
lemma Inter_inter [code]:
37023
efc202e1677e added theory More_List
haftmann
parents: 36176
diff changeset
   265
  "Inter (Set xs) = foldr inter xs (Coset [])"
33947
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   266
  "Inter (Coset []) = empty"
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   267
  unfolding Inter_def Inf_inf by simp_all
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   268
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   269
lemma Union_union [code]:
37023
efc202e1677e added theory More_List
haftmann
parents: 36176
diff changeset
   270
  "Union (Set xs) = foldr union xs empty"
33947
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   271
  "Union (Coset []) = Coset []"
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   272
  unfolding Union_def Sup_sup by simp_all
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   273
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
   274
hide_const (open) is_empty empty remove
33947
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   275
  set_eq subset_eq subset inter union subtract Inf Sup Inter Union
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   276
38304
df7d5143db55 executable relation operations contributed by Tjark Weber
haftmann
parents: 37595
diff changeset
   277
df7d5143db55 executable relation operations contributed by Tjark Weber
haftmann
parents: 37595
diff changeset
   278
subsection {* Operations on relations *}
df7d5143db55 executable relation operations contributed by Tjark Weber
haftmann
parents: 37595
diff changeset
   279
df7d5143db55 executable relation operations contributed by Tjark Weber
haftmann
parents: 37595
diff changeset
   280
text {* Initially contributed by Tjark Weber. *}
df7d5143db55 executable relation operations contributed by Tjark Weber
haftmann
parents: 37595
diff changeset
   281
df7d5143db55 executable relation operations contributed by Tjark Weber
haftmann
parents: 37595
diff changeset
   282
lemma bounded_Collect_code [code_unfold]:
df7d5143db55 executable relation operations contributed by Tjark Weber
haftmann
parents: 37595
diff changeset
   283
  "{x\<in>S. P x} = project P S"
df7d5143db55 executable relation operations contributed by Tjark Weber
haftmann
parents: 37595
diff changeset
   284
  by (auto simp add: project_def)
df7d5143db55 executable relation operations contributed by Tjark Weber
haftmann
parents: 37595
diff changeset
   285
df7d5143db55 executable relation operations contributed by Tjark Weber
haftmann
parents: 37595
diff changeset
   286
lemma Id_on_code [code]:
df7d5143db55 executable relation operations contributed by Tjark Weber
haftmann
parents: 37595
diff changeset
   287
  "Id_on (Set xs) = Set [(x,x). x \<leftarrow> xs]"
df7d5143db55 executable relation operations contributed by Tjark Weber
haftmann
parents: 37595
diff changeset
   288
  by (auto simp add: Id_on_def)
df7d5143db55 executable relation operations contributed by Tjark Weber
haftmann
parents: 37595
diff changeset
   289
df7d5143db55 executable relation operations contributed by Tjark Weber
haftmann
parents: 37595
diff changeset
   290
lemma Domain_fst [code]:
df7d5143db55 executable relation operations contributed by Tjark Weber
haftmann
parents: 37595
diff changeset
   291
  "Domain r = fst ` r"
df7d5143db55 executable relation operations contributed by Tjark Weber
haftmann
parents: 37595
diff changeset
   292
  by (auto simp add: image_def Bex_def)
df7d5143db55 executable relation operations contributed by Tjark Weber
haftmann
parents: 37595
diff changeset
   293
df7d5143db55 executable relation operations contributed by Tjark Weber
haftmann
parents: 37595
diff changeset
   294
lemma Range_snd [code]:
df7d5143db55 executable relation operations contributed by Tjark Weber
haftmann
parents: 37595
diff changeset
   295
  "Range r = snd ` r"
df7d5143db55 executable relation operations contributed by Tjark Weber
haftmann
parents: 37595
diff changeset
   296
  by (auto simp add: image_def Bex_def)
df7d5143db55 executable relation operations contributed by Tjark Weber
haftmann
parents: 37595
diff changeset
   297
df7d5143db55 executable relation operations contributed by Tjark Weber
haftmann
parents: 37595
diff changeset
   298
lemma irrefl_code [code]:
df7d5143db55 executable relation operations contributed by Tjark Weber
haftmann
parents: 37595
diff changeset
   299
  "irrefl r \<longleftrightarrow> (\<forall>(x, y)\<in>r. x \<noteq> y)"
df7d5143db55 executable relation operations contributed by Tjark Weber
haftmann
parents: 37595
diff changeset
   300
  by (auto simp add: irrefl_def)
df7d5143db55 executable relation operations contributed by Tjark Weber
haftmann
parents: 37595
diff changeset
   301
df7d5143db55 executable relation operations contributed by Tjark Weber
haftmann
parents: 37595
diff changeset
   302
lemma trans_def [code]:
df7d5143db55 executable relation operations contributed by Tjark Weber
haftmann
parents: 37595
diff changeset
   303
  "trans r \<longleftrightarrow> (\<forall>(x, y1)\<in>r. \<forall>(y2, z)\<in>r. y1 = y2 \<longrightarrow> (x, z)\<in>r)"
df7d5143db55 executable relation operations contributed by Tjark Weber
haftmann
parents: 37595
diff changeset
   304
  by (auto simp add: trans_def)
df7d5143db55 executable relation operations contributed by Tjark Weber
haftmann
parents: 37595
diff changeset
   305
df7d5143db55 executable relation operations contributed by Tjark Weber
haftmann
parents: 37595
diff changeset
   306
definition "exTimes A B = Sigma A (%_. B)"
df7d5143db55 executable relation operations contributed by Tjark Weber
haftmann
parents: 37595
diff changeset
   307
df7d5143db55 executable relation operations contributed by Tjark Weber
haftmann
parents: 37595
diff changeset
   308
lemma [code_unfold]:
df7d5143db55 executable relation operations contributed by Tjark Weber
haftmann
parents: 37595
diff changeset
   309
  "Sigma A (%_. B) = exTimes A B"
df7d5143db55 executable relation operations contributed by Tjark Weber
haftmann
parents: 37595
diff changeset
   310
  by (simp add: exTimes_def)
df7d5143db55 executable relation operations contributed by Tjark Weber
haftmann
parents: 37595
diff changeset
   311
df7d5143db55 executable relation operations contributed by Tjark Weber
haftmann
parents: 37595
diff changeset
   312
lemma exTimes_code [code]:
df7d5143db55 executable relation operations contributed by Tjark Weber
haftmann
parents: 37595
diff changeset
   313
  "exTimes (Set xs) (Set ys) = Set [(x,y). x \<leftarrow> xs, y \<leftarrow> ys]"
df7d5143db55 executable relation operations contributed by Tjark Weber
haftmann
parents: 37595
diff changeset
   314
  by (auto simp add: exTimes_def)
df7d5143db55 executable relation operations contributed by Tjark Weber
haftmann
parents: 37595
diff changeset
   315
df7d5143db55 executable relation operations contributed by Tjark Weber
haftmann
parents: 37595
diff changeset
   316
lemma rel_comp_code [code]:
df7d5143db55 executable relation operations contributed by Tjark Weber
haftmann
parents: 37595
diff changeset
   317
  "(Set xys) O (Set yzs) = Set (remdups [(fst xy, snd yz). xy \<leftarrow> xys, yz \<leftarrow> yzs, snd xy = fst yz])"
df7d5143db55 executable relation operations contributed by Tjark Weber
haftmann
parents: 37595
diff changeset
   318
 by (auto simp add: Bex_def)
df7d5143db55 executable relation operations contributed by Tjark Weber
haftmann
parents: 37595
diff changeset
   319
df7d5143db55 executable relation operations contributed by Tjark Weber
haftmann
parents: 37595
diff changeset
   320
lemma acyclic_code [code]:
df7d5143db55 executable relation operations contributed by Tjark Weber
haftmann
parents: 37595
diff changeset
   321
  "acyclic r = irrefl (r^+)"
df7d5143db55 executable relation operations contributed by Tjark Weber
haftmann
parents: 37595
diff changeset
   322
  by (simp add: acyclic_def irrefl_def)
df7d5143db55 executable relation operations contributed by Tjark Weber
haftmann
parents: 37595
diff changeset
   323
df7d5143db55 executable relation operations contributed by Tjark Weber
haftmann
parents: 37595
diff changeset
   324
lemma wf_code [code]:
df7d5143db55 executable relation operations contributed by Tjark Weber
haftmann
parents: 37595
diff changeset
   325
  "wf (Set xs) = acyclic (Set xs)"
df7d5143db55 executable relation operations contributed by Tjark Weber
haftmann
parents: 37595
diff changeset
   326
  by (simp add: wf_iff_acyclic_if_finite)
df7d5143db55 executable relation operations contributed by Tjark Weber
haftmann
parents: 37595
diff changeset
   327
33947
2d0d08b5b048 added Crude_Executable_Set
haftmann
parents:
diff changeset
   328
end