src/HOL/Library/Cset.thy
author wenzelm
Tue, 05 Apr 2011 14:25:18 +0200
changeset 42224 578a51fae383
parent 41505 6d19301074cf
child 43241 93b1183e43e5
permissions -rw-r--r--
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
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
40672
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
     4
header {* A dedicated set type which is executable on its finite part *}
31807
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
     5
40672
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
     6
theory Cset
37024
e938a0b5286e renamed List_Set to the now more appropriate More_Set
haftmann
parents: 37023
diff changeset
     7
imports More_Set More_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
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    10
subsection {* Lifting *}
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    11
40672
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
    12
typedef (open) 'a set = "UNIV :: 'a set set"
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
    13
  morphisms member Set by rule+
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
    14
hide_type (open) set
31807
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    15
40672
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
    16
lemma member_Set [simp]:
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
    17
  "member (Set A) = A"
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
    18
  by (rule Set_inverse) rule
34048
369509057220 using existing lattice classes
haftmann
parents: 33939
diff changeset
    19
40672
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
    20
lemma Set_member [simp]:
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
    21
  "Set (member A) = A"
37699
f110a9fa8766 tuned proof
haftmann
parents: 37595
diff changeset
    22
  by (fact member_inverse)
37468
a2a3b62fc819 quickcheck for fsets
haftmann
parents: 37024
diff changeset
    23
40672
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
    24
lemma Set_inject [simp]:
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
    25
  "Set A = Set B \<longleftrightarrow> A = B"
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
    26
  by (simp add: Set_inject)
37468
a2a3b62fc819 quickcheck for fsets
haftmann
parents: 37024
diff changeset
    27
40672
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
    28
lemma set_eq_iff:
39380
5a2662c1e44a established emerging canonical names *_eqI and *_eq_iff
haftmann
parents: 39302
diff changeset
    29
  "A = B \<longleftrightarrow> member A = member B"
5a2662c1e44a established emerging canonical names *_eqI and *_eq_iff
haftmann
parents: 39302
diff changeset
    30
  by (simp add: member_inject)
40672
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
    31
hide_fact (open) set_eq_iff
39380
5a2662c1e44a established emerging canonical names *_eqI and *_eq_iff
haftmann
parents: 39302
diff changeset
    32
40672
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
    33
lemma set_eqI:
37473
013f78aed840 extensionality rule fset_eqI
haftmann
parents: 37468
diff changeset
    34
  "member A = member B \<Longrightarrow> A = B"
40672
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
    35
  by (simp add: Cset.set_eq_iff)
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
    36
hide_fact (open) set_eqI
37473
013f78aed840 extensionality rule fset_eqI
haftmann
parents: 37468
diff changeset
    37
37468
a2a3b62fc819 quickcheck for fsets
haftmann
parents: 37024
diff changeset
    38
declare mem_def [simp]
31807
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    39
40672
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
    40
definition set :: "'a list \<Rightarrow> 'a Cset.set" where
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
    41
  "set xs = Set (List.set xs)"
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
    42
hide_const (open) set
31807
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    43
40672
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
    44
lemma member_set [simp]:
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
    45
  "member (Cset.set xs) = set xs"
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
    46
  by (simp add: set_def)
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
    47
hide_fact (open) member_set
31807
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    48
40672
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
    49
definition coset :: "'a list \<Rightarrow> 'a Cset.set" where
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
    50
  "coset xs = Set (- set xs)"
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
    51
hide_const (open) coset
32880
b8bee63c7202 sets and cosets
haftmann
parents: 32139
diff changeset
    52
40672
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
    53
lemma member_coset [simp]:
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
    54
  "member (Cset.coset xs) = - set xs"
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
    55
  by (simp add: coset_def)
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
    56
hide_fact (open) member_coset
32880
b8bee63c7202 sets and cosets
haftmann
parents: 32139
diff changeset
    57
40672
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
    58
code_datatype Cset.set Cset.coset
32880
b8bee63c7202 sets and cosets
haftmann
parents: 32139
diff changeset
    59
b8bee63c7202 sets and cosets
haftmann
parents: 32139
diff changeset
    60
lemma member_code [code]:
40672
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
    61
  "member (Cset.set xs) = List.member xs"
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
    62
  "member (Cset.coset xs) = Not \<circ> List.member xs"
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 39200
diff changeset
    63
  by (simp_all add: fun_eq_iff member_def fun_Compl_def bool_Compl_def)
32880
b8bee63c7202 sets and cosets
haftmann
parents: 32139
diff changeset
    64
b8bee63c7202 sets and cosets
haftmann
parents: 32139
diff changeset
    65
lemma member_image_UNIV [simp]:
b8bee63c7202 sets and cosets
haftmann
parents: 32139
diff changeset
    66
  "member ` UNIV = UNIV"
b8bee63c7202 sets and cosets
haftmann
parents: 32139
diff changeset
    67
proof -
40672
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
    68
  have "\<And>A \<Colon> 'a set. \<exists>B \<Colon> 'a Cset.set. A = member B"
32880
b8bee63c7202 sets and cosets
haftmann
parents: 32139
diff changeset
    69
  proof
b8bee63c7202 sets and cosets
haftmann
parents: 32139
diff changeset
    70
    fix A :: "'a set"
40672
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
    71
    show "A = member (Set A)" by simp
32880
b8bee63c7202 sets and cosets
haftmann
parents: 32139
diff changeset
    72
  qed
b8bee63c7202 sets and cosets
haftmann
parents: 32139
diff changeset
    73
  then show ?thesis by (simp add: image_def)
b8bee63c7202 sets and cosets
haftmann
parents: 32139
diff changeset
    74
qed
31807
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    75
37468
a2a3b62fc819 quickcheck for fsets
haftmann
parents: 37024
diff changeset
    76
definition (in term_syntax)
a2a3b62fc819 quickcheck for fsets
haftmann
parents: 37024
diff changeset
    77
  setify :: "'a\<Colon>typerep list \<times> (unit \<Rightarrow> Code_Evaluation.term)
40672
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
    78
    \<Rightarrow> 'a Cset.set \<times> (unit \<Rightarrow> Code_Evaluation.term)" where
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
    79
  [code_unfold]: "setify xs = Code_Evaluation.valtermify Cset.set {\<cdot>} xs"
37468
a2a3b62fc819 quickcheck for fsets
haftmann
parents: 37024
diff changeset
    80
37751
89e16802b6cc nicer xsymbol syntax for fcomp and scomp
haftmann
parents: 37699
diff changeset
    81
notation fcomp (infixl "\<circ>>" 60)
89e16802b6cc nicer xsymbol syntax for fcomp and scomp
haftmann
parents: 37699
diff changeset
    82
notation scomp (infixl "\<circ>\<rightarrow>" 60)
37468
a2a3b62fc819 quickcheck for fsets
haftmann
parents: 37024
diff changeset
    83
40672
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
    84
instantiation Cset.set :: (random) random
37468
a2a3b62fc819 quickcheck for fsets
haftmann
parents: 37024
diff changeset
    85
begin
a2a3b62fc819 quickcheck for fsets
haftmann
parents: 37024
diff changeset
    86
a2a3b62fc819 quickcheck for fsets
haftmann
parents: 37024
diff changeset
    87
definition
37751
89e16802b6cc nicer xsymbol syntax for fcomp and scomp
haftmann
parents: 37699
diff changeset
    88
  "Quickcheck.random i = Quickcheck.random i \<circ>\<rightarrow> (\<lambda>xs. Pair (setify xs))"
37468
a2a3b62fc819 quickcheck for fsets
haftmann
parents: 37024
diff changeset
    89
a2a3b62fc819 quickcheck for fsets
haftmann
parents: 37024
diff changeset
    90
instance ..
a2a3b62fc819 quickcheck for fsets
haftmann
parents: 37024
diff changeset
    91
a2a3b62fc819 quickcheck for fsets
haftmann
parents: 37024
diff changeset
    92
end
a2a3b62fc819 quickcheck for fsets
haftmann
parents: 37024
diff changeset
    93
37751
89e16802b6cc nicer xsymbol syntax for fcomp and scomp
haftmann
parents: 37699
diff changeset
    94
no_notation fcomp (infixl "\<circ>>" 60)
89e16802b6cc nicer xsymbol syntax for fcomp and scomp
haftmann
parents: 37699
diff changeset
    95
no_notation scomp (infixl "\<circ>\<rightarrow>" 60)
37468
a2a3b62fc819 quickcheck for fsets
haftmann
parents: 37024
diff changeset
    96
31807
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
    97
34048
369509057220 using existing lattice classes
haftmann
parents: 33939
diff changeset
    98
subsection {* Lattice instantiation *}
369509057220 using existing lattice classes
haftmann
parents: 33939
diff changeset
    99
40672
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   100
instantiation Cset.set :: (type) boolean_algebra
34048
369509057220 using existing lattice classes
haftmann
parents: 33939
diff changeset
   101
begin
369509057220 using existing lattice classes
haftmann
parents: 33939
diff changeset
   102
40672
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   103
definition less_eq_set :: "'a Cset.set \<Rightarrow> 'a Cset.set \<Rightarrow> bool" where
34048
369509057220 using existing lattice classes
haftmann
parents: 33939
diff changeset
   104
  [simp]: "A \<le> B \<longleftrightarrow> member A \<subseteq> member B"
369509057220 using existing lattice classes
haftmann
parents: 33939
diff changeset
   105
40672
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   106
definition less_set :: "'a Cset.set \<Rightarrow> 'a Cset.set \<Rightarrow> bool" where
34048
369509057220 using existing lattice classes
haftmann
parents: 33939
diff changeset
   107
  [simp]: "A < B \<longleftrightarrow> member A \<subset> member B"
369509057220 using existing lattice classes
haftmann
parents: 33939
diff changeset
   108
40672
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   109
definition inf_set :: "'a Cset.set \<Rightarrow> 'a Cset.set \<Rightarrow> 'a Cset.set" where
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   110
  [simp]: "inf A B = Set (member A \<inter> member B)"
34048
369509057220 using existing lattice classes
haftmann
parents: 33939
diff changeset
   111
40672
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   112
definition sup_set :: "'a Cset.set \<Rightarrow> 'a Cset.set \<Rightarrow> 'a Cset.set" where
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   113
  [simp]: "sup A B = Set (member A \<union> member B)"
34048
369509057220 using existing lattice classes
haftmann
parents: 33939
diff changeset
   114
40672
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   115
definition bot_set :: "'a Cset.set" where
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   116
  [simp]: "bot = Set {}"
34048
369509057220 using existing lattice classes
haftmann
parents: 33939
diff changeset
   117
40672
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   118
definition top_set :: "'a Cset.set" where
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   119
  [simp]: "top = Set UNIV"
34048
369509057220 using existing lattice classes
haftmann
parents: 33939
diff changeset
   120
40672
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   121
definition uminus_set :: "'a Cset.set \<Rightarrow> 'a Cset.set" where
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   122
  [simp]: "- A = Set (- (member A))"
34048
369509057220 using existing lattice classes
haftmann
parents: 33939
diff changeset
   123
40672
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   124
definition minus_set :: "'a Cset.set \<Rightarrow> 'a Cset.set \<Rightarrow> 'a Cset.set" where
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   125
  [simp]: "A - B = Set (member A - member B)"
34048
369509057220 using existing lattice classes
haftmann
parents: 33939
diff changeset
   126
369509057220 using existing lattice classes
haftmann
parents: 33939
diff changeset
   127
instance proof
40672
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   128
qed (auto intro: Cset.set_eqI)
34048
369509057220 using existing lattice classes
haftmann
parents: 33939
diff changeset
   129
369509057220 using existing lattice classes
haftmann
parents: 33939
diff changeset
   130
end
369509057220 using existing lattice classes
haftmann
parents: 33939
diff changeset
   131
40672
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   132
instantiation Cset.set :: (type) complete_lattice
34048
369509057220 using existing lattice classes
haftmann
parents: 33939
diff changeset
   133
begin
369509057220 using existing lattice classes
haftmann
parents: 33939
diff changeset
   134
40672
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   135
definition Inf_set :: "'a Cset.set set \<Rightarrow> 'a Cset.set" where
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   136
  [simp]: "Inf_set As = Set (Inf (image member As))"
34048
369509057220 using existing lattice classes
haftmann
parents: 33939
diff changeset
   137
40672
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   138
definition Sup_set :: "'a Cset.set set \<Rightarrow> 'a Cset.set" where
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   139
  [simp]: "Sup_set As = Set (Sup (image member As))"
34048
369509057220 using existing lattice classes
haftmann
parents: 33939
diff changeset
   140
369509057220 using existing lattice classes
haftmann
parents: 33939
diff changeset
   141
instance proof
369509057220 using existing lattice classes
haftmann
parents: 33939
diff changeset
   142
qed (auto simp add: le_fun_def le_bool_def)
369509057220 using existing lattice classes
haftmann
parents: 33939
diff changeset
   143
369509057220 using existing lattice classes
haftmann
parents: 33939
diff changeset
   144
end
369509057220 using existing lattice classes
haftmann
parents: 33939
diff changeset
   145
37023
efc202e1677e added theory More_List
haftmann
parents: 36176
diff changeset
   146
31807
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   147
subsection {* Basic operations *}
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   148
40672
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   149
definition is_empty :: "'a Cset.set \<Rightarrow> bool" where
37024
e938a0b5286e renamed List_Set to the now more appropriate More_Set
haftmann
parents: 37023
diff changeset
   150
  [simp]: "is_empty A \<longleftrightarrow> More_Set.is_empty (member A)"
31807
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   151
40672
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   152
lemma is_empty_set [code]:
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   153
  "is_empty (Cset.set xs) \<longleftrightarrow> List.null xs"
31846
89c37daebfdd added Inter, Union
haftmann
parents: 31807
diff changeset
   154
  by (simp add: is_empty_set)
40672
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   155
hide_fact (open) is_empty_set
31807
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   156
40672
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   157
lemma empty_set [code]:
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   158
  "bot = Cset.set []"
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   159
  by (simp add: set_def)
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   160
hide_fact (open) empty_set
31807
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   161
40672
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   162
lemma UNIV_set [code]:
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   163
  "top = Cset.coset []"
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   164
  by (simp add: coset_def)
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   165
hide_fact (open) UNIV_set
31807
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   166
40672
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   167
definition insert :: "'a \<Rightarrow> 'a Cset.set \<Rightarrow> 'a Cset.set" where
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   168
  [simp]: "insert x A = Set (Set.insert x (member A))"
31807
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   169
40672
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   170
lemma insert_set [code]:
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   171
  "insert x (Cset.set xs) = Cset.set (List.insert x xs)"
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   172
  "insert x (Cset.coset xs) = Cset.coset (removeAll x xs)"
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   173
  by (simp_all add: set_def coset_def)
31807
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   174
40672
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   175
definition remove :: "'a \<Rightarrow> 'a Cset.set \<Rightarrow> 'a Cset.set" where
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   176
  [simp]: "remove x A = Set (More_Set.remove x (member A))"
31807
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   177
40672
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   178
lemma remove_set [code]:
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   179
  "remove x (Cset.set xs) = Cset.set (removeAll x xs)"
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   180
  "remove x (Cset.coset xs) = Cset.coset (List.insert x xs)"
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   181
  by (simp_all add: set_def coset_def remove_set_compl)
37024
e938a0b5286e renamed List_Set to the now more appropriate More_Set
haftmann
parents: 37023
diff changeset
   182
    (simp add: More_Set.remove_def)
31807
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   183
40672
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   184
definition map :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a Cset.set \<Rightarrow> 'b Cset.set" where
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   185
  [simp]: "map f A = Set (image f (member A))"
31807
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   186
40672
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   187
lemma map_set [code]:
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   188
  "map f (Cset.set xs) = Cset.set (remdups (List.map f xs))"
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   189
  by (simp add: set_def)
31807
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   190
41505
6d19301074cf "enriched_type" replaces less specific "type_lifting"
haftmann
parents: 41372
diff changeset
   191
enriched_type map: map
41372
551eb49a6e91 tuned type_lifting declarations
haftmann
parents: 40968
diff changeset
   192
  by (simp_all add: fun_eq_iff image_compose)
40604
c0770657c8de mapper for fset type
haftmann
parents: 39929
diff changeset
   193
40672
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   194
definition filter :: "('a \<Rightarrow> bool) \<Rightarrow> 'a Cset.set \<Rightarrow> 'a Cset.set" where
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   195
  [simp]: "filter P A = Set (More_Set.project P (member A))"
31807
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   196
40672
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   197
lemma filter_set [code]:
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   198
  "filter P (Cset.set xs) = Cset.set (List.filter P xs)"
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   199
  by (simp add: set_def project_set)
31807
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   200
40672
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   201
definition forall :: "('a \<Rightarrow> bool) \<Rightarrow> 'a Cset.set \<Rightarrow> bool" where
31846
89c37daebfdd added Inter, Union
haftmann
parents: 31807
diff changeset
   202
  [simp]: "forall P A \<longleftrightarrow> Ball (member A) P"
31807
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   203
40672
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   204
lemma forall_set [code]:
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   205
  "forall P (Cset.set xs) \<longleftrightarrow> list_all P xs"
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   206
  by (simp add: set_def list_all_iff)
31807
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   207
40672
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   208
definition exists :: "('a \<Rightarrow> bool) \<Rightarrow> 'a Cset.set \<Rightarrow> bool" where
31846
89c37daebfdd added Inter, Union
haftmann
parents: 31807
diff changeset
   209
  [simp]: "exists P A \<longleftrightarrow> Bex (member A) P"
31807
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   210
40672
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   211
lemma exists_set [code]:
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   212
  "exists P (Cset.set xs) \<longleftrightarrow> list_ex P xs"
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   213
  by (simp add: set_def list_ex_iff)
31846
89c37daebfdd added Inter, Union
haftmann
parents: 31807
diff changeset
   214
40672
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   215
definition card :: "'a Cset.set \<Rightarrow> nat" where
31849
431d8588bcad renamed theory Code_Set to Fset
haftmann
parents: 31847
diff changeset
   216
  [simp]: "card A = Finite_Set.card (member A)"
431d8588bcad renamed theory Code_Set to Fset
haftmann
parents: 31847
diff changeset
   217
40672
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   218
lemma card_set [code]:
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   219
  "card (Cset.set xs) = length (remdups xs)"
31849
431d8588bcad renamed theory Code_Set to Fset
haftmann
parents: 31847
diff changeset
   220
proof -
431d8588bcad renamed theory Code_Set to Fset
haftmann
parents: 31847
diff changeset
   221
  have "Finite_Set.card (set (remdups xs)) = length (remdups xs)"
431d8588bcad renamed theory Code_Set to Fset
haftmann
parents: 31847
diff changeset
   222
    by (rule distinct_card) simp
40672
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   223
  then show ?thesis by (simp add: set_def)
31849
431d8588bcad renamed theory Code_Set to Fset
haftmann
parents: 31847
diff changeset
   224
qed
431d8588bcad renamed theory Code_Set to Fset
haftmann
parents: 31847
diff changeset
   225
40672
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   226
lemma compl_set [simp, code]:
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   227
  "- Cset.set xs = Cset.coset xs"
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   228
  by (simp add: set_def coset_def)
37023
efc202e1677e added theory More_List
haftmann
parents: 36176
diff changeset
   229
40672
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   230
lemma compl_coset [simp, code]:
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   231
  "- Cset.coset xs = Cset.set xs"
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   232
  by (simp add: set_def coset_def)
37023
efc202e1677e added theory More_List
haftmann
parents: 36176
diff changeset
   233
31846
89c37daebfdd added Inter, Union
haftmann
parents: 31807
diff changeset
   234
89c37daebfdd added Inter, Union
haftmann
parents: 31807
diff changeset
   235
subsection {* Derived operations *}
89c37daebfdd added Inter, Union
haftmann
parents: 31807
diff changeset
   236
40672
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   237
lemma subset_eq_forall [code]:
34048
369509057220 using existing lattice classes
haftmann
parents: 33939
diff changeset
   238
  "A \<le> B \<longleftrightarrow> forall (member B) A"
31846
89c37daebfdd added Inter, Union
haftmann
parents: 31807
diff changeset
   239
  by (simp add: subset_eq)
89c37daebfdd added Inter, Union
haftmann
parents: 31807
diff changeset
   240
40672
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   241
lemma subset_subset_eq [code]:
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   242
  "A < B \<longleftrightarrow> A \<le> B \<and> \<not> B \<le> (A :: 'a Cset.set)"
34048
369509057220 using existing lattice classes
haftmann
parents: 33939
diff changeset
   243
  by (fact less_le_not_le)
31846
89c37daebfdd added Inter, Union
haftmann
parents: 31807
diff changeset
   244
40672
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   245
instantiation Cset.set :: (type) equal
37468
a2a3b62fc819 quickcheck for fsets
haftmann
parents: 37024
diff changeset
   246
begin
a2a3b62fc819 quickcheck for fsets
haftmann
parents: 37024
diff changeset
   247
39190
a2775776be3f adding code attribute to definition of equality of finite sets for executability of equality of finite sets
bulwahn
parents: 38857
diff changeset
   248
definition [code]:
40672
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   249
  "HOL.equal A B \<longleftrightarrow> A \<le> B \<and> B \<le> (A :: 'a Cset.set)"
37468
a2a3b62fc819 quickcheck for fsets
haftmann
parents: 37024
diff changeset
   250
a2a3b62fc819 quickcheck for fsets
haftmann
parents: 37024
diff changeset
   251
instance proof
40672
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   252
qed (simp add: equal_set_def set_eq [symmetric] Cset.set_eq_iff)
37468
a2a3b62fc819 quickcheck for fsets
haftmann
parents: 37024
diff changeset
   253
a2a3b62fc819 quickcheck for fsets
haftmann
parents: 37024
diff changeset
   254
end
31846
89c37daebfdd added Inter, Union
haftmann
parents: 31807
diff changeset
   255
38857
97775f3e8722 renamed class/constant eq to equal; tuned some instantiations
haftmann
parents: 37765
diff changeset
   256
lemma [code nbe]:
40672
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   257
  "HOL.equal (A :: 'a Cset.set) A \<longleftrightarrow> True"
38857
97775f3e8722 renamed class/constant eq to equal; tuned some instantiations
haftmann
parents: 37765
diff changeset
   258
  by (fact equal_refl)
97775f3e8722 renamed class/constant eq to equal; tuned some instantiations
haftmann
parents: 37765
diff changeset
   259
31807
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   260
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   261
subsection {* Functorial operations *}
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   262
32880
b8bee63c7202 sets and cosets
haftmann
parents: 32139
diff changeset
   263
lemma inter_project [code]:
40672
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   264
  "inf A (Cset.set xs) = Cset.set (List.filter (member A) xs)"
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   265
  "inf A (Cset.coset xs) = foldr remove xs A"
31807
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   266
proof -
40672
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   267
  show "inf A (Cset.set xs) = Cset.set (List.filter (member A) xs)"
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   268
    by (simp add: inter project_def set_def)
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   269
  have *: "\<And>x::'a. remove = (\<lambda>x. Set \<circ> More_Set.remove x \<circ> member)"
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 39200
diff changeset
   270
    by (simp add: fun_eq_iff)
40672
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   271
  have "member \<circ> fold (\<lambda>x. Set \<circ> More_Set.remove x \<circ> member) xs =
37024
e938a0b5286e renamed List_Set to the now more appropriate More_Set
haftmann
parents: 37023
diff changeset
   272
    fold More_Set.remove xs \<circ> member"
39921
45f95e4de831 lemmas fold_commute and fold_commute_apply
haftmann
parents: 39380
diff changeset
   273
    by (rule fold_commute) (simp add: fun_eq_iff)
37024
e938a0b5286e renamed List_Set to the now more appropriate More_Set
haftmann
parents: 37023
diff changeset
   274
  then have "fold More_Set.remove xs (member A) = 
40672
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   275
    member (fold (\<lambda>x. Set \<circ> More_Set.remove x \<circ> member) xs A)"
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 39200
diff changeset
   276
    by (simp add: fun_eq_iff)
40672
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   277
  then have "inf A (Cset.coset xs) = fold remove xs A"
37023
efc202e1677e added theory More_List
haftmann
parents: 36176
diff changeset
   278
    by (simp add: Diff_eq [symmetric] minus_set *)
40672
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   279
  moreover have "\<And>x y :: 'a. Cset.remove y \<circ> Cset.remove x = Cset.remove x \<circ> Cset.remove y"
37024
e938a0b5286e renamed List_Set to the now more appropriate More_Set
haftmann
parents: 37023
diff changeset
   280
    by (auto simp add: More_Set.remove_def * intro: ext)
40672
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   281
  ultimately show "inf A (Cset.coset xs) = foldr remove xs A"
37023
efc202e1677e added theory More_List
haftmann
parents: 36176
diff changeset
   282
    by (simp add: foldr_fold)
31807
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   283
qed
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   284
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   285
lemma subtract_remove [code]:
40672
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   286
  "A - Cset.set xs = foldr remove xs A"
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   287
  "A - Cset.coset xs = Cset.set (List.filter (member A) xs)"
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   288
  by (simp_all only: diff_eq compl_set compl_coset inter_project)
32880
b8bee63c7202 sets and cosets
haftmann
parents: 32139
diff changeset
   289
b8bee63c7202 sets and cosets
haftmann
parents: 32139
diff changeset
   290
lemma union_insert [code]:
40672
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   291
  "sup (Cset.set xs) A = foldr insert xs A"
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   292
  "sup (Cset.coset xs) A = Cset.coset (List.filter (Not \<circ> member A) xs)"
32880
b8bee63c7202 sets and cosets
haftmann
parents: 32139
diff changeset
   293
proof -
40672
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   294
  have *: "\<And>x::'a. insert = (\<lambda>x. Set \<circ> Set.insert x \<circ> member)"
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 39200
diff changeset
   295
    by (simp add: fun_eq_iff)
40672
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   296
  have "member \<circ> fold (\<lambda>x. Set \<circ> Set.insert x \<circ> member) xs =
37023
efc202e1677e added theory More_List
haftmann
parents: 36176
diff changeset
   297
    fold Set.insert xs \<circ> member"
39921
45f95e4de831 lemmas fold_commute and fold_commute_apply
haftmann
parents: 39380
diff changeset
   298
    by (rule fold_commute) (simp add: fun_eq_iff)
37023
efc202e1677e added theory More_List
haftmann
parents: 36176
diff changeset
   299
  then have "fold Set.insert xs (member A) =
40672
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   300
    member (fold (\<lambda>x. Set \<circ> Set.insert x \<circ> member) xs A)"
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 39200
diff changeset
   301
    by (simp add: fun_eq_iff)
40672
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   302
  then have "sup (Cset.set xs) A = fold insert xs A"
37023
efc202e1677e added theory More_List
haftmann
parents: 36176
diff changeset
   303
    by (simp add: union_set *)
40672
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   304
  moreover have "\<And>x y :: 'a. Cset.insert y \<circ> Cset.insert x = Cset.insert x \<circ> Cset.insert y"
37023
efc202e1677e added theory More_List
haftmann
parents: 36176
diff changeset
   305
    by (auto simp add: * intro: ext)
40672
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   306
  ultimately show "sup (Cset.set xs) A = foldr insert xs A"
37023
efc202e1677e added theory More_List
haftmann
parents: 36176
diff changeset
   307
    by (simp add: foldr_fold)
40672
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   308
  show "sup (Cset.coset xs) A = Cset.coset (List.filter (Not \<circ> member A) xs)"
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   309
    by (auto simp add: coset_def)
31807
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   310
qed
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   311
34048
369509057220 using existing lattice classes
haftmann
parents: 33939
diff changeset
   312
context complete_lattice
369509057220 using existing lattice classes
haftmann
parents: 33939
diff changeset
   313
begin
31807
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   314
40672
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   315
definition Infimum :: "'a Cset.set \<Rightarrow> 'a" where
34048
369509057220 using existing lattice classes
haftmann
parents: 33939
diff changeset
   316
  [simp]: "Infimum A = Inf (member A)"
31807
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   317
34048
369509057220 using existing lattice classes
haftmann
parents: 33939
diff changeset
   318
lemma Infimum_inf [code]:
40672
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   319
  "Infimum (Cset.set As) = foldr inf As top"
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   320
  "Infimum (Cset.coset []) = bot"
37023
efc202e1677e added theory More_List
haftmann
parents: 36176
diff changeset
   321
  by (simp_all add: Inf_set_foldr Inf_UNIV)
31807
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   322
40672
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   323
definition Supremum :: "'a Cset.set \<Rightarrow> 'a" where
34048
369509057220 using existing lattice classes
haftmann
parents: 33939
diff changeset
   324
  [simp]: "Supremum A = Sup (member A)"
369509057220 using existing lattice classes
haftmann
parents: 33939
diff changeset
   325
369509057220 using existing lattice classes
haftmann
parents: 33939
diff changeset
   326
lemma Supremum_sup [code]:
40672
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   327
  "Supremum (Cset.set As) = foldr sup As bot"
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   328
  "Supremum (Cset.coset []) = top"
37023
efc202e1677e added theory More_List
haftmann
parents: 36176
diff changeset
   329
  by (simp_all add: Sup_set_foldr Sup_UNIV)
34048
369509057220 using existing lattice classes
haftmann
parents: 33939
diff changeset
   330
369509057220 using existing lattice classes
haftmann
parents: 33939
diff changeset
   331
end
31807
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   332
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   333
31846
89c37daebfdd added Inter, Union
haftmann
parents: 31807
diff changeset
   334
subsection {* Simplified simprules *}
89c37daebfdd added Inter, Union
haftmann
parents: 31807
diff changeset
   335
89c37daebfdd added Inter, Union
haftmann
parents: 31807
diff changeset
   336
lemma is_empty_simp [simp]:
89c37daebfdd added Inter, Union
haftmann
parents: 31807
diff changeset
   337
  "is_empty A \<longleftrightarrow> member A = {}"
37024
e938a0b5286e renamed List_Set to the now more appropriate More_Set
haftmann
parents: 37023
diff changeset
   338
  by (simp add: More_Set.is_empty_def)
31846
89c37daebfdd added Inter, Union
haftmann
parents: 31807
diff changeset
   339
declare is_empty_def [simp del]
89c37daebfdd added Inter, Union
haftmann
parents: 31807
diff changeset
   340
89c37daebfdd added Inter, Union
haftmann
parents: 31807
diff changeset
   341
lemma remove_simp [simp]:
40672
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   342
  "remove x A = Set (member A - {x})"
37024
e938a0b5286e renamed List_Set to the now more appropriate More_Set
haftmann
parents: 37023
diff changeset
   343
  by (simp add: More_Set.remove_def)
31846
89c37daebfdd added Inter, Union
haftmann
parents: 31807
diff changeset
   344
declare remove_def [simp del]
89c37daebfdd added Inter, Union
haftmann
parents: 31807
diff changeset
   345
31847
7de0e20ca24d Executable_Set now based on Code_Set
haftmann
parents: 31846
diff changeset
   346
lemma filter_simp [simp]:
40672
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40604
diff changeset
   347
  "filter P A = Set {x \<in> member A. P x}"
37024
e938a0b5286e renamed List_Set to the now more appropriate More_Set
haftmann
parents: 37023
diff changeset
   348
  by (simp add: More_Set.project_def)
31847
7de0e20ca24d Executable_Set now based on Code_Set
haftmann
parents: 31846
diff changeset
   349
declare filter_def [simp del]
31846
89c37daebfdd added Inter, Union
haftmann
parents: 31807
diff changeset
   350
89c37daebfdd added Inter, Union
haftmann
parents: 31807
diff changeset
   351
declare mem_def [simp del]
89c37daebfdd added Inter, Union
haftmann
parents: 31807
diff changeset
   352
31849
431d8588bcad renamed theory Code_Set to Fset
haftmann
parents: 31847
diff changeset
   353
37468
a2a3b62fc819 quickcheck for fsets
haftmann
parents: 37024
diff changeset
   354
hide_const (open) setify is_empty insert remove map filter forall exists card
34048
369509057220 using existing lattice classes
haftmann
parents: 33939
diff changeset
   355
  Inter Union
31849
431d8588bcad renamed theory Code_Set to Fset
haftmann
parents: 31847
diff changeset
   356
31807
039893a9a77d added List_Set and Code_Set theories
haftmann
parents:
diff changeset
   357
end