author | wenzelm |
Wed, 15 Feb 2012 23:19:30 +0100 | |
changeset 46497 | 89ccf66aa73d |
parent 46305 | 8ea02e499d53 |
child 47232 | e2f0176149d0 |
permissions | -rw-r--r-- |
45319
2b002c6b0f7d
renaming Quotient_Set and List_Quotient_Set to Quotient_Cset and List_Quotient_Cset to avoid name clash with existing Quotient_Set (again, cf. 66823a0066db)
bulwahn
parents:
45267
diff
changeset
|
1 |
(* Title: HOL/Quotient_Examples/List_Quotient_Cset.thy |
43800
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
2 |
Author: Florian Haftmann, Alexander Krauss, TU Muenchen |
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
3 |
*) |
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
4 |
|
45319
2b002c6b0f7d
renaming Quotient_Set and List_Quotient_Set to Quotient_Cset and List_Quotient_Cset to avoid name clash with existing Quotient_Set (again, cf. 66823a0066db)
bulwahn
parents:
45267
diff
changeset
|
5 |
header {* Implementation of type Quotient_Cset.set based on lists. Code equations obtained via quotient lifting. *} |
43800
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
6 |
|
45319
2b002c6b0f7d
renaming Quotient_Set and List_Quotient_Set to Quotient_Cset and List_Quotient_Cset to avoid name clash with existing Quotient_Set (again, cf. 66823a0066db)
bulwahn
parents:
45267
diff
changeset
|
7 |
theory List_Quotient_Cset |
2b002c6b0f7d
renaming Quotient_Set and List_Quotient_Set to Quotient_Cset and List_Quotient_Cset to avoid name clash with existing Quotient_Set (again, cf. 66823a0066db)
bulwahn
parents:
45267
diff
changeset
|
8 |
imports Quotient_Cset |
43800
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
9 |
begin |
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
10 |
|
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
11 |
lemma [quot_respect]: "((op = ===> set_eq ===> set_eq) ===> op = ===> set_eq ===> set_eq) |
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
12 |
foldr foldr" |
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
13 |
by (simp add: fun_rel_eq) |
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
14 |
|
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
15 |
lemma [quot_preserve]: "((id ---> abs_set ---> rep_set) ---> id ---> rep_set ---> abs_set) foldr = foldr" |
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
16 |
apply (rule ext)+ |
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
17 |
by (induct_tac xa) (auto simp: Quotient_abs_rep[OF Quotient_set]) |
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
18 |
|
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
19 |
|
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
20 |
subsection {* Relationship to lists *} |
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
21 |
|
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
22 |
(*FIXME: maybe define on sets first and then lift -> more canonical*) |
45319
2b002c6b0f7d
renaming Quotient_Set and List_Quotient_Set to Quotient_Cset and List_Quotient_Cset to avoid name clash with existing Quotient_Set (again, cf. 66823a0066db)
bulwahn
parents:
45267
diff
changeset
|
23 |
definition coset :: "'a list \<Rightarrow> 'a Quotient_Cset.set" where |
2b002c6b0f7d
renaming Quotient_Set and List_Quotient_Set to Quotient_Cset and List_Quotient_Cset to avoid name clash with existing Quotient_Set (again, cf. 66823a0066db)
bulwahn
parents:
45267
diff
changeset
|
24 |
"coset xs = Quotient_Cset.uminus (Quotient_Cset.set xs)" |
43800
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
25 |
|
45319
2b002c6b0f7d
renaming Quotient_Set and List_Quotient_Set to Quotient_Cset and List_Quotient_Cset to avoid name clash with existing Quotient_Set (again, cf. 66823a0066db)
bulwahn
parents:
45267
diff
changeset
|
26 |
code_datatype Quotient_Cset.set List_Quotient_Cset.coset |
43800
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
27 |
|
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
28 |
lemma member_code [code]: |
45319
2b002c6b0f7d
renaming Quotient_Set and List_Quotient_Set to Quotient_Cset and List_Quotient_Cset to avoid name clash with existing Quotient_Set (again, cf. 66823a0066db)
bulwahn
parents:
45267
diff
changeset
|
29 |
"member x (Quotient_Cset.set xs) \<longleftrightarrow> List.member xs x" |
43800
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
30 |
"member x (coset xs) \<longleftrightarrow> \<not> List.member xs x" |
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
31 |
unfolding coset_def |
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
32 |
apply (lifting in_set_member) |
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
33 |
by descending (simp add: in_set_member) |
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
34 |
|
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
35 |
definition (in term_syntax) |
46305 | 36 |
csetify :: "'a\<Colon>typerep list \<times> (unit \<Rightarrow> Code_Evaluation.term) |
45319
2b002c6b0f7d
renaming Quotient_Set and List_Quotient_Set to Quotient_Cset and List_Quotient_Cset to avoid name clash with existing Quotient_Set (again, cf. 66823a0066db)
bulwahn
parents:
45267
diff
changeset
|
37 |
\<Rightarrow> 'a Quotient_Cset.set \<times> (unit \<Rightarrow> Code_Evaluation.term)" where |
46305 | 38 |
[code_unfold]: "csetify xs = Code_Evaluation.valtermify Quotient_Cset.set {\<cdot>} xs" |
43800
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
39 |
|
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
40 |
notation fcomp (infixl "\<circ>>" 60) |
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
41 |
notation scomp (infixl "\<circ>\<rightarrow>" 60) |
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
42 |
|
45319
2b002c6b0f7d
renaming Quotient_Set and List_Quotient_Set to Quotient_Cset and List_Quotient_Cset to avoid name clash with existing Quotient_Set (again, cf. 66823a0066db)
bulwahn
parents:
45267
diff
changeset
|
43 |
instantiation Quotient_Cset.set :: (random) random |
43800
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
44 |
begin |
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
45 |
|
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
46 |
definition |
46305 | 47 |
"Quickcheck.random i = Quickcheck.random i \<circ>\<rightarrow> (\<lambda>xs. Pair (csetify xs))" |
43800
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
48 |
|
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
49 |
instance .. |
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
50 |
|
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
51 |
end |
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
52 |
|
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
53 |
no_notation fcomp (infixl "\<circ>>" 60) |
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
54 |
no_notation scomp (infixl "\<circ>\<rightarrow>" 60) |
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
55 |
|
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
56 |
subsection {* Basic operations *} |
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
57 |
|
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
58 |
lemma is_empty_set [code]: |
45319
2b002c6b0f7d
renaming Quotient_Set and List_Quotient_Set to Quotient_Cset and List_Quotient_Cset to avoid name clash with existing Quotient_Set (again, cf. 66823a0066db)
bulwahn
parents:
45267
diff
changeset
|
59 |
"Quotient_Cset.is_empty (Quotient_Cset.set xs) \<longleftrightarrow> List.null xs" |
43800
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
60 |
by (lifting is_empty_set) |
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
61 |
hide_fact (open) is_empty_set |
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
62 |
|
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
63 |
lemma empty_set [code]: |
45319
2b002c6b0f7d
renaming Quotient_Set and List_Quotient_Set to Quotient_Cset and List_Quotient_Cset to avoid name clash with existing Quotient_Set (again, cf. 66823a0066db)
bulwahn
parents:
45267
diff
changeset
|
64 |
"Quotient_Cset.empty = Quotient_Cset.set []" |
43800
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
65 |
by (lifting set.simps(1)[symmetric]) |
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
66 |
hide_fact (open) empty_set |
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
67 |
|
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
68 |
lemma UNIV_set [code]: |
45319
2b002c6b0f7d
renaming Quotient_Set and List_Quotient_Set to Quotient_Cset and List_Quotient_Cset to avoid name clash with existing Quotient_Set (again, cf. 66823a0066db)
bulwahn
parents:
45267
diff
changeset
|
69 |
"Quotient_Cset.UNIV = coset []" |
43800
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
70 |
unfolding coset_def by descending simp |
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
71 |
hide_fact (open) UNIV_set |
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
72 |
|
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
73 |
lemma remove_set [code]: |
45319
2b002c6b0f7d
renaming Quotient_Set and List_Quotient_Set to Quotient_Cset and List_Quotient_Cset to avoid name clash with existing Quotient_Set (again, cf. 66823a0066db)
bulwahn
parents:
45267
diff
changeset
|
74 |
"Quotient_Cset.remove x (Quotient_Cset.set xs) = Quotient_Cset.set (removeAll x xs)" |
2b002c6b0f7d
renaming Quotient_Set and List_Quotient_Set to Quotient_Cset and List_Quotient_Cset to avoid name clash with existing Quotient_Set (again, cf. 66823a0066db)
bulwahn
parents:
45267
diff
changeset
|
75 |
"Quotient_Cset.remove x (coset xs) = coset (List.insert x xs)" |
43800
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
76 |
unfolding coset_def |
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
77 |
apply descending |
45986
c9e50153e5ae
moved various set operations to theory Set (resp. Product_Type)
haftmann
parents:
45319
diff
changeset
|
78 |
apply (simp add: Set.remove_def) |
43800
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
79 |
apply descending |
46148
04a8da7dcffc
moved lemmas about List.set and set operations to List theory
haftmann
parents:
45986
diff
changeset
|
80 |
by (auto simp add: set_eq_iff) |
43800
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
81 |
|
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
82 |
lemma insert_set [code]: |
45319
2b002c6b0f7d
renaming Quotient_Set and List_Quotient_Set to Quotient_Cset and List_Quotient_Cset to avoid name clash with existing Quotient_Set (again, cf. 66823a0066db)
bulwahn
parents:
45267
diff
changeset
|
83 |
"Quotient_Cset.insert x (Quotient_Cset.set xs) = Quotient_Cset.set (List.insert x xs)" |
2b002c6b0f7d
renaming Quotient_Set and List_Quotient_Set to Quotient_Cset and List_Quotient_Cset to avoid name clash with existing Quotient_Set (again, cf. 66823a0066db)
bulwahn
parents:
45267
diff
changeset
|
84 |
"Quotient_Cset.insert x (coset xs) = coset (removeAll x xs)" |
43800
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
85 |
unfolding coset_def |
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
86 |
apply (lifting set_insert[symmetric]) |
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
87 |
by descending simp |
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
88 |
|
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
89 |
lemma map_set [code]: |
45319
2b002c6b0f7d
renaming Quotient_Set and List_Quotient_Set to Quotient_Cset and List_Quotient_Cset to avoid name clash with existing Quotient_Set (again, cf. 66823a0066db)
bulwahn
parents:
45267
diff
changeset
|
90 |
"Quotient_Cset.map f (Quotient_Cset.set xs) = Quotient_Cset.set (remdups (List.map f xs))" |
43800
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
91 |
by descending simp |
44293
83c4f8ba0aa3
Quotient_Examples: Cset, List_Cset: Lift Inf and Sup directly.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
43926
diff
changeset
|
92 |
|
43800
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
93 |
lemma filter_set [code]: |
45319
2b002c6b0f7d
renaming Quotient_Set and List_Quotient_Set to Quotient_Cset and List_Quotient_Cset to avoid name clash with existing Quotient_Set (again, cf. 66823a0066db)
bulwahn
parents:
45267
diff
changeset
|
94 |
"Quotient_Cset.filter P (Quotient_Cset.set xs) = Quotient_Cset.set (List.filter P xs)" |
46148
04a8da7dcffc
moved lemmas about List.set and set operations to List theory
haftmann
parents:
45986
diff
changeset
|
95 |
by descending (simp add: set_eq_iff) |
43800
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
96 |
|
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
97 |
lemma forall_set [code]: |
45319
2b002c6b0f7d
renaming Quotient_Set and List_Quotient_Set to Quotient_Cset and List_Quotient_Cset to avoid name clash with existing Quotient_Set (again, cf. 66823a0066db)
bulwahn
parents:
45267
diff
changeset
|
98 |
"Quotient_Cset.forall (Quotient_Cset.set xs) P \<longleftrightarrow> list_all P xs" |
43800
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
99 |
(* FIXME: why does (lifting Ball_set_list_all) fail? *) |
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
100 |
by descending (fact Ball_set_list_all) |
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
101 |
|
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
102 |
lemma exists_set [code]: |
45319
2b002c6b0f7d
renaming Quotient_Set and List_Quotient_Set to Quotient_Cset and List_Quotient_Cset to avoid name clash with existing Quotient_Set (again, cf. 66823a0066db)
bulwahn
parents:
45267
diff
changeset
|
103 |
"Quotient_Cset.exists (Quotient_Cset.set xs) P \<longleftrightarrow> list_ex P xs" |
43800
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
104 |
by descending (fact Bex_set_list_ex) |
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
105 |
|
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
106 |
lemma card_set [code]: |
45319
2b002c6b0f7d
renaming Quotient_Set and List_Quotient_Set to Quotient_Cset and List_Quotient_Cset to avoid name clash with existing Quotient_Set (again, cf. 66823a0066db)
bulwahn
parents:
45267
diff
changeset
|
107 |
"Quotient_Cset.card (Quotient_Cset.set xs) = length (remdups xs)" |
43800
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
108 |
by (lifting length_remdups_card_conv[symmetric]) |
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
109 |
|
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
110 |
lemma compl_set [simp, code]: |
45319
2b002c6b0f7d
renaming Quotient_Set and List_Quotient_Set to Quotient_Cset and List_Quotient_Cset to avoid name clash with existing Quotient_Set (again, cf. 66823a0066db)
bulwahn
parents:
45267
diff
changeset
|
111 |
"Quotient_Cset.uminus (Quotient_Cset.set xs) = coset xs" |
43800
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
112 |
unfolding coset_def by descending simp |
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
113 |
|
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
114 |
lemma compl_coset [simp, code]: |
45319
2b002c6b0f7d
renaming Quotient_Set and List_Quotient_Set to Quotient_Cset and List_Quotient_Cset to avoid name clash with existing Quotient_Set (again, cf. 66823a0066db)
bulwahn
parents:
45267
diff
changeset
|
115 |
"Quotient_Cset.uminus (coset xs) = Quotient_Cset.set xs" |
43800
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
116 |
unfolding coset_def by descending simp |
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
117 |
|
44293
83c4f8ba0aa3
Quotient_Examples: Cset, List_Cset: Lift Inf and Sup directly.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
43926
diff
changeset
|
118 |
lemma Inf_inf [code]: |
45319
2b002c6b0f7d
renaming Quotient_Set and List_Quotient_Set to Quotient_Cset and List_Quotient_Cset to avoid name clash with existing Quotient_Set (again, cf. 66823a0066db)
bulwahn
parents:
45267
diff
changeset
|
119 |
"Quotient_Cset.Inf (Quotient_Cset.set (xs\<Colon>'a\<Colon>complete_lattice list)) = foldr inf xs top" |
2b002c6b0f7d
renaming Quotient_Set and List_Quotient_Set to Quotient_Cset and List_Quotient_Cset to avoid name clash with existing Quotient_Set (again, cf. 66823a0066db)
bulwahn
parents:
45267
diff
changeset
|
120 |
"Quotient_Cset.Inf (coset ([]\<Colon>'a\<Colon>complete_lattice list)) = bot" |
2b002c6b0f7d
renaming Quotient_Set and List_Quotient_Set to Quotient_Cset and List_Quotient_Cset to avoid name clash with existing Quotient_Set (again, cf. 66823a0066db)
bulwahn
parents:
45267
diff
changeset
|
121 |
unfolding List_Quotient_Cset.UNIV_set[symmetric] |
44293
83c4f8ba0aa3
Quotient_Examples: Cset, List_Cset: Lift Inf and Sup directly.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
43926
diff
changeset
|
122 |
by (lifting Inf_set_foldr Inf_UNIV) |
43800
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
123 |
|
44293
83c4f8ba0aa3
Quotient_Examples: Cset, List_Cset: Lift Inf and Sup directly.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
43926
diff
changeset
|
124 |
lemma Sup_sup [code]: |
45319
2b002c6b0f7d
renaming Quotient_Set and List_Quotient_Set to Quotient_Cset and List_Quotient_Cset to avoid name clash with existing Quotient_Set (again, cf. 66823a0066db)
bulwahn
parents:
45267
diff
changeset
|
125 |
"Quotient_Cset.Sup (Quotient_Cset.set (xs\<Colon>'a\<Colon>complete_lattice list)) = foldr sup xs bot" |
2b002c6b0f7d
renaming Quotient_Set and List_Quotient_Set to Quotient_Cset and List_Quotient_Cset to avoid name clash with existing Quotient_Set (again, cf. 66823a0066db)
bulwahn
parents:
45267
diff
changeset
|
126 |
"Quotient_Cset.Sup (coset ([]\<Colon>'a\<Colon>complete_lattice list)) = top" |
2b002c6b0f7d
renaming Quotient_Set and List_Quotient_Set to Quotient_Cset and List_Quotient_Cset to avoid name clash with existing Quotient_Set (again, cf. 66823a0066db)
bulwahn
parents:
45267
diff
changeset
|
127 |
unfolding List_Quotient_Cset.UNIV_set[symmetric] |
44293
83c4f8ba0aa3
Quotient_Examples: Cset, List_Cset: Lift Inf and Sup directly.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
43926
diff
changeset
|
128 |
by (lifting Sup_set_foldr Sup_UNIV) |
43800
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
129 |
|
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
130 |
subsection {* Derived operations *} |
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
131 |
|
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
132 |
lemma subset_eq_forall [code]: |
45319
2b002c6b0f7d
renaming Quotient_Set and List_Quotient_Set to Quotient_Cset and List_Quotient_Cset to avoid name clash with existing Quotient_Set (again, cf. 66823a0066db)
bulwahn
parents:
45267
diff
changeset
|
133 |
"Quotient_Cset.subset A B \<longleftrightarrow> Quotient_Cset.forall A (\<lambda>x. member x B)" |
43800
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
134 |
by descending blast |
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
135 |
|
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
136 |
lemma subset_subset_eq [code]: |
45319
2b002c6b0f7d
renaming Quotient_Set and List_Quotient_Set to Quotient_Cset and List_Quotient_Cset to avoid name clash with existing Quotient_Set (again, cf. 66823a0066db)
bulwahn
parents:
45267
diff
changeset
|
137 |
"Quotient_Cset.psubset A B \<longleftrightarrow> Quotient_Cset.subset A B \<and> \<not> Quotient_Cset.subset B A" |
43800
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
138 |
by descending blast |
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
139 |
|
45319
2b002c6b0f7d
renaming Quotient_Set and List_Quotient_Set to Quotient_Cset and List_Quotient_Cset to avoid name clash with existing Quotient_Set (again, cf. 66823a0066db)
bulwahn
parents:
45267
diff
changeset
|
140 |
instantiation Quotient_Cset.set :: (type) equal |
43800
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
141 |
begin |
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
142 |
|
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
143 |
definition [code]: |
45319
2b002c6b0f7d
renaming Quotient_Set and List_Quotient_Set to Quotient_Cset and List_Quotient_Cset to avoid name clash with existing Quotient_Set (again, cf. 66823a0066db)
bulwahn
parents:
45267
diff
changeset
|
144 |
"HOL.equal A B \<longleftrightarrow> Quotient_Cset.subset A B \<and> Quotient_Cset.subset B A" |
43800
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
145 |
|
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
146 |
instance |
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
147 |
apply intro_classes |
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
148 |
unfolding equal_set_def |
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
149 |
by descending auto |
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
150 |
|
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
151 |
end |
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
152 |
|
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
153 |
lemma [code nbe]: |
45319
2b002c6b0f7d
renaming Quotient_Set and List_Quotient_Set to Quotient_Cset and List_Quotient_Cset to avoid name clash with existing Quotient_Set (again, cf. 66823a0066db)
bulwahn
parents:
45267
diff
changeset
|
154 |
"HOL.equal (A :: 'a Quotient_Cset.set) A \<longleftrightarrow> True" |
43800
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
155 |
by (fact equal_refl) |
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
156 |
|
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
157 |
|
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
158 |
subsection {* Functorial operations *} |
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
159 |
|
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
160 |
lemma inter_project [code]: |
45319
2b002c6b0f7d
renaming Quotient_Set and List_Quotient_Set to Quotient_Cset and List_Quotient_Cset to avoid name clash with existing Quotient_Set (again, cf. 66823a0066db)
bulwahn
parents:
45267
diff
changeset
|
161 |
"Quotient_Cset.inter A (Quotient_Cset.set xs) = Quotient_Cset.set (List.filter (\<lambda>x. Quotient_Cset.member x A) xs)" |
2b002c6b0f7d
renaming Quotient_Set and List_Quotient_Set to Quotient_Cset and List_Quotient_Cset to avoid name clash with existing Quotient_Set (again, cf. 66823a0066db)
bulwahn
parents:
45267
diff
changeset
|
162 |
"Quotient_Cset.inter A (coset xs) = foldr Quotient_Cset.remove xs A" |
43800
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
163 |
apply descending |
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
164 |
apply auto |
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
165 |
unfolding coset_def |
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
166 |
apply descending |
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
167 |
apply simp |
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
168 |
by (metis diff_eq minus_set_foldr) |
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
169 |
|
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
170 |
lemma subtract_remove [code]: |
45319
2b002c6b0f7d
renaming Quotient_Set and List_Quotient_Set to Quotient_Cset and List_Quotient_Cset to avoid name clash with existing Quotient_Set (again, cf. 66823a0066db)
bulwahn
parents:
45267
diff
changeset
|
171 |
"Quotient_Cset.minus A (Quotient_Cset.set xs) = foldr Quotient_Cset.remove xs A" |
2b002c6b0f7d
renaming Quotient_Set and List_Quotient_Set to Quotient_Cset and List_Quotient_Cset to avoid name clash with existing Quotient_Set (again, cf. 66823a0066db)
bulwahn
parents:
45267
diff
changeset
|
172 |
"Quotient_Cset.minus A (coset xs) = Quotient_Cset.set (List.filter (\<lambda>x. member x A) xs)" |
43800
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
173 |
unfolding coset_def |
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
174 |
apply (lifting minus_set_foldr) |
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
175 |
by descending auto |
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
176 |
|
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
177 |
lemma union_insert [code]: |
45319
2b002c6b0f7d
renaming Quotient_Set and List_Quotient_Set to Quotient_Cset and List_Quotient_Cset to avoid name clash with existing Quotient_Set (again, cf. 66823a0066db)
bulwahn
parents:
45267
diff
changeset
|
178 |
"Quotient_Cset.union (Quotient_Cset.set xs) A = foldr Quotient_Cset.insert xs A" |
2b002c6b0f7d
renaming Quotient_Set and List_Quotient_Set to Quotient_Cset and List_Quotient_Cset to avoid name clash with existing Quotient_Set (again, cf. 66823a0066db)
bulwahn
parents:
45267
diff
changeset
|
179 |
"Quotient_Cset.union (coset xs) A = coset (List.filter (\<lambda>x. \<not> member x A) xs)" |
43800
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
180 |
unfolding coset_def |
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
181 |
apply (lifting union_set_foldr) |
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
182 |
by descending auto |
9959c8732edf
experimental variants of Library/Cset.thy and Library/Dlist_Cset.thy defined via quotient package
krauss
parents:
diff
changeset
|
183 |
|
43926 | 184 |
lemma UNION_code [code]: |
45319
2b002c6b0f7d
renaming Quotient_Set and List_Quotient_Set to Quotient_Cset and List_Quotient_Cset to avoid name clash with existing Quotient_Set (again, cf. 66823a0066db)
bulwahn
parents:
45267
diff
changeset
|
185 |
"Quotient_Cset.UNION (Quotient_Cset.set []) f = Quotient_Cset.set []" |
2b002c6b0f7d
renaming Quotient_Set and List_Quotient_Set to Quotient_Cset and List_Quotient_Cset to avoid name clash with existing Quotient_Set (again, cf. 66823a0066db)
bulwahn
parents:
45267
diff
changeset
|
186 |
"Quotient_Cset.UNION (Quotient_Cset.set (x#xs)) f = |
2b002c6b0f7d
renaming Quotient_Set and List_Quotient_Set to Quotient_Cset and List_Quotient_Cset to avoid name clash with existing Quotient_Set (again, cf. 66823a0066db)
bulwahn
parents:
45267
diff
changeset
|
187 |
Quotient_Cset.union (f x) (Quotient_Cset.UNION (Quotient_Cset.set xs) f)" |
43926 | 188 |
by (descending, simp)+ |
189 |
||
190 |
||
44293
83c4f8ba0aa3
Quotient_Examples: Cset, List_Cset: Lift Inf and Sup directly.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
43926
diff
changeset
|
191 |
end |