src/HOL/More_Set.thy
author haftmann
Fri, 06 Jan 2012 20:39:50 +0100
changeset 46143 c932c80d3eae
parent 46133 d9fe85d3d2cd
child 46146 6baea4fca6bd
permissions -rw-r--r--
farewell to theory More_List
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
31807
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
     1
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
     2
(* Author: Florian Haftmann, TU Muenchen *)
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
     3
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
     4
header {* Relating (finite) sets and lists *}
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
     5
37024
e938a0b5286e renamed List_Set to the now more appropriate More_Set
haftmann
parents: 37023
diff changeset
     6
theory More_Set
46143
c932c80d3eae farewell to theory More_List
haftmann
parents: 46133
diff changeset
     7
imports List
31807
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
     8
begin
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
     9
42871
1c0b99f950d9 names of fold_set locales resemble name of characteristic property more closely
haftmann
parents: 42868
diff changeset
    10
lemma comp_fun_idem_remove:
45986
c9e50153e5ae moved various set operations to theory Set (resp. Product_Type)
haftmann
parents: 45974
diff changeset
    11
  "comp_fun_idem Set.remove"
31807
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    12
proof -
45986
c9e50153e5ae moved various set operations to theory Set (resp. Product_Type)
haftmann
parents: 45974
diff changeset
    13
  have rem: "Set.remove = (\<lambda>x A. A - {x})" by (simp add: fun_eq_iff remove_def)
42871
1c0b99f950d9 names of fold_set locales resemble name of characteristic property more closely
haftmann
parents: 42868
diff changeset
    14
  show ?thesis by (simp only: comp_fun_idem_remove rem)
31807
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    15
qed
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    16
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    17
lemma minus_fold_remove:
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    18
  assumes "finite A"
45986
c9e50153e5ae moved various set operations to theory Set (resp. Product_Type)
haftmann
parents: 45974
diff changeset
    19
  shows "B - A = Finite_Set.fold Set.remove B A"
31807
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    20
proof -
45986
c9e50153e5ae moved various set operations to theory Set (resp. Product_Type)
haftmann
parents: 45974
diff changeset
    21
  have rem: "Set.remove = (\<lambda>x A. A - {x})" by (simp add: fun_eq_iff remove_def)
31807
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    22
  show ?thesis by (simp only: rem assms minus_fold_remove)
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    23
qed
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    24
46028
9f113cdf3d66 attribute code_abbrev superseedes code_unfold_post
haftmann
parents: 45990
diff changeset
    25
lemma bounded_Collect_code: (* FIXME delete candidate *)
45986
c9e50153e5ae moved various set operations to theory Set (resp. Product_Type)
haftmann
parents: 45974
diff changeset
    26
  "{x \<in> A. P x} = Set.project P A"
45012
060f76635bfe tuned specification and lemma distribution among theories; tuned proofs
haftmann
parents: 44326
diff changeset
    27
  by (simp add: project_def)
060f76635bfe tuned specification and lemma distribution among theories; tuned proofs
haftmann
parents: 44326
diff changeset
    28
31807
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    29
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    30
subsection {* Basic set operations *}
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    31
46127
af3b95160b59 cleanup of code declarations
haftmann
parents: 46037
diff changeset
    32
lemma is_empty_set [code]:
45986
c9e50153e5ae moved various set operations to theory Set (resp. Product_Type)
haftmann
parents: 45974
diff changeset
    33
  "Set.is_empty (set xs) \<longleftrightarrow> List.null xs"
c9e50153e5ae moved various set operations to theory Set (resp. Product_Type)
haftmann
parents: 45974
diff changeset
    34
  by (simp add: Set.is_empty_def null_def)
31807
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    35
46127
af3b95160b59 cleanup of code declarations
haftmann
parents: 46037
diff changeset
    36
lemma empty_set [code]:
31807
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    37
  "{} = set []"
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    38
  by simp
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    39
32880
b8bee63c7202 sets and cosets
haftmann
parents: 31851
diff changeset
    40
lemma insert_set_compl:
34977
27ceb64d41ea dropped some redundancies
haftmann
parents: 34007
diff changeset
    41
  "insert x (- set xs) = - set (removeAll x xs)"
27ceb64d41ea dropped some redundancies
haftmann
parents: 34007
diff changeset
    42
  by auto
31807
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    43
32880
b8bee63c7202 sets and cosets
haftmann
parents: 31851
diff changeset
    44
lemma remove_set_compl:
45986
c9e50153e5ae moved various set operations to theory Set (resp. Product_Type)
haftmann
parents: 45974
diff changeset
    45
  "Set.remove x (- set xs) = - set (List.insert x xs)"
44326
2b088d74beb3 tuned proof
haftmann
parents: 42871
diff changeset
    46
  by (auto simp add: remove_def List.insert_def)
32880
b8bee63c7202 sets and cosets
haftmann
parents: 31851
diff changeset
    47
31807
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    48
lemma image_set:
31846
89c37daebfdd added Inter, Union
haftmann
parents: 31807
diff changeset
    49
  "image f (set xs) = set (map f xs)"
31807
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    50
  by simp
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    51
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    52
lemma project_set:
45986
c9e50153e5ae moved various set operations to theory Set (resp. Product_Type)
haftmann
parents: 45974
diff changeset
    53
  "Set.project P (set xs) = set (filter P xs)"
31807
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    54
  by (auto simp add: project_def)
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    55
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    56
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    57
subsection {* Functorial set operations *}
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    58
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    59
lemma union_set:
37023
efc202e1677e added theory More_List
haftmann
parents: 34977
diff changeset
    60
  "set xs \<union> A = fold Set.insert xs A"
31807
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    61
proof -
42871
1c0b99f950d9 names of fold_set locales resemble name of characteristic property more closely
haftmann
parents: 42868
diff changeset
    62
  interpret comp_fun_idem Set.insert
1c0b99f950d9 names of fold_set locales resemble name of characteristic property more closely
haftmann
parents: 42868
diff changeset
    63
    by (fact comp_fun_idem_insert)
45990
b7b905b23b2a incorporated More_Set and More_List into the Main body -- to be consolidated later
haftmann
parents: 45986
diff changeset
    64
  show ?thesis by (simp add: union_fold_insert fold_set_fold)
31807
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    65
qed
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    66
37023
efc202e1677e added theory More_List
haftmann
parents: 34977
diff changeset
    67
lemma union_set_foldr:
efc202e1677e added theory More_List
haftmann
parents: 34977
diff changeset
    68
  "set xs \<union> A = foldr Set.insert xs A"
efc202e1677e added theory More_List
haftmann
parents: 34977
diff changeset
    69
proof -
efc202e1677e added theory More_List
haftmann
parents: 34977
diff changeset
    70
  have "\<And>x y :: 'a. insert y \<circ> insert x = insert x \<circ> insert y"
45012
060f76635bfe tuned specification and lemma distribution among theories; tuned proofs
haftmann
parents: 44326
diff changeset
    71
    by auto
37023
efc202e1677e added theory More_List
haftmann
parents: 34977
diff changeset
    72
  then show ?thesis by (simp add: union_set foldr_fold)
efc202e1677e added theory More_List
haftmann
parents: 34977
diff changeset
    73
qed
efc202e1677e added theory More_List
haftmann
parents: 34977
diff changeset
    74
31807
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    75
lemma minus_set:
45986
c9e50153e5ae moved various set operations to theory Set (resp. Product_Type)
haftmann
parents: 45974
diff changeset
    76
  "A - set xs = fold Set.remove xs A"
31807
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    77
proof -
45986
c9e50153e5ae moved various set operations to theory Set (resp. Product_Type)
haftmann
parents: 45974
diff changeset
    78
  interpret comp_fun_idem Set.remove
42871
1c0b99f950d9 names of fold_set locales resemble name of characteristic property more closely
haftmann
parents: 42868
diff changeset
    79
    by (fact comp_fun_idem_remove)
31807
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    80
  show ?thesis
45990
b7b905b23b2a incorporated More_Set and More_List into the Main body -- to be consolidated later
haftmann
parents: 45986
diff changeset
    81
    by (simp add: minus_fold_remove [of _ A] fold_set_fold)
31807
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    82
qed
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    83
37023
efc202e1677e added theory More_List
haftmann
parents: 34977
diff changeset
    84
lemma minus_set_foldr:
45986
c9e50153e5ae moved various set operations to theory Set (resp. Product_Type)
haftmann
parents: 45974
diff changeset
    85
  "A - set xs = foldr Set.remove xs A"
37023
efc202e1677e added theory More_List
haftmann
parents: 34977
diff changeset
    86
proof -
45986
c9e50153e5ae moved various set operations to theory Set (resp. Product_Type)
haftmann
parents: 45974
diff changeset
    87
  have "\<And>x y :: 'a. Set.remove y \<circ> Set.remove x = Set.remove x \<circ> Set.remove y"
45012
060f76635bfe tuned specification and lemma distribution among theories; tuned proofs
haftmann
parents: 44326
diff changeset
    88
    by (auto simp add: remove_def)
37023
efc202e1677e added theory More_List
haftmann
parents: 34977
diff changeset
    89
  then show ?thesis by (simp add: minus_set foldr_fold)
efc202e1677e added theory More_List
haftmann
parents: 34977
diff changeset
    90
qed
efc202e1677e added theory More_List
haftmann
parents: 34977
diff changeset
    91
31807
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    92
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    93
subsection {* Derived set operations *}
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    94
46127
af3b95160b59 cleanup of code declarations
haftmann
parents: 46037
diff changeset
    95
lemma member [code]:
31807
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    96
  "a \<in> A \<longleftrightarrow> (\<exists>x\<in>A. a = x)"
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    97
  by simp
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    98
46127
af3b95160b59 cleanup of code declarations
haftmann
parents: 46037
diff changeset
    99
lemma subset [code]:
31807
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   100
  "A \<subset> B \<longleftrightarrow> A \<subseteq> B \<and> \<not> B \<subseteq> A"
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   101
  by (fact less_le_not_le)
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   102
46127
af3b95160b59 cleanup of code declarations
haftmann
parents: 46037
diff changeset
   103
lemma set_eq [code]:
31807
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   104
  "A = B \<longleftrightarrow> A \<subseteq> B \<and> B \<subseteq> A"
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   105
  by (fact eq_iff)
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   106
46127
af3b95160b59 cleanup of code declarations
haftmann
parents: 46037
diff changeset
   107
lemma inter [code]:
45986
c9e50153e5ae moved various set operations to theory Set (resp. Product_Type)
haftmann
parents: 45974
diff changeset
   108
  "A \<inter> B = Set.project (\<lambda>x. x \<in> A) B"
31807
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   109
  by (auto simp add: project_def)
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   110
37023
efc202e1677e added theory More_List
haftmann
parents: 34977
diff changeset
   111
45974
2b043ed911ac added setup for executable code
haftmann
parents: 45012
diff changeset
   112
subsection {* Code generator setup *}
2b043ed911ac added setup for executable code
haftmann
parents: 45012
diff changeset
   113
2b043ed911ac added setup for executable code
haftmann
parents: 45012
diff changeset
   114
definition coset :: "'a list \<Rightarrow> 'a set" where
2b043ed911ac added setup for executable code
haftmann
parents: 45012
diff changeset
   115
  [simp]: "coset xs = - set xs"
2b043ed911ac added setup for executable code
haftmann
parents: 45012
diff changeset
   116
2b043ed911ac added setup for executable code
haftmann
parents: 45012
diff changeset
   117
code_datatype set coset
2b043ed911ac added setup for executable code
haftmann
parents: 45012
diff changeset
   118
2b043ed911ac added setup for executable code
haftmann
parents: 45012
diff changeset
   119
2b043ed911ac added setup for executable code
haftmann
parents: 45012
diff changeset
   120
subsection {* Basic operations *}
2b043ed911ac added setup for executable code
haftmann
parents: 45012
diff changeset
   121
2b043ed911ac added setup for executable code
haftmann
parents: 45012
diff changeset
   122
lemma [code]:
2b043ed911ac added setup for executable code
haftmann
parents: 45012
diff changeset
   123
  "x \<in> set xs \<longleftrightarrow> List.member xs x"
2b043ed911ac added setup for executable code
haftmann
parents: 45012
diff changeset
   124
  "x \<in> coset xs \<longleftrightarrow> \<not> List.member xs x"
2b043ed911ac added setup for executable code
haftmann
parents: 45012
diff changeset
   125
  by (simp_all add: member_def)
2b043ed911ac added setup for executable code
haftmann
parents: 45012
diff changeset
   126
2b043ed911ac added setup for executable code
haftmann
parents: 45012
diff changeset
   127
lemma UNIV_coset [code]:
2b043ed911ac added setup for executable code
haftmann
parents: 45012
diff changeset
   128
  "UNIV = coset []"
2b043ed911ac added setup for executable code
haftmann
parents: 45012
diff changeset
   129
  by simp
2b043ed911ac added setup for executable code
haftmann
parents: 45012
diff changeset
   130
2b043ed911ac added setup for executable code
haftmann
parents: 45012
diff changeset
   131
lemma insert_code [code]:
2b043ed911ac added setup for executable code
haftmann
parents: 45012
diff changeset
   132
  "insert x (set xs) = set (List.insert x xs)"
2b043ed911ac added setup for executable code
haftmann
parents: 45012
diff changeset
   133
  "insert x (coset xs) = coset (removeAll x xs)"
2b043ed911ac added setup for executable code
haftmann
parents: 45012
diff changeset
   134
  by simp_all
2b043ed911ac added setup for executable code
haftmann
parents: 45012
diff changeset
   135
2b043ed911ac added setup for executable code
haftmann
parents: 45012
diff changeset
   136
lemma remove_code [code]:
45986
c9e50153e5ae moved various set operations to theory Set (resp. Product_Type)
haftmann
parents: 45974
diff changeset
   137
  "Set.remove x (set xs) = set (removeAll x xs)"
c9e50153e5ae moved various set operations to theory Set (resp. Product_Type)
haftmann
parents: 45974
diff changeset
   138
  "Set.remove x (coset xs) = coset (List.insert x xs)"
45974
2b043ed911ac added setup for executable code
haftmann
parents: 45012
diff changeset
   139
  by (simp_all add: remove_def Compl_insert)
2b043ed911ac added setup for executable code
haftmann
parents: 45012
diff changeset
   140
2b043ed911ac added setup for executable code
haftmann
parents: 45012
diff changeset
   141
lemma Ball_set [code]:
2b043ed911ac added setup for executable code
haftmann
parents: 45012
diff changeset
   142
  "Ball (set xs) P \<longleftrightarrow> list_all P xs"
2b043ed911ac added setup for executable code
haftmann
parents: 45012
diff changeset
   143
  by (simp add: list_all_iff)
2b043ed911ac added setup for executable code
haftmann
parents: 45012
diff changeset
   144
2b043ed911ac added setup for executable code
haftmann
parents: 45012
diff changeset
   145
lemma Bex_set [code]:
2b043ed911ac added setup for executable code
haftmann
parents: 45012
diff changeset
   146
  "Bex (set xs) P \<longleftrightarrow> list_ex P xs"
2b043ed911ac added setup for executable code
haftmann
parents: 45012
diff changeset
   147
  by (simp add: list_ex_iff)
2b043ed911ac added setup for executable code
haftmann
parents: 45012
diff changeset
   148
2b043ed911ac added setup for executable code
haftmann
parents: 45012
diff changeset
   149
lemma card_set [code]:
2b043ed911ac added setup for executable code
haftmann
parents: 45012
diff changeset
   150
  "card (set xs) = length (remdups xs)"
2b043ed911ac added setup for executable code
haftmann
parents: 45012
diff changeset
   151
proof -
2b043ed911ac added setup for executable code
haftmann
parents: 45012
diff changeset
   152
  have "card (set (remdups xs)) = length (remdups xs)"
2b043ed911ac added setup for executable code
haftmann
parents: 45012
diff changeset
   153
    by (rule distinct_card) simp
2b043ed911ac added setup for executable code
haftmann
parents: 45012
diff changeset
   154
  then show ?thesis by simp
2b043ed911ac added setup for executable code
haftmann
parents: 45012
diff changeset
   155
qed
2b043ed911ac added setup for executable code
haftmann
parents: 45012
diff changeset
   156
2b043ed911ac added setup for executable code
haftmann
parents: 45012
diff changeset
   157
2b043ed911ac added setup for executable code
haftmann
parents: 45012
diff changeset
   158
subsection {* Functorial operations *}
2b043ed911ac added setup for executable code
haftmann
parents: 45012
diff changeset
   159
2b043ed911ac added setup for executable code
haftmann
parents: 45012
diff changeset
   160
lemma inter_code [code]:
2b043ed911ac added setup for executable code
haftmann
parents: 45012
diff changeset
   161
  "A \<inter> set xs = set (List.filter (\<lambda>x. x \<in> A) xs)"
45986
c9e50153e5ae moved various set operations to theory Set (resp. Product_Type)
haftmann
parents: 45974
diff changeset
   162
  "A \<inter> coset xs = foldr Set.remove xs A"
45974
2b043ed911ac added setup for executable code
haftmann
parents: 45012
diff changeset
   163
  by (simp add: inter project_def) (simp add: Diff_eq [symmetric] minus_set_foldr)
2b043ed911ac added setup for executable code
haftmann
parents: 45012
diff changeset
   164
2b043ed911ac added setup for executable code
haftmann
parents: 45012
diff changeset
   165
lemma subtract_code [code]:
45986
c9e50153e5ae moved various set operations to theory Set (resp. Product_Type)
haftmann
parents: 45974
diff changeset
   166
  "A - set xs = foldr Set.remove xs A"
45974
2b043ed911ac added setup for executable code
haftmann
parents: 45012
diff changeset
   167
  "A - coset xs = set (List.filter (\<lambda>x. x \<in> A) xs)"
2b043ed911ac added setup for executable code
haftmann
parents: 45012
diff changeset
   168
  by (auto simp add: minus_set_foldr)
2b043ed911ac added setup for executable code
haftmann
parents: 45012
diff changeset
   169
2b043ed911ac added setup for executable code
haftmann
parents: 45012
diff changeset
   170
lemma union_code [code]:
2b043ed911ac added setup for executable code
haftmann
parents: 45012
diff changeset
   171
  "set xs \<union> A = foldr insert xs A"
2b043ed911ac added setup for executable code
haftmann
parents: 45012
diff changeset
   172
  "coset xs \<union> A = coset (List.filter (\<lambda>x. x \<notin> A) xs)"
2b043ed911ac added setup for executable code
haftmann
parents: 45012
diff changeset
   173
  by (auto simp add: union_set_foldr)
2b043ed911ac added setup for executable code
haftmann
parents: 45012
diff changeset
   174
2b043ed911ac added setup for executable code
haftmann
parents: 45012
diff changeset
   175
definition Inf :: "'a::complete_lattice set \<Rightarrow> 'a" where
46028
9f113cdf3d66 attribute code_abbrev superseedes code_unfold_post
haftmann
parents: 45990
diff changeset
   176
  [simp, code_abbrev]: "Inf = Complete_Lattices.Inf"
45974
2b043ed911ac added setup for executable code
haftmann
parents: 45012
diff changeset
   177
2b043ed911ac added setup for executable code
haftmann
parents: 45012
diff changeset
   178
hide_const (open) Inf
2b043ed911ac added setup for executable code
haftmann
parents: 45012
diff changeset
   179
2b043ed911ac added setup for executable code
haftmann
parents: 45012
diff changeset
   180
definition Sup :: "'a::complete_lattice set \<Rightarrow> 'a" where
46028
9f113cdf3d66 attribute code_abbrev superseedes code_unfold_post
haftmann
parents: 45990
diff changeset
   181
  [simp, code_abbrev]: "Sup = Complete_Lattices.Sup"
45974
2b043ed911ac added setup for executable code
haftmann
parents: 45012
diff changeset
   182
2b043ed911ac added setup for executable code
haftmann
parents: 45012
diff changeset
   183
hide_const (open) Sup
2b043ed911ac added setup for executable code
haftmann
parents: 45012
diff changeset
   184
2b043ed911ac added setup for executable code
haftmann
parents: 45012
diff changeset
   185
lemma Inf_code [code]:
2b043ed911ac added setup for executable code
haftmann
parents: 45012
diff changeset
   186
  "More_Set.Inf (set xs) = foldr inf xs top"
2b043ed911ac added setup for executable code
haftmann
parents: 45012
diff changeset
   187
  "More_Set.Inf (coset []) = bot"
2b043ed911ac added setup for executable code
haftmann
parents: 45012
diff changeset
   188
  by (simp_all add: Inf_set_foldr)
2b043ed911ac added setup for executable code
haftmann
parents: 45012
diff changeset
   189
2b043ed911ac added setup for executable code
haftmann
parents: 45012
diff changeset
   190
lemma Sup_sup [code]:
2b043ed911ac added setup for executable code
haftmann
parents: 45012
diff changeset
   191
  "More_Set.Sup (set xs) = foldr sup xs bot"
2b043ed911ac added setup for executable code
haftmann
parents: 45012
diff changeset
   192
  "More_Set.Sup (coset []) = top"
2b043ed911ac added setup for executable code
haftmann
parents: 45012
diff changeset
   193
  by (simp_all add: Sup_set_foldr)
2b043ed911ac added setup for executable code
haftmann
parents: 45012
diff changeset
   194
46037
49a436ada6a2 added implementation of pred_of_set
haftmann
parents: 46028
diff changeset
   195
(* FIXME: better implement conversion by bisection *)
49a436ada6a2 added implementation of pred_of_set
haftmann
parents: 46028
diff changeset
   196
49a436ada6a2 added implementation of pred_of_set
haftmann
parents: 46028
diff changeset
   197
lemma pred_of_set_fold_sup:
49a436ada6a2 added implementation of pred_of_set
haftmann
parents: 46028
diff changeset
   198
  assumes "finite A"
49a436ada6a2 added implementation of pred_of_set
haftmann
parents: 46028
diff changeset
   199
  shows "pred_of_set A = Finite_Set.fold sup bot (Predicate.single ` A)" (is "?lhs = ?rhs")
49a436ada6a2 added implementation of pred_of_set
haftmann
parents: 46028
diff changeset
   200
proof (rule sym)
49a436ada6a2 added implementation of pred_of_set
haftmann
parents: 46028
diff changeset
   201
  interpret comp_fun_idem "sup :: 'a Predicate.pred \<Rightarrow> 'a Predicate.pred \<Rightarrow> 'a Predicate.pred"
49a436ada6a2 added implementation of pred_of_set
haftmann
parents: 46028
diff changeset
   202
    by (fact comp_fun_idem_sup)
49a436ada6a2 added implementation of pred_of_set
haftmann
parents: 46028
diff changeset
   203
  from `finite A` show "?rhs = ?lhs" by (induct A) (auto intro!: pred_eqI)
49a436ada6a2 added implementation of pred_of_set
haftmann
parents: 46028
diff changeset
   204
qed
49a436ada6a2 added implementation of pred_of_set
haftmann
parents: 46028
diff changeset
   205
49a436ada6a2 added implementation of pred_of_set
haftmann
parents: 46028
diff changeset
   206
lemma pred_of_set_set_fold_sup:
49a436ada6a2 added implementation of pred_of_set
haftmann
parents: 46028
diff changeset
   207
  "pred_of_set (set xs) = fold sup (map Predicate.single xs) bot"
49a436ada6a2 added implementation of pred_of_set
haftmann
parents: 46028
diff changeset
   208
proof -
49a436ada6a2 added implementation of pred_of_set
haftmann
parents: 46028
diff changeset
   209
  interpret comp_fun_idem "sup :: 'a Predicate.pred \<Rightarrow> 'a Predicate.pred \<Rightarrow> 'a Predicate.pred"
49a436ada6a2 added implementation of pred_of_set
haftmann
parents: 46028
diff changeset
   210
    by (fact comp_fun_idem_sup)
49a436ada6a2 added implementation of pred_of_set
haftmann
parents: 46028
diff changeset
   211
  show ?thesis by (simp add: pred_of_set_fold_sup fold_set_fold [symmetric])
49a436ada6a2 added implementation of pred_of_set
haftmann
parents: 46028
diff changeset
   212
qed
49a436ada6a2 added implementation of pred_of_set
haftmann
parents: 46028
diff changeset
   213
49a436ada6a2 added implementation of pred_of_set
haftmann
parents: 46028
diff changeset
   214
lemma pred_of_set_set_foldr_sup [code]:
49a436ada6a2 added implementation of pred_of_set
haftmann
parents: 46028
diff changeset
   215
  "pred_of_set (set xs) = foldr sup (map Predicate.single xs) bot"
49a436ada6a2 added implementation of pred_of_set
haftmann
parents: 46028
diff changeset
   216
  by (simp add: pred_of_set_set_fold_sup ac_simps foldr_fold fun_eq_iff)
49a436ada6a2 added implementation of pred_of_set
haftmann
parents: 46028
diff changeset
   217
45974
2b043ed911ac added setup for executable code
haftmann
parents: 45012
diff changeset
   218
2b043ed911ac added setup for executable code
haftmann
parents: 45012
diff changeset
   219
subsection {* Operations on relations *}
2b043ed911ac added setup for executable code
haftmann
parents: 45012
diff changeset
   220
46127
af3b95160b59 cleanup of code declarations
haftmann
parents: 46037
diff changeset
   221
lemma product_code [code]:
af3b95160b59 cleanup of code declarations
haftmann
parents: 46037
diff changeset
   222
  "Product_Type.product (set xs) (set ys) = set [(x, y). x \<leftarrow> xs, y \<leftarrow> ys]"
af3b95160b59 cleanup of code declarations
haftmann
parents: 46037
diff changeset
   223
  by (auto simp add: Product_Type.product_def)
45974
2b043ed911ac added setup for executable code
haftmann
parents: 45012
diff changeset
   224
46127
af3b95160b59 cleanup of code declarations
haftmann
parents: 46037
diff changeset
   225
lemma Id_on_set [code]:
af3b95160b59 cleanup of code declarations
haftmann
parents: 46037
diff changeset
   226
  "Id_on (set xs) = set [(x, x). x \<leftarrow> xs]"
af3b95160b59 cleanup of code declarations
haftmann
parents: 46037
diff changeset
   227
  by (auto simp add: Id_on_def)
45974
2b043ed911ac added setup for executable code
haftmann
parents: 45012
diff changeset
   228
46127
af3b95160b59 cleanup of code declarations
haftmann
parents: 46037
diff changeset
   229
lemma trancl_set_ntrancl [code]: "trancl (set xs) = ntrancl (card (set xs) - 1) (set xs)"
af3b95160b59 cleanup of code declarations
haftmann
parents: 46037
diff changeset
   230
  by (simp add: finite_trancl_ntranl)
45974
2b043ed911ac added setup for executable code
haftmann
parents: 45012
diff changeset
   231
46127
af3b95160b59 cleanup of code declarations
haftmann
parents: 46037
diff changeset
   232
lemma set_rel_comp [code]:
af3b95160b59 cleanup of code declarations
haftmann
parents: 46037
diff changeset
   233
  "set xys O set yzs = set ([(fst xy, snd yz). xy \<leftarrow> xys, yz \<leftarrow> yzs, snd xy = fst yz])"
af3b95160b59 cleanup of code declarations
haftmann
parents: 46037
diff changeset
   234
  by (auto simp add: Bex_def)
45974
2b043ed911ac added setup for executable code
haftmann
parents: 45012
diff changeset
   235
46127
af3b95160b59 cleanup of code declarations
haftmann
parents: 46037
diff changeset
   236
lemma wf_set [code]:
af3b95160b59 cleanup of code declarations
haftmann
parents: 46037
diff changeset
   237
  "wf (set xs) = acyclic (set xs)"
af3b95160b59 cleanup of code declarations
haftmann
parents: 46037
diff changeset
   238
  by (simp add: wf_iff_acyclic_if_finite)
45974
2b043ed911ac added setup for executable code
haftmann
parents: 45012
diff changeset
   239
2b043ed911ac added setup for executable code
haftmann
parents: 45012
diff changeset
   240
end
2b043ed911ac added setup for executable code
haftmann
parents: 45012
diff changeset
   241