author | wenzelm |
Mon, 30 Sep 2024 23:32:26 +0200 | |
changeset 81090 | 843dba3d307a |
parent 81019 | dd59daa3c37a |
child 81091 | c007e6d9941d |
permissions | -rw-r--r-- |
63920
003622e08379
resolve the name clash of HOL/Library/FSet and HOL/Quotient_Examples/FSet
kuncar
parents:
63167
diff
changeset
|
1 |
(* Title: HOL/Quotient_Examples/Quotient_FSet.thy |
36465 | 2 |
Author: Cezary Kaliszyk, TU Munich |
3 |
Author: Christian Urban, TU Munich |
|
36280 | 4 |
|
41467 | 5 |
Type of finite sets. |
36280 | 6 |
*) |
36465 | 7 |
|
54336
9a21e5c6e5d9
declare Quotient_Examples/FSet.thy as almost obsolete
kuncar
parents:
53374
diff
changeset
|
8 |
(******************************************************************** |
9a21e5c6e5d9
declare Quotient_Examples/FSet.thy as almost obsolete
kuncar
parents:
53374
diff
changeset
|
9 |
WARNING: There is a formalization of 'a fset as a subtype of sets in |
9a21e5c6e5d9
declare Quotient_Examples/FSet.thy as almost obsolete
kuncar
parents:
53374
diff
changeset
|
10 |
HOL/Library/FSet.thy using Lifting/Transfer. The user should use |
9a21e5c6e5d9
declare Quotient_Examples/FSet.thy as almost obsolete
kuncar
parents:
53374
diff
changeset
|
11 |
that file rather than this file unless there are some very specific |
9a21e5c6e5d9
declare Quotient_Examples/FSet.thy as almost obsolete
kuncar
parents:
53374
diff
changeset
|
12 |
reasons. |
9a21e5c6e5d9
declare Quotient_Examples/FSet.thy as almost obsolete
kuncar
parents:
53374
diff
changeset
|
13 |
*********************************************************************) |
9a21e5c6e5d9
declare Quotient_Examples/FSet.thy as almost obsolete
kuncar
parents:
53374
diff
changeset
|
14 |
|
63920
003622e08379
resolve the name clash of HOL/Library/FSet and HOL/Quotient_Examples/FSet
kuncar
parents:
63167
diff
changeset
|
15 |
theory Quotient_FSet |
66453
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
wenzelm
parents:
63920
diff
changeset
|
16 |
imports "HOL-Library.Multiset" "HOL-Library.Quotient_List" |
36280 | 17 |
begin |
18 |
||
63167 | 19 |
text \<open> |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
20 |
The type of finite sets is created by a quotient construction |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
21 |
over lists. The definition of the equivalence: |
63167 | 22 |
\<close> |
36280 | 23 |
|
40467
dc0439fdd7c5
more appropriate specification packages; fun_rel_def is no simp rule by default
haftmann
parents:
40034
diff
changeset
|
24 |
definition |
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80786
diff
changeset
|
25 |
list_eq :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool" (infix \<open>\<approx>\<close> 50) |
36280 | 26 |
where |
40952 | 27 |
[simp]: "xs \<approx> ys \<longleftrightarrow> set xs = set ys" |
36280 | 28 |
|
40822 | 29 |
lemma list_eq_reflp: |
30 |
"reflp list_eq" |
|
31 |
by (auto intro: reflpI) |
|
32 |
||
33 |
lemma list_eq_symp: |
|
34 |
"symp list_eq" |
|
35 |
by (auto intro: sympI) |
|
36 |
||
37 |
lemma list_eq_transp: |
|
38 |
"transp list_eq" |
|
39 |
by (auto intro: transpI) |
|
40 |
||
36280 | 41 |
lemma list_eq_equivp: |
40822 | 42 |
"equivp list_eq" |
43 |
by (auto intro: equivpI list_eq_reflp list_eq_symp list_eq_transp) |
|
36280 | 44 |
|
63167 | 45 |
text \<open>The \<open>fset\<close> type\<close> |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
46 |
|
36280 | 47 |
quotient_type |
48 |
'a fset = "'a list" / "list_eq" |
|
49 |
by (rule list_eq_equivp) |
|
50 |
||
63167 | 51 |
text \<open> |
40953 | 52 |
Definitions for sublist, cardinality, |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
53 |
intersection, difference and respectful fold over |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
54 |
lists. |
63167 | 55 |
\<close> |
36280 | 56 |
|
40953 | 57 |
declare List.member_def [simp] |
36280 | 58 |
|
40034 | 59 |
definition |
36280 | 60 |
sub_list :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool" |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
61 |
where |
40034 | 62 |
[simp]: "sub_list xs ys \<longleftrightarrow> set xs \<subseteq> set ys" |
36280 | 63 |
|
40034 | 64 |
definition |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
65 |
card_list :: "'a list \<Rightarrow> nat" |
36280 | 66 |
where |
40034 | 67 |
[simp]: "card_list xs = card (set xs)" |
36675
806ea6e282e4
fminus and some more theorems ported from Finite_Set.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
36646
diff
changeset
|
68 |
|
40034 | 69 |
definition |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
70 |
inter_list :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" |
36675
806ea6e282e4
fminus and some more theorems ported from Finite_Set.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
36646
diff
changeset
|
71 |
where |
68386 | 72 |
[simp]: "inter_list xs ys = [x \<leftarrow> xs. x \<in> set xs \<and> x \<in> set ys]" |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
73 |
|
40034 | 74 |
definition |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
75 |
diff_list :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
76 |
where |
68386 | 77 |
[simp]: "diff_list xs ys = [x \<leftarrow> xs. x \<notin> set ys]" |
36280 | 78 |
|
79 |
definition |
|
40954
ecca598474dd
conventional point-free characterization of rsp_fold
haftmann
parents:
40953
diff
changeset
|
80 |
rsp_fold :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> bool" |
36280 | 81 |
where |
40954
ecca598474dd
conventional point-free characterization of rsp_fold
haftmann
parents:
40953
diff
changeset
|
82 |
"rsp_fold f \<longleftrightarrow> (\<forall>u v. f u \<circ> f v = f v \<circ> f u)" |
36280 | 83 |
|
40961 | 84 |
lemma rsp_foldI: |
85 |
"(\<And>u v. f u \<circ> f v = f v \<circ> f u) \<Longrightarrow> rsp_fold f" |
|
86 |
by (simp add: rsp_fold_def) |
|
87 |
||
88 |
lemma rsp_foldE: |
|
89 |
assumes "rsp_fold f" |
|
90 |
obtains "f u \<circ> f v = f v \<circ> f u" |
|
91 |
using assms by (simp add: rsp_fold_def) |
|
92 |
||
40962
069cd1c1f39b
more intimate definition of fold_list / fold_once in terms of fold
haftmann
parents:
40961
diff
changeset
|
93 |
definition |
069cd1c1f39b
more intimate definition of fold_list / fold_once in terms of fold
haftmann
parents:
40961
diff
changeset
|
94 |
fold_once :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'a list \<Rightarrow> 'b \<Rightarrow> 'b" |
36280 | 95 |
where |
40962
069cd1c1f39b
more intimate definition of fold_list / fold_once in terms of fold
haftmann
parents:
40961
diff
changeset
|
96 |
"fold_once f xs = (if rsp_fold f then fold f (remdups xs) else id)" |
36280 | 97 |
|
40962
069cd1c1f39b
more intimate definition of fold_list / fold_once in terms of fold
haftmann
parents:
40961
diff
changeset
|
98 |
lemma fold_once_default [simp]: |
069cd1c1f39b
more intimate definition of fold_list / fold_once in terms of fold
haftmann
parents:
40961
diff
changeset
|
99 |
"\<not> rsp_fold f \<Longrightarrow> fold_once f xs = id" |
069cd1c1f39b
more intimate definition of fold_list / fold_once in terms of fold
haftmann
parents:
40961
diff
changeset
|
100 |
by (simp add: fold_once_def) |
40961 | 101 |
|
40962
069cd1c1f39b
more intimate definition of fold_list / fold_once in terms of fold
haftmann
parents:
40961
diff
changeset
|
102 |
lemma fold_once_fold_remdups: |
069cd1c1f39b
more intimate definition of fold_list / fold_once in terms of fold
haftmann
parents:
40961
diff
changeset
|
103 |
"rsp_fold f \<Longrightarrow> fold_once f xs = fold f (remdups xs)" |
069cd1c1f39b
more intimate definition of fold_list / fold_once in terms of fold
haftmann
parents:
40961
diff
changeset
|
104 |
by (simp add: fold_once_def) |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
105 |
|
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
106 |
|
63167 | 107 |
section \<open>Quotient composition lemmas\<close> |
36280 | 108 |
|
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
109 |
lemma list_all2_refl': |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
110 |
assumes q: "equivp R" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
111 |
shows "(list_all2 R) r r" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
112 |
by (rule list_all2_refl) (metis equivp_def q) |
36280 | 113 |
|
114 |
lemma compose_list_refl: |
|
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
115 |
assumes q: "equivp R" |
67399 | 116 |
shows "(list_all2 R OOO (\<approx>)) r r" |
36280 | 117 |
proof |
36465 | 118 |
have *: "r \<approx> r" by (rule equivp_reflp[OF fset_equivp]) |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
119 |
show "list_all2 R r r" by (rule list_all2_refl'[OF q]) |
67399 | 120 |
with * show "((\<approx>) OO list_all2 R) r r" .. |
36280 | 121 |
qed |
122 |
||
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
123 |
lemma map_list_eq_cong: "b \<approx> ba \<Longrightarrow> map f b \<approx> map f ba" |
40467
dc0439fdd7c5
more appropriate specification packages; fun_rel_def is no simp rule by default
haftmann
parents:
40034
diff
changeset
|
124 |
by (simp only: list_eq_def set_map) |
36280 | 125 |
|
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
126 |
lemma quotient_compose_list_g: |
47308 | 127 |
assumes q: "Quotient3 R Abs Rep" |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
128 |
and e: "equivp R" |
67399 | 129 |
shows "Quotient3 ((list_all2 R) OOO (\<approx>)) |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
130 |
(abs_fset \<circ> (map Abs)) ((map Rep) \<circ> rep_fset)" |
47308 | 131 |
unfolding Quotient3_def comp_def |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
132 |
proof (intro conjI allI) |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
133 |
fix a r s |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
134 |
show "abs_fset (map Abs (map Rep (rep_fset a))) = a" |
47308 | 135 |
by (simp add: abs_o_rep[OF q] Quotient3_abs_rep[OF Quotient3_fset] List.map.id) |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
136 |
have b: "list_all2 R (map Rep (rep_fset a)) (map Rep (rep_fset a))" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
137 |
by (rule list_all2_refl'[OF e]) |
67399 | 138 |
have c: "((\<approx>) OO list_all2 R) (map Rep (rep_fset a)) (map Rep (rep_fset a))" |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
139 |
by (rule, rule equivp_reflp[OF fset_equivp]) (rule b) |
67399 | 140 |
show "(list_all2 R OOO (\<approx>)) (map Rep (rep_fset a)) (map Rep (rep_fset a))" |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
141 |
by (rule, rule list_all2_refl'[OF e]) (rule c) |
67399 | 142 |
show "(list_all2 R OOO (\<approx>)) r s = ((list_all2 R OOO (\<approx>)) r r \<and> |
143 |
(list_all2 R OOO (\<approx>)) s s \<and> abs_fset (map Abs r) = abs_fset (map Abs s))" |
|
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
144 |
proof (intro iffI conjI) |
67399 | 145 |
show "(list_all2 R OOO (\<approx>)) r r" by (rule compose_list_refl[OF e]) |
146 |
show "(list_all2 R OOO (\<approx>)) s s" by (rule compose_list_refl[OF e]) |
|
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
147 |
next |
67399 | 148 |
assume a: "(list_all2 R OOO (\<approx>)) r s" |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
149 |
then have b: "map Abs r \<approx> map Abs s" |
47434
b75ce48a93ee
dropped abbreviation "pred_comp"; introduced infix notation "P OO Q" for "relcompp P Q"
griff
parents:
47198
diff
changeset
|
150 |
proof (elim relcomppE) |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
151 |
fix b ba |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
152 |
assume c: "list_all2 R r b" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
153 |
assume d: "b \<approx> ba" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
154 |
assume e: "list_all2 R ba s" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
155 |
have f: "map Abs r = map Abs b" |
47308 | 156 |
using Quotient3_rel[OF list_quotient3[OF q]] c by blast |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
157 |
have "map Abs ba = map Abs s" |
47308 | 158 |
using Quotient3_rel[OF list_quotient3[OF q]] e by blast |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
159 |
then have g: "map Abs s = map Abs ba" by simp |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
160 |
then show "map Abs r \<approx> map Abs s" using d f map_list_eq_cong by simp |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
161 |
qed |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
162 |
then show "abs_fset (map Abs r) = abs_fset (map Abs s)" |
47308 | 163 |
using Quotient3_rel[OF Quotient3_fset] by blast |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
164 |
next |
67399 | 165 |
assume a: "(list_all2 R OOO (\<approx>)) r r \<and> (list_all2 R OOO (\<approx>)) s s |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
166 |
\<and> abs_fset (map Abs r) = abs_fset (map Abs s)" |
67399 | 167 |
then have s: "(list_all2 R OOO (\<approx>)) s s" by simp |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
168 |
have d: "map Abs r \<approx> map Abs s" |
47308 | 169 |
by (subst Quotient3_rel [OF Quotient3_fset, symmetric]) (simp add: a) |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
170 |
have b: "map Rep (map Abs r) \<approx> map Rep (map Abs s)" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
171 |
by (rule map_list_eq_cong[OF d]) |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
172 |
have y: "list_all2 R (map Rep (map Abs s)) s" |
47308 | 173 |
by (fact rep_abs_rsp_left[OF list_quotient3[OF q], OF list_all2_refl'[OF e, of s]]) |
67399 | 174 |
have c: "((\<approx>) OO list_all2 R) (map Rep (map Abs r)) s" |
47434
b75ce48a93ee
dropped abbreviation "pred_comp"; introduced infix notation "P OO Q" for "relcompp P Q"
griff
parents:
47198
diff
changeset
|
175 |
by (rule relcomppI) (rule b, rule y) |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
176 |
have z: "list_all2 R r (map Rep (map Abs r))" |
47308 | 177 |
by (fact rep_abs_rsp[OF list_quotient3[OF q], OF list_all2_refl'[OF e, of r]]) |
67399 | 178 |
then show "(list_all2 R OOO (\<approx>)) r s" |
47434
b75ce48a93ee
dropped abbreviation "pred_comp"; introduced infix notation "P OO Q" for "relcompp P Q"
griff
parents:
47198
diff
changeset
|
179 |
using a c relcomppI by simp |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
180 |
qed |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
181 |
qed |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
182 |
|
36280 | 183 |
lemma quotient_compose_list[quot_thm]: |
67399 | 184 |
shows "Quotient3 ((list_all2 (\<approx>)) OOO (\<approx>)) |
36280 | 185 |
(abs_fset \<circ> (map abs_fset)) ((map rep_fset) \<circ> rep_fset)" |
47308 | 186 |
by (rule quotient_compose_list_g, rule Quotient3_fset, rule list_eq_equivp) |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
187 |
|
36280 | 188 |
|
63167 | 189 |
section \<open>Quotient definitions for fsets\<close> |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
190 |
|
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
191 |
|
63167 | 192 |
subsection \<open>Finite sets are a bounded, distributive lattice with minus\<close> |
36280 | 193 |
|
37634
2116425cebc8
cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents:
37492
diff
changeset
|
194 |
instantiation fset :: (type) "{bounded_lattice_bot, distrib_lattice, minus}" |
36280 | 195 |
begin |
196 |
||
197 |
quotient_definition |
|
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
198 |
"bot :: 'a fset" |
47092 | 199 |
is "Nil :: 'a list" done |
36280 | 200 |
|
201 |
abbreviation |
|
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80786
diff
changeset
|
202 |
empty_fset (\<open>{||}\<close>) |
36280 | 203 |
where |
204 |
"{||} \<equiv> bot :: 'a fset" |
|
205 |
||
206 |
quotient_definition |
|
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
207 |
"less_eq_fset :: ('a fset \<Rightarrow> 'a fset \<Rightarrow> bool)" |
47092 | 208 |
is "sub_list :: ('a list \<Rightarrow> 'a list \<Rightarrow> bool)" by simp |
36280 | 209 |
|
210 |
abbreviation |
|
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80786
diff
changeset
|
211 |
subset_fset :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> bool" (infix \<open>|\<subseteq>|\<close> 50) |
36280 | 212 |
where |
213 |
"xs |\<subseteq>| ys \<equiv> xs \<le> ys" |
|
214 |
||
215 |
definition |
|
39995 | 216 |
less_fset :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> bool" |
217 |
where |
|
218 |
"xs < ys \<equiv> xs \<le> ys \<and> xs \<noteq> (ys::'a fset)" |
|
36280 | 219 |
|
220 |
abbreviation |
|
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80786
diff
changeset
|
221 |
psubset_fset :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> bool" (infix \<open>|\<subset>|\<close> 50) |
36280 | 222 |
where |
223 |
"xs |\<subset>| ys \<equiv> xs < ys" |
|
224 |
||
225 |
quotient_definition |
|
39995 | 226 |
"sup :: 'a fset \<Rightarrow> 'a fset \<Rightarrow> 'a fset" |
47092 | 227 |
is "append :: 'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" by simp |
36280 | 228 |
|
229 |
abbreviation |
|
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80786
diff
changeset
|
230 |
union_fset (infixl \<open>|\<union>|\<close> 65) |
36280 | 231 |
where |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
232 |
"xs |\<union>| ys \<equiv> sup xs (ys::'a fset)" |
36280 | 233 |
|
234 |
quotient_definition |
|
39995 | 235 |
"inf :: 'a fset \<Rightarrow> 'a fset \<Rightarrow> 'a fset" |
47092 | 236 |
is "inter_list :: 'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" by simp |
36280 | 237 |
|
238 |
abbreviation |
|
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80786
diff
changeset
|
239 |
inter_fset (infixl \<open>|\<inter>|\<close> 65) |
36280 | 240 |
where |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
241 |
"xs |\<inter>| ys \<equiv> inf xs (ys::'a fset)" |
36280 | 242 |
|
36675
806ea6e282e4
fminus and some more theorems ported from Finite_Set.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
36646
diff
changeset
|
243 |
quotient_definition |
37634
2116425cebc8
cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents:
37492
diff
changeset
|
244 |
"minus :: 'a fset \<Rightarrow> 'a fset \<Rightarrow> 'a fset" |
47092 | 245 |
is "diff_list :: 'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" by fastforce |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
246 |
|
36280 | 247 |
instance |
248 |
proof |
|
249 |
fix x y z :: "'a fset" |
|
37634
2116425cebc8
cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents:
37492
diff
changeset
|
250 |
show "x |\<subset>| y \<longleftrightarrow> x |\<subseteq>| y \<and> \<not> y |\<subseteq>| x" |
40467
dc0439fdd7c5
more appropriate specification packages; fun_rel_def is no simp rule by default
haftmann
parents:
40034
diff
changeset
|
251 |
by (unfold less_fset_def, descending) auto |
46441
992a1688303f
Generalize the compositional preservation theorems
Cezary Kaliszyk <cezarykaliszyk@gmail.com>
parents:
46416
diff
changeset
|
252 |
show "x |\<subseteq>| x" by (descending) (simp) |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
253 |
show "{||} |\<subseteq>| x" by (descending) (simp) |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
254 |
show "x |\<subseteq>| x |\<union>| y" by (descending) (simp) |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
255 |
show "y |\<subseteq>| x |\<union>| y" by (descending) (simp) |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
256 |
show "x |\<inter>| y |\<subseteq>| x" by (descending) (auto) |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
257 |
show "x |\<inter>| y |\<subseteq>| y" by (descending) (auto) |
37634
2116425cebc8
cleaned by using descending instead of lifting
Christian Urban <urbanc@in.tum.de>
parents:
37492
diff
changeset
|
258 |
show "x |\<union>| (y |\<inter>| z) = x |\<union>| y |\<inter>| (x |\<union>| z)" |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
259 |
by (descending) (auto) |
36280 | 260 |
next |
261 |
fix x y z :: "'a fset" |
|
262 |
assume a: "x |\<subseteq>| y" |
|
263 |
assume b: "y |\<subseteq>| z" |
|
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
264 |
show "x |\<subseteq>| z" using a b by (descending) (simp) |
36280 | 265 |
next |
266 |
fix x y :: "'a fset" |
|
267 |
assume a: "x |\<subseteq>| y" |
|
268 |
assume b: "y |\<subseteq>| x" |
|
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
269 |
show "x = y" using a b by (descending) (auto) |
36280 | 270 |
next |
271 |
fix x y z :: "'a fset" |
|
272 |
assume a: "y |\<subseteq>| x" |
|
273 |
assume b: "z |\<subseteq>| x" |
|
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
274 |
show "y |\<union>| z |\<subseteq>| x" using a b by (descending) (simp) |
36280 | 275 |
next |
276 |
fix x y z :: "'a fset" |
|
277 |
assume a: "x |\<subseteq>| y" |
|
278 |
assume b: "x |\<subseteq>| z" |
|
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
279 |
show "x |\<subseteq>| y |\<inter>| z" using a b by (descending) (auto) |
36280 | 280 |
qed |
281 |
||
282 |
end |
|
283 |
||
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
284 |
|
63167 | 285 |
subsection \<open>Other constants for fsets\<close> |
36280 | 286 |
|
287 |
quotient_definition |
|
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
288 |
"insert_fset :: 'a \<Rightarrow> 'a fset \<Rightarrow> 'a fset" |
47092 | 289 |
is "Cons" by auto |
36280 | 290 |
|
291 |
syntax |
|
81090 | 292 |
"_fset" :: "args => 'a fset" (\<open>(\<open>indent=2 notation=\<open>mixfix finite set enumeration\<close>\<close>{|_|})\<close>) |
80768 | 293 |
syntax_consts |
81019
dd59daa3c37a
clarified inner-syntax markup, notably for enumerations: prefer "notation=mixfix" over "entity" via 'syntax_consts' (see also 70076ba563d2);
wenzelm
parents:
80914
diff
changeset
|
294 |
insert_fset |
36280 | 295 |
translations |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
296 |
"{|x, xs|}" == "CONST insert_fset x {|xs|}" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
297 |
"{|x|}" == "CONST insert_fset x {||}" |
36280 | 298 |
|
299 |
quotient_definition |
|
40953 | 300 |
fset_member |
36280 | 301 |
where |
47092 | 302 |
"fset_member :: 'a fset \<Rightarrow> 'a \<Rightarrow> bool" is "List.member" by fastforce |
40953 | 303 |
|
304 |
abbreviation |
|
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80786
diff
changeset
|
305 |
in_fset :: "'a \<Rightarrow> 'a fset \<Rightarrow> bool" (infix \<open>|\<in>|\<close> 50) |
40953 | 306 |
where |
307 |
"x |\<in>| S \<equiv> fset_member S x" |
|
36280 | 308 |
|
309 |
abbreviation |
|
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80786
diff
changeset
|
310 |
notin_fset :: "'a \<Rightarrow> 'a fset \<Rightarrow> bool" (infix \<open>|\<notin>|\<close> 50) |
36280 | 311 |
where |
312 |
"x |\<notin>| S \<equiv> \<not> (x |\<in>| S)" |
|
313 |
||
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
314 |
|
63167 | 315 |
subsection \<open>Other constants on the Quotient Type\<close> |
36280 | 316 |
|
317 |
quotient_definition |
|
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
318 |
"card_fset :: 'a fset \<Rightarrow> nat" |
47092 | 319 |
is card_list by simp |
36280 | 320 |
|
321 |
quotient_definition |
|
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
322 |
"map_fset :: ('a \<Rightarrow> 'b) \<Rightarrow> 'a fset \<Rightarrow> 'b fset" |
47092 | 323 |
is map by simp |
36280 | 324 |
|
325 |
quotient_definition |
|
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
326 |
"remove_fset :: 'a \<Rightarrow> 'a fset \<Rightarrow> 'a fset" |
47092 | 327 |
is removeAll by simp |
36280 | 328 |
|
329 |
quotient_definition |
|
39996
c02078ff8691
FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
39995
diff
changeset
|
330 |
"fset :: 'a fset \<Rightarrow> 'a set" |
47092 | 331 |
is "set" by simp |
332 |
||
333 |
lemma fold_once_set_equiv: |
|
334 |
assumes "xs \<approx> ys" |
|
335 |
shows "fold_once f xs = fold_once f ys" |
|
336 |
proof (cases "rsp_fold f") |
|
337 |
case False then show ?thesis by simp |
|
338 |
next |
|
339 |
case True |
|
340 |
then have "\<And>x y. x \<in> set (remdups xs) \<Longrightarrow> y \<in> set (remdups xs) \<Longrightarrow> f x \<circ> f y = f y \<circ> f x" |
|
341 |
by (rule rsp_foldE) |
|
60515 | 342 |
moreover from assms have "mset (remdups xs) = mset (remdups ys)" |
343 |
by (simp add: set_eq_iff_mset_remdups_eq) |
|
47092 | 344 |
ultimately have "fold f (remdups xs) = fold f (remdups ys)" |
345 |
by (rule fold_multiset_equiv) |
|
346 |
with True show ?thesis by (simp add: fold_once_fold_remdups) |
|
347 |
qed |
|
36280 | 348 |
|
349 |
quotient_definition |
|
40961 | 350 |
"fold_fset :: ('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'a fset \<Rightarrow> 'b \<Rightarrow> 'b" |
47092 | 351 |
is fold_once by (rule fold_once_set_equiv) |
352 |
||
353 |
lemma concat_rsp_pre: |
|
67399 | 354 |
assumes a: "list_all2 (\<approx>) x x'" |
47092 | 355 |
and b: "x' \<approx> y'" |
67399 | 356 |
and c: "list_all2 (\<approx>) y' y" |
47092 | 357 |
and d: "\<exists>x\<in>set x. xa \<in> set x" |
358 |
shows "\<exists>x\<in>set y. xa \<in> set x" |
|
359 |
proof - |
|
360 |
obtain xb where e: "xb \<in> set x" and f: "xa \<in> set xb" using d by auto |
|
361 |
have "\<exists>y. y \<in> set x' \<and> xb \<approx> y" by (rule list_all2_find_element[OF e a]) |
|
362 |
then obtain ya where h: "ya \<in> set x'" and i: "xb \<approx> ya" by auto |
|
363 |
have "ya \<in> set y'" using b h by simp |
|
364 |
then have "\<exists>yb. yb \<in> set y \<and> ya \<approx> yb" using c by (rule list_all2_find_element) |
|
365 |
then show ?thesis using f i by auto |
|
366 |
qed |
|
36280 | 367 |
|
368 |
quotient_definition |
|
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
369 |
"concat_fset :: ('a fset) fset \<Rightarrow> 'a fset" |
47092 | 370 |
is concat |
47434
b75ce48a93ee
dropped abbreviation "pred_comp"; introduced infix notation "P OO Q" for "relcompp P Q"
griff
parents:
47198
diff
changeset
|
371 |
proof (elim relcomppE) |
47092 | 372 |
fix a b ba bb |
67399 | 373 |
assume a: "list_all2 (\<approx>) a ba" |
374 |
with list_symp [OF list_eq_symp] have a': "list_all2 (\<approx>) ba a" by (rule sympE) |
|
47092 | 375 |
assume b: "ba \<approx> bb" |
376 |
with list_eq_symp have b': "bb \<approx> ba" by (rule sympE) |
|
67399 | 377 |
assume c: "list_all2 (\<approx>) bb b" |
378 |
with list_symp [OF list_eq_symp] have c': "list_all2 (\<approx>) b bb" by (rule sympE) |
|
47092 | 379 |
have "\<forall>x. (\<exists>xa\<in>set a. x \<in> set xa) = (\<exists>xa\<in>set b. x \<in> set xa)" |
380 |
proof |
|
381 |
fix x |
|
382 |
show "(\<exists>xa\<in>set a. x \<in> set xa) = (\<exists>xa\<in>set b. x \<in> set xa)" |
|
383 |
proof |
|
384 |
assume d: "\<exists>xa\<in>set a. x \<in> set xa" |
|
385 |
show "\<exists>xa\<in>set b. x \<in> set xa" by (rule concat_rsp_pre[OF a b c d]) |
|
386 |
next |
|
387 |
assume e: "\<exists>xa\<in>set b. x \<in> set xa" |
|
388 |
show "\<exists>xa\<in>set a. x \<in> set xa" by (rule concat_rsp_pre[OF c' b' a' e]) |
|
389 |
qed |
|
390 |
qed |
|
391 |
then show "concat a \<approx> concat b" by auto |
|
392 |
qed |
|
36280 | 393 |
|
36639
6243b49498ea
added function ffilter and some lemmas from Finite_Set to the FSet theory
bulwahn
parents:
36524
diff
changeset
|
394 |
quotient_definition |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
395 |
"filter_fset :: ('a \<Rightarrow> bool) \<Rightarrow> 'a fset \<Rightarrow> 'a fset" |
47092 | 396 |
is filter by force |
36639
6243b49498ea
added function ffilter and some lemmas from Finite_Set to the FSet theory
bulwahn
parents:
36524
diff
changeset
|
397 |
|
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
398 |
|
63167 | 399 |
subsection \<open>Compositional respectfulness and preservation lemmas\<close> |
36280 | 400 |
|
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
401 |
lemma Nil_rsp2 [quot_respect]: |
67399 | 402 |
shows "(list_all2 (\<approx>) OOO (\<approx>)) Nil Nil" |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
403 |
by (rule compose_list_refl, rule list_eq_equivp) |
36280 | 404 |
|
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
405 |
lemma Cons_rsp2 [quot_respect]: |
67399 | 406 |
shows "((\<approx>) ===> list_all2 (\<approx>) OOO (\<approx>) ===> list_all2 (\<approx>) OOO (\<approx>)) Cons Cons" |
55945 | 407 |
apply (auto intro!: rel_funI) |
47434
b75ce48a93ee
dropped abbreviation "pred_comp"; introduced infix notation "P OO Q" for "relcompp P Q"
griff
parents:
47198
diff
changeset
|
408 |
apply (rule_tac b="x # b" in relcomppI) |
36280 | 409 |
apply auto |
47434
b75ce48a93ee
dropped abbreviation "pred_comp"; introduced infix notation "P OO Q" for "relcompp P Q"
griff
parents:
47198
diff
changeset
|
410 |
apply (rule_tac b="x # ba" in relcomppI) |
36280 | 411 |
apply auto |
412 |
done |
|
413 |
||
46441
992a1688303f
Generalize the compositional preservation theorems
Cezary Kaliszyk <cezarykaliszyk@gmail.com>
parents:
46416
diff
changeset
|
414 |
lemma Nil_prs2 [quot_preserve]: |
47308 | 415 |
assumes "Quotient3 R Abs Rep" |
46441
992a1688303f
Generalize the compositional preservation theorems
Cezary Kaliszyk <cezarykaliszyk@gmail.com>
parents:
46416
diff
changeset
|
416 |
shows "(Abs \<circ> map f) [] = Abs []" |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
417 |
by simp |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
418 |
|
46441
992a1688303f
Generalize the compositional preservation theorems
Cezary Kaliszyk <cezarykaliszyk@gmail.com>
parents:
46416
diff
changeset
|
419 |
lemma Cons_prs2 [quot_preserve]: |
47308 | 420 |
assumes q: "Quotient3 R1 Abs1 Rep1" |
421 |
and r: "Quotient3 R2 Abs2 Rep2" |
|
67399 | 422 |
shows "(Rep1 ---> (map Rep1 \<circ> Rep2) ---> (Abs2 \<circ> map Abs1)) (#) = (id ---> Rep2 ---> Abs2) (#)" |
47308 | 423 |
by (auto simp add: fun_eq_iff comp_def Quotient3_abs_rep [OF q]) |
36280 | 424 |
|
46441
992a1688303f
Generalize the compositional preservation theorems
Cezary Kaliszyk <cezarykaliszyk@gmail.com>
parents:
46416
diff
changeset
|
425 |
lemma append_prs2 [quot_preserve]: |
47308 | 426 |
assumes q: "Quotient3 R1 Abs1 Rep1" |
427 |
and r: "Quotient3 R2 Abs2 Rep2" |
|
67399 | 428 |
shows "((map Rep1 \<circ> Rep2) ---> (map Rep1 \<circ> Rep2) ---> (Abs2 \<circ> map Abs1)) (@) = |
429 |
(Rep2 ---> Rep2 ---> Abs2) (@)" |
|
46663 | 430 |
by (simp add: fun_eq_iff abs_o_rep[OF q] List.map.id) |
36280 | 431 |
|
37492
ab36b1a50ca8
Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
36675
diff
changeset
|
432 |
lemma list_all2_app_l: |
36280 | 433 |
assumes a: "reflp R" |
37492
ab36b1a50ca8
Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
36675
diff
changeset
|
434 |
and b: "list_all2 R l r" |
ab36b1a50ca8
Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
36675
diff
changeset
|
435 |
shows "list_all2 R (z @ l) (z @ r)" |
40822 | 436 |
using a b by (induct z) (auto elim: reflpE) |
36280 | 437 |
|
438 |
lemma append_rsp2_pre0: |
|
67399 | 439 |
assumes a:"list_all2 (\<approx>) x x'" |
440 |
shows "list_all2 (\<approx>) (x @ z) (x' @ z)" |
|
36280 | 441 |
using a apply (induct x x' rule: list_induct2') |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
442 |
by simp_all (rule list_all2_refl'[OF list_eq_equivp]) |
36280 | 443 |
|
444 |
lemma append_rsp2_pre1: |
|
67399 | 445 |
assumes a:"list_all2 (\<approx>) x x'" |
446 |
shows "list_all2 (\<approx>) (z @ x) (z @ x')" |
|
36280 | 447 |
using a apply (induct x x' arbitrary: z rule: list_induct2') |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
448 |
apply (rule list_all2_refl'[OF list_eq_equivp]) |
40467
dc0439fdd7c5
more appropriate specification packages; fun_rel_def is no simp rule by default
haftmann
parents:
40034
diff
changeset
|
449 |
apply (simp_all del: list_eq_def) |
37492
ab36b1a50ca8
Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
36675
diff
changeset
|
450 |
apply (rule list_all2_app_l) |
40822 | 451 |
apply (simp_all add: reflpI) |
36280 | 452 |
done |
453 |
||
454 |
lemma append_rsp2_pre: |
|
67399 | 455 |
assumes "list_all2 (\<approx>) x x'" |
456 |
and "list_all2 (\<approx>) z z'" |
|
457 |
shows "list_all2 (\<approx>) (x @ z) (x' @ z')" |
|
40822 | 458 |
using assms by (rule list_all2_appendI) |
36280 | 459 |
|
46441
992a1688303f
Generalize the compositional preservation theorems
Cezary Kaliszyk <cezarykaliszyk@gmail.com>
parents:
46416
diff
changeset
|
460 |
lemma compositional_rsp3: |
992a1688303f
Generalize the compositional preservation theorems
Cezary Kaliszyk <cezarykaliszyk@gmail.com>
parents:
46416
diff
changeset
|
461 |
assumes "(R1 ===> R2 ===> R3) C C" and "(R4 ===> R5 ===> R6) C C" |
992a1688303f
Generalize the compositional preservation theorems
Cezary Kaliszyk <cezarykaliszyk@gmail.com>
parents:
46416
diff
changeset
|
462 |
shows "(R1 OOO R4 ===> R2 OOO R5 ===> R3 OOO R6) C C" |
55945 | 463 |
by (auto intro!: rel_funI) |
464 |
(metis (full_types) assms rel_funE relcomppI) |
|
46441
992a1688303f
Generalize the compositional preservation theorems
Cezary Kaliszyk <cezarykaliszyk@gmail.com>
parents:
46416
diff
changeset
|
465 |
|
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
466 |
lemma append_rsp2 [quot_respect]: |
67399 | 467 |
"(list_all2 (\<approx>) OOO (\<approx>) ===> list_all2 (\<approx>) OOO (\<approx>) ===> list_all2 (\<approx>) OOO (\<approx>)) append append" |
47092 | 468 |
by (intro compositional_rsp3) |
55945 | 469 |
(auto intro!: rel_funI simp add: append_rsp2_pre) |
36280 | 470 |
|
46404
7736068b9f56
Quotient FSet: Add compositional respectfulness and preservation for map and lift map_concat
Cezary Kaliszyk <cezarykaliszyk@gmail.com>
parents:
46133
diff
changeset
|
471 |
lemma map_rsp2 [quot_respect]: |
67399 | 472 |
"(((\<approx>) ===> (\<approx>)) ===> list_all2 (\<approx>) OOO (\<approx>) ===> list_all2 (\<approx>) OOO (\<approx>)) map map" |
55945 | 473 |
proof (auto intro!: rel_funI) |
46404
7736068b9f56
Quotient FSet: Add compositional respectfulness and preservation for map and lift map_concat
Cezary Kaliszyk <cezarykaliszyk@gmail.com>
parents:
46133
diff
changeset
|
474 |
fix f f' :: "'a list \<Rightarrow> 'b list" |
7736068b9f56
Quotient FSet: Add compositional respectfulness and preservation for map and lift map_concat
Cezary Kaliszyk <cezarykaliszyk@gmail.com>
parents:
46133
diff
changeset
|
475 |
fix xa ya x y :: "'a list list" |
67399 | 476 |
assume fs: "((\<approx>) ===> (\<approx>)) f f'" and x: "list_all2 (\<approx>) xa x" and xy: "set x = set y" and y: "list_all2 (\<approx>) y ya" |
477 |
have a: "(list_all2 (\<approx>)) (map f xa) (map f x)" |
|
46404
7736068b9f56
Quotient FSet: Add compositional respectfulness and preservation for map and lift map_concat
Cezary Kaliszyk <cezarykaliszyk@gmail.com>
parents:
46133
diff
changeset
|
478 |
using x |
7736068b9f56
Quotient FSet: Add compositional respectfulness and preservation for map and lift map_concat
Cezary Kaliszyk <cezarykaliszyk@gmail.com>
parents:
46133
diff
changeset
|
479 |
by (induct xa x rule: list_induct2') |
55945 | 480 |
(simp_all, metis fs rel_funE list_eq_def) |
46404
7736068b9f56
Quotient FSet: Add compositional respectfulness and preservation for map and lift map_concat
Cezary Kaliszyk <cezarykaliszyk@gmail.com>
parents:
46133
diff
changeset
|
481 |
have b: "set (map f x) = set (map f y)" |
7736068b9f56
Quotient FSet: Add compositional respectfulness and preservation for map and lift map_concat
Cezary Kaliszyk <cezarykaliszyk@gmail.com>
parents:
46133
diff
changeset
|
482 |
using xy fs |
7736068b9f56
Quotient FSet: Add compositional respectfulness and preservation for map and lift map_concat
Cezary Kaliszyk <cezarykaliszyk@gmail.com>
parents:
46133
diff
changeset
|
483 |
by (induct x y rule: list_induct2') |
7736068b9f56
Quotient FSet: Add compositional respectfulness and preservation for map and lift map_concat
Cezary Kaliszyk <cezarykaliszyk@gmail.com>
parents:
46133
diff
changeset
|
484 |
(simp_all, metis image_insert) |
67399 | 485 |
have c: "(list_all2 (\<approx>)) (map f y) (map f' ya)" |
46404
7736068b9f56
Quotient FSet: Add compositional respectfulness and preservation for map and lift map_concat
Cezary Kaliszyk <cezarykaliszyk@gmail.com>
parents:
46133
diff
changeset
|
486 |
using y fs |
7736068b9f56
Quotient FSet: Add compositional respectfulness and preservation for map and lift map_concat
Cezary Kaliszyk <cezarykaliszyk@gmail.com>
parents:
46133
diff
changeset
|
487 |
by (induct y ya rule: list_induct2') |
7736068b9f56
Quotient FSet: Add compositional respectfulness and preservation for map and lift map_concat
Cezary Kaliszyk <cezarykaliszyk@gmail.com>
parents:
46133
diff
changeset
|
488 |
(simp_all, metis apply_rsp' list_eq_def) |
67399 | 489 |
show "(list_all2 (\<approx>) OOO (\<approx>)) (map f xa) (map f' ya)" |
47434
b75ce48a93ee
dropped abbreviation "pred_comp"; introduced infix notation "P OO Q" for "relcompp P Q"
griff
parents:
47198
diff
changeset
|
490 |
by (metis a b c list_eq_def relcomppI) |
46404
7736068b9f56
Quotient FSet: Add compositional respectfulness and preservation for map and lift map_concat
Cezary Kaliszyk <cezarykaliszyk@gmail.com>
parents:
46133
diff
changeset
|
491 |
qed |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
492 |
|
46404
7736068b9f56
Quotient FSet: Add compositional respectfulness and preservation for map and lift map_concat
Cezary Kaliszyk <cezarykaliszyk@gmail.com>
parents:
46133
diff
changeset
|
493 |
lemma map_prs2 [quot_preserve]: |
46441
992a1688303f
Generalize the compositional preservation theorems
Cezary Kaliszyk <cezarykaliszyk@gmail.com>
parents:
46416
diff
changeset
|
494 |
shows "((abs_fset ---> rep_fset) ---> (map rep_fset \<circ> rep_fset) ---> abs_fset \<circ> map abs_fset) map = (id ---> rep_fset ---> abs_fset) map" |
992a1688303f
Generalize the compositional preservation theorems
Cezary Kaliszyk <cezarykaliszyk@gmail.com>
parents:
46416
diff
changeset
|
495 |
by (auto simp add: fun_eq_iff) |
47308 | 496 |
(simp only: map_map[symmetric] map_prs_aux[OF Quotient3_fset Quotient3_fset]) |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
497 |
|
63167 | 498 |
section \<open>Lifted theorems\<close> |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
499 |
|
63167 | 500 |
subsection \<open>fset\<close> |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
501 |
|
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
502 |
lemma fset_simps [simp]: |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
503 |
shows "fset {||} = {}" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
504 |
and "fset (insert_fset x S) = insert x (fset S)" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
505 |
by (descending, simp)+ |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
506 |
|
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
507 |
lemma finite_fset [simp]: |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
508 |
shows "finite (fset S)" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
509 |
by (descending) (simp) |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
510 |
|
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
511 |
lemma fset_cong: |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
512 |
shows "fset S = fset T \<longleftrightarrow> S = T" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
513 |
by (descending) (simp) |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
514 |
|
44204
3cdc4176638c
Quotient Package: make quotient_type work with separate set type
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
41467
diff
changeset
|
515 |
lemma filter_fset [simp]: |
3cdc4176638c
Quotient Package: make quotient_type work with separate set type
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
41467
diff
changeset
|
516 |
shows "fset (filter_fset P xs) = Collect P \<inter> fset xs" |
3cdc4176638c
Quotient Package: make quotient_type work with separate set type
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
41467
diff
changeset
|
517 |
by (descending) (auto) |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
518 |
|
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
519 |
lemma remove_fset [simp]: |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
520 |
shows "fset (remove_fset x xs) = fset xs - {x}" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
521 |
by (descending) (simp) |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
522 |
|
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
523 |
lemma inter_fset [simp]: |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
524 |
shows "fset (xs |\<inter>| ys) = fset xs \<inter> fset ys" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
525 |
by (descending) (auto) |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
526 |
|
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
527 |
lemma union_fset [simp]: |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
528 |
shows "fset (xs |\<union>| ys) = fset xs \<union> fset ys" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
529 |
by (lifting set_append) |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
530 |
|
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
531 |
lemma minus_fset [simp]: |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
532 |
shows "fset (xs - ys) = fset xs - fset ys" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
533 |
by (descending) (auto) |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
534 |
|
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
535 |
|
63167 | 536 |
subsection \<open>in_fset\<close> |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
537 |
|
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
538 |
lemma in_fset: |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
539 |
shows "x |\<in>| S \<longleftrightarrow> x \<in> fset S" |
40953 | 540 |
by descending simp |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
541 |
|
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
542 |
lemma notin_fset: |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
543 |
shows "x |\<notin>| S \<longleftrightarrow> x \<notin> fset S" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
544 |
by (simp add: in_fset) |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
545 |
|
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
546 |
lemma notin_empty_fset: |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
547 |
shows "x |\<notin>| {||}" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
548 |
by (simp add: in_fset) |
36280 | 549 |
|
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
550 |
lemma fset_eq_iff: |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
551 |
shows "S = T \<longleftrightarrow> (\<forall>x. (x |\<in>| S) = (x |\<in>| T))" |
40953 | 552 |
by descending auto |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
553 |
|
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
554 |
lemma none_in_empty_fset: |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
555 |
shows "(\<forall>x. x |\<notin>| S) \<longleftrightarrow> S = {||}" |
40953 | 556 |
by descending simp |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
557 |
|
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
558 |
|
63167 | 559 |
subsection \<open>insert_fset\<close> |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
560 |
|
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
561 |
lemma in_insert_fset_iff [simp]: |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
562 |
shows "x |\<in>| insert_fset y S \<longleftrightarrow> x = y \<or> x |\<in>| S" |
40953 | 563 |
by descending simp |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
564 |
|
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
565 |
lemma |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
566 |
shows insert_fsetI1: "x |\<in>| insert_fset x S" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
567 |
and insert_fsetI2: "x |\<in>| S \<Longrightarrow> x |\<in>| insert_fset y S" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
568 |
by simp_all |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
569 |
|
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
570 |
lemma insert_absorb_fset [simp]: |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
571 |
shows "x |\<in>| S \<Longrightarrow> insert_fset x S = S" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
572 |
by (descending) (auto) |
36280 | 573 |
|
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
574 |
lemma empty_not_insert_fset[simp]: |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
575 |
shows "{||} \<noteq> insert_fset x S" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
576 |
and "insert_fset x S \<noteq> {||}" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
577 |
by (descending, simp)+ |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
578 |
|
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
579 |
lemma insert_fset_left_comm: |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
580 |
shows "insert_fset x (insert_fset y S) = insert_fset y (insert_fset x S)" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
581 |
by (descending) (auto) |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
582 |
|
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
583 |
lemma insert_fset_left_idem: |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
584 |
shows "insert_fset x (insert_fset x S) = insert_fset x S" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
585 |
by (descending) (auto) |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
586 |
|
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
587 |
lemma singleton_fset_eq[simp]: |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
588 |
shows "{|x|} = {|y|} \<longleftrightarrow> x = y" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
589 |
by (descending) (auto) |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
590 |
|
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
591 |
lemma in_fset_mdef: |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
592 |
shows "x |\<in>| F \<longleftrightarrow> x |\<notin>| (F - {|x|}) \<and> F = insert_fset x (F - {|x|})" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
593 |
by (descending) (auto) |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
594 |
|
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
595 |
|
63167 | 596 |
subsection \<open>union_fset\<close> |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
597 |
|
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
598 |
lemmas [simp] = |
45605 | 599 |
sup_bot_left[where 'a="'a fset"] |
600 |
sup_bot_right[where 'a="'a fset"] |
|
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
601 |
|
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
602 |
lemma union_insert_fset [simp]: |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
603 |
shows "insert_fset x S |\<union>| T = insert_fset x (S |\<union>| T)" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
604 |
by (lifting append.simps(2)) |
36280 | 605 |
|
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
606 |
lemma singleton_union_fset_left: |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
607 |
shows "{|a|} |\<union>| S = insert_fset a S" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
608 |
by simp |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
609 |
|
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
610 |
lemma singleton_union_fset_right: |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
611 |
shows "S |\<union>| {|a|} = insert_fset a S" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
612 |
by (subst sup.commute) simp |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
613 |
|
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
614 |
lemma in_union_fset: |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
615 |
shows "x |\<in>| S |\<union>| T \<longleftrightarrow> x |\<in>| S \<or> x |\<in>| T" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
616 |
by (descending) (simp) |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
617 |
|
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
618 |
|
63167 | 619 |
subsection \<open>minus_fset\<close> |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
620 |
|
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
621 |
lemma minus_in_fset: |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
622 |
shows "x |\<in>| (xs - ys) \<longleftrightarrow> x |\<in>| xs \<and> x |\<notin>| ys" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
623 |
by (descending) (simp) |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
624 |
|
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
625 |
lemma minus_insert_fset: |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
626 |
shows "insert_fset x xs - ys = (if x |\<in>| ys then xs - ys else insert_fset x (xs - ys))" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
627 |
by (descending) (auto) |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
628 |
|
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
629 |
lemma minus_insert_in_fset[simp]: |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
630 |
shows "x |\<in>| ys \<Longrightarrow> insert_fset x xs - ys = xs - ys" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
631 |
by (simp add: minus_insert_fset) |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
632 |
|
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
633 |
lemma minus_insert_notin_fset[simp]: |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
634 |
shows "x |\<notin>| ys \<Longrightarrow> insert_fset x xs - ys = insert_fset x (xs - ys)" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
635 |
by (simp add: minus_insert_fset) |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
636 |
|
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
637 |
lemma in_minus_fset: |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
638 |
shows "x |\<in>| F - S \<Longrightarrow> x |\<notin>| S" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
639 |
unfolding in_fset minus_fset |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
640 |
by blast |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
641 |
|
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
642 |
lemma notin_minus_fset: |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
643 |
shows "x |\<in>| S \<Longrightarrow> x |\<notin>| F - S" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
644 |
unfolding in_fset minus_fset |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
645 |
by blast |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
646 |
|
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
647 |
|
63167 | 648 |
subsection \<open>remove_fset\<close> |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
649 |
|
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
650 |
lemma in_remove_fset: |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
651 |
shows "x |\<in>| remove_fset y S \<longleftrightarrow> x |\<in>| S \<and> x \<noteq> y" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
652 |
by (descending) (simp) |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
653 |
|
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
654 |
lemma notin_remove_fset: |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
655 |
shows "x |\<notin>| remove_fset x S" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
656 |
by (descending) (simp) |
36280 | 657 |
|
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
658 |
lemma notin_remove_ident_fset: |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
659 |
shows "x |\<notin>| S \<Longrightarrow> remove_fset x S = S" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
660 |
by (descending) (simp) |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
661 |
|
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
662 |
lemma remove_fset_cases: |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
663 |
shows "S = {||} \<or> (\<exists>x. x |\<in>| S \<and> S = insert_fset x (remove_fset x S))" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
664 |
by (descending) (auto simp add: insert_absorb) |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
665 |
|
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
666 |
|
63167 | 667 |
subsection \<open>inter_fset\<close> |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
668 |
|
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
669 |
lemma inter_empty_fset_l: |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
670 |
shows "{||} |\<inter>| S = {||}" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
671 |
by simp |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
672 |
|
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
673 |
lemma inter_empty_fset_r: |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
674 |
shows "S |\<inter>| {||} = {||}" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
675 |
by simp |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
676 |
|
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
677 |
lemma inter_insert_fset: |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
678 |
shows "insert_fset x S |\<inter>| T = (if x |\<in>| T then insert_fset x (S |\<inter>| T) else S |\<inter>| T)" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
679 |
by (descending) (auto) |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
680 |
|
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
681 |
lemma in_inter_fset: |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
682 |
shows "x |\<in>| (S |\<inter>| T) \<longleftrightarrow> x |\<in>| S \<and> x |\<in>| T" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
683 |
by (descending) (simp) |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
684 |
|
36280 | 685 |
|
63167 | 686 |
subsection \<open>subset_fset and psubset_fset\<close> |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
687 |
|
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
688 |
lemma subset_fset: |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
689 |
shows "xs |\<subseteq>| ys \<longleftrightarrow> fset xs \<subseteq> fset ys" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
690 |
by (descending) (simp) |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
691 |
|
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
692 |
lemma psubset_fset: |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
693 |
shows "xs |\<subset>| ys \<longleftrightarrow> fset xs \<subset> fset ys" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
694 |
unfolding less_fset_def |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
695 |
by (descending) (auto) |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
696 |
|
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
697 |
lemma subset_insert_fset: |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
698 |
shows "(insert_fset x xs) |\<subseteq>| ys \<longleftrightarrow> x |\<in>| ys \<and> xs |\<subseteq>| ys" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
699 |
by (descending) (simp) |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
700 |
|
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
701 |
lemma subset_in_fset: |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
702 |
shows "xs |\<subseteq>| ys = (\<forall>x. x |\<in>| xs \<longrightarrow> x |\<in>| ys)" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
703 |
by (descending) (auto) |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
704 |
|
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
705 |
lemma subset_empty_fset: |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
706 |
shows "xs |\<subseteq>| {||} \<longleftrightarrow> xs = {||}" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
707 |
by (descending) (simp) |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
708 |
|
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
709 |
lemma not_psubset_empty_fset: |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
710 |
shows "\<not> xs |\<subset>| {||}" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
711 |
by (metis fset_simps(1) psubset_fset not_psubset_empty) |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
712 |
|
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
713 |
|
63167 | 714 |
subsection \<open>map_fset\<close> |
36280 | 715 |
|
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
716 |
lemma map_fset_simps [simp]: |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
717 |
shows "map_fset f {||} = {||}" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
718 |
and "map_fset f (insert_fset x S) = insert_fset (f x) (map_fset f S)" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
719 |
by (descending, simp)+ |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
720 |
|
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
721 |
lemma map_fset_image [simp]: |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
722 |
shows "fset (map_fset f S) = f ` (fset S)" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
723 |
by (descending) (simp) |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
724 |
|
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
725 |
lemma inj_map_fset_cong: |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
726 |
shows "inj f \<Longrightarrow> map_fset f S = map_fset f T \<longleftrightarrow> S = T" |
40467
dc0439fdd7c5
more appropriate specification packages; fun_rel_def is no simp rule by default
haftmann
parents:
40034
diff
changeset
|
727 |
by (descending) (metis inj_vimage_image_eq list_eq_def set_map) |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
728 |
|
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
729 |
lemma map_union_fset: |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
730 |
shows "map_fset f (S |\<union>| T) = map_fset f S |\<union>| map_fset f T" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
731 |
by (descending) (simp) |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
732 |
|
51371 | 733 |
lemma in_fset_map_fset[simp]: "a |\<in>| map_fset f X = (\<exists>b. b |\<in>| X \<and> a = f b)" |
734 |
by descending auto |
|
735 |
||
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
736 |
|
63167 | 737 |
subsection \<open>card_fset\<close> |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
738 |
|
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
739 |
lemma card_fset: |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
740 |
shows "card_fset xs = card (fset xs)" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
741 |
by (descending) (simp) |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
742 |
|
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
743 |
lemma card_insert_fset_iff [simp]: |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
744 |
shows "card_fset (insert_fset x S) = (if x |\<in>| S then card_fset S else Suc (card_fset S))" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
745 |
by (descending) (simp add: insert_absorb) |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
746 |
|
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
747 |
lemma card_fset_0[simp]: |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
748 |
shows "card_fset S = 0 \<longleftrightarrow> S = {||}" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
749 |
by (descending) (simp) |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
750 |
|
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
751 |
lemma card_empty_fset[simp]: |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
752 |
shows "card_fset {||} = 0" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
753 |
by (simp add: card_fset) |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
754 |
|
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
755 |
lemma card_fset_1: |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
756 |
shows "card_fset S = 1 \<longleftrightarrow> (\<exists>x. S = {|x|})" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
757 |
by (descending) (auto simp add: card_Suc_eq) |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
758 |
|
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
759 |
lemma card_fset_gt_0: |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
760 |
shows "x \<in> fset S \<Longrightarrow> 0 < card_fset S" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
761 |
by (descending) (auto simp add: card_gt_0_iff) |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
762 |
|
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
763 |
lemma card_notin_fset: |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
764 |
shows "(x |\<notin>| S) = (card_fset (insert_fset x S) = Suc (card_fset S))" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
765 |
by simp |
36280 | 766 |
|
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
767 |
lemma card_fset_Suc: |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
768 |
shows "card_fset S = Suc n \<Longrightarrow> \<exists>x T. x |\<notin>| T \<and> S = insert_fset x T \<and> card_fset T = n" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
769 |
apply(descending) |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
770 |
apply(auto dest!: card_eq_SucD) |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
771 |
by (metis Diff_insert_absorb set_removeAll) |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
772 |
|
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
773 |
lemma card_remove_fset_iff [simp]: |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
774 |
shows "card_fset (remove_fset y S) = (if y |\<in>| S then card_fset S - 1 else card_fset S)" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
775 |
by (descending) (simp) |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
776 |
|
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
777 |
lemma card_Suc_exists_in_fset: |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
778 |
shows "card_fset S = Suc n \<Longrightarrow> \<exists>a. a |\<in>| S" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
779 |
by (drule card_fset_Suc) (auto) |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
780 |
|
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
781 |
lemma in_card_fset_not_0: |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
782 |
shows "a |\<in>| A \<Longrightarrow> card_fset A \<noteq> 0" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
783 |
by (descending) (auto) |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
784 |
|
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
785 |
lemma card_fset_mono: |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
786 |
shows "xs |\<subseteq>| ys \<Longrightarrow> card_fset xs \<le> card_fset ys" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
787 |
unfolding card_fset psubset_fset |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
788 |
by (simp add: card_mono subset_fset) |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
789 |
|
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
790 |
lemma card_subset_fset_eq: |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
791 |
shows "xs |\<subseteq>| ys \<Longrightarrow> card_fset ys \<le> card_fset xs \<Longrightarrow> xs = ys" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
792 |
unfolding card_fset subset_fset |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
793 |
by (auto dest: card_seteq[OF finite_fset] simp add: fset_cong) |
36675
806ea6e282e4
fminus and some more theorems ported from Finite_Set.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
36646
diff
changeset
|
794 |
|
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
795 |
lemma psubset_card_fset_mono: |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
796 |
shows "xs |\<subset>| ys \<Longrightarrow> card_fset xs < card_fset ys" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
797 |
unfolding card_fset subset_fset |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
798 |
by (metis finite_fset psubset_fset psubset_card_mono) |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
799 |
|
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
800 |
lemma card_union_inter_fset: |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
801 |
shows "card_fset xs + card_fset ys = card_fset (xs |\<union>| ys) + card_fset (xs |\<inter>| ys)" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
802 |
unfolding card_fset union_fset inter_fset |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
803 |
by (rule card_Un_Int[OF finite_fset finite_fset]) |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
804 |
|
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
805 |
lemma card_union_disjoint_fset: |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
806 |
shows "xs |\<inter>| ys = {||} \<Longrightarrow> card_fset (xs |\<union>| ys) = card_fset xs + card_fset ys" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
807 |
unfolding card_fset union_fset |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
808 |
apply (rule card_Un_disjoint[OF finite_fset finite_fset]) |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
809 |
by (metis inter_fset fset_simps(1)) |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
810 |
|
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
811 |
lemma card_remove_fset_less1: |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
812 |
shows "x |\<in>| xs \<Longrightarrow> card_fset (remove_fset x xs) < card_fset xs" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
813 |
unfolding card_fset in_fset remove_fset |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
814 |
by (rule card_Diff1_less[OF finite_fset]) |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
815 |
|
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
816 |
lemma card_remove_fset_less2: |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
817 |
shows "x |\<in>| xs \<Longrightarrow> y |\<in>| xs \<Longrightarrow> card_fset (remove_fset y (remove_fset x xs)) < card_fset xs" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
818 |
unfolding card_fset remove_fset in_fset |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
819 |
by (rule card_Diff2_less[OF finite_fset]) |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
820 |
|
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
821 |
lemma card_remove_fset_le1: |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
822 |
shows "card_fset (remove_fset x xs) \<le> card_fset xs" |
74224
e04ec2b9ed97
some fixes connected with card_Diff_singleton
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
823 |
by simp |
36280 | 824 |
|
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
825 |
lemma card_psubset_fset: |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
826 |
shows "ys |\<subseteq>| xs \<Longrightarrow> card_fset ys < card_fset xs \<Longrightarrow> ys |\<subset>| xs" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
827 |
unfolding card_fset psubset_fset subset_fset |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
828 |
by (rule card_psubset[OF finite_fset]) |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
829 |
|
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
830 |
lemma card_map_fset_le: |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
831 |
shows "card_fset (map_fset f xs) \<le> card_fset xs" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
832 |
unfolding card_fset map_fset_image |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
833 |
by (rule card_image_le[OF finite_fset]) |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
834 |
|
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
835 |
lemma card_minus_insert_fset[simp]: |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
836 |
assumes "a |\<in>| A" and "a |\<notin>| B" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
837 |
shows "card_fset (A - insert_fset a B) = card_fset (A - B) - 1" |
74224
e04ec2b9ed97
some fixes connected with card_Diff_singleton
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
838 |
using assms by (simp add: in_fset card_fset) |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
839 |
|
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
840 |
lemma card_minus_subset_fset: |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
841 |
assumes "B |\<subseteq>| A" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
842 |
shows "card_fset (A - B) = card_fset A - card_fset B" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
843 |
using assms |
74224
e04ec2b9ed97
some fixes connected with card_Diff_singleton
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
844 |
by (simp add: subset_fset card_fset card_Diff_subset) |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
845 |
|
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
846 |
lemma card_minus_fset: |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
847 |
shows "card_fset (A - B) = card_fset A - card_fset (A |\<inter>| B)" |
74224
e04ec2b9ed97
some fixes connected with card_Diff_singleton
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
848 |
by (simp add: card_fset card_Diff_subset_Int) |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
849 |
|
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
850 |
|
63167 | 851 |
subsection \<open>concat_fset\<close> |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
852 |
|
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
853 |
lemma concat_empty_fset [simp]: |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
854 |
shows "concat_fset {||} = {||}" |
46416
5f5665a0b973
Make automatic derivation of raw/quotient types more greedy to allow descending and quot_lifted for compound quotients.
Cezary Kaliszyk <cezarykaliszyk@gmail.com>
parents:
46404
diff
changeset
|
855 |
by descending simp |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
856 |
|
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
857 |
lemma concat_insert_fset [simp]: |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
858 |
shows "concat_fset (insert_fset x S) = x |\<union>| concat_fset S" |
46416
5f5665a0b973
Make automatic derivation of raw/quotient types more greedy to allow descending and quot_lifted for compound quotients.
Cezary Kaliszyk <cezarykaliszyk@gmail.com>
parents:
46404
diff
changeset
|
859 |
by descending simp |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
860 |
|
46441
992a1688303f
Generalize the compositional preservation theorems
Cezary Kaliszyk <cezarykaliszyk@gmail.com>
parents:
46416
diff
changeset
|
861 |
lemma concat_union_fset [simp]: |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
862 |
shows "concat_fset (xs |\<union>| ys) = concat_fset xs |\<union>| concat_fset ys" |
46416
5f5665a0b973
Make automatic derivation of raw/quotient types more greedy to allow descending and quot_lifted for compound quotients.
Cezary Kaliszyk <cezarykaliszyk@gmail.com>
parents:
46404
diff
changeset
|
863 |
by descending simp |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
864 |
|
46404
7736068b9f56
Quotient FSet: Add compositional respectfulness and preservation for map and lift map_concat
Cezary Kaliszyk <cezarykaliszyk@gmail.com>
parents:
46133
diff
changeset
|
865 |
lemma map_concat_fset: |
7736068b9f56
Quotient FSet: Add compositional respectfulness and preservation for map and lift map_concat
Cezary Kaliszyk <cezarykaliszyk@gmail.com>
parents:
46133
diff
changeset
|
866 |
shows "map_fset f (concat_fset xs) = concat_fset (map_fset (map_fset f) xs)" |
7736068b9f56
Quotient FSet: Add compositional respectfulness and preservation for map and lift map_concat
Cezary Kaliszyk <cezarykaliszyk@gmail.com>
parents:
46133
diff
changeset
|
867 |
by (lifting map_concat) |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
868 |
|
63167 | 869 |
subsection \<open>filter_fset\<close> |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
870 |
|
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
871 |
lemma subset_filter_fset: |
40961 | 872 |
"filter_fset P xs |\<subseteq>| filter_fset Q xs = (\<forall> x. x |\<in>| xs \<longrightarrow> P x \<longrightarrow> Q x)" |
873 |
by descending auto |
|
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
874 |
|
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
875 |
lemma eq_filter_fset: |
40961 | 876 |
"(filter_fset P xs = filter_fset Q xs) = (\<forall>x. x |\<in>| xs \<longrightarrow> P x = Q x)" |
877 |
by descending auto |
|
36280 | 878 |
|
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
879 |
lemma psubset_filter_fset: |
40961 | 880 |
"(\<And>x. x |\<in>| xs \<Longrightarrow> P x \<Longrightarrow> Q x) \<Longrightarrow> (x |\<in>| xs & \<not> P x & Q x) \<Longrightarrow> |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
881 |
filter_fset P xs |\<subset>| filter_fset Q xs" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
882 |
unfolding less_fset_def by (auto simp add: subset_filter_fset eq_filter_fset) |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
883 |
|
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
884 |
|
63167 | 885 |
subsection \<open>fold_fset\<close> |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
886 |
|
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
887 |
lemma fold_empty_fset: |
40961 | 888 |
"fold_fset f {||} = id" |
40962
069cd1c1f39b
more intimate definition of fold_list / fold_once in terms of fold
haftmann
parents:
40961
diff
changeset
|
889 |
by descending (simp add: fold_once_def) |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
890 |
|
40961 | 891 |
lemma fold_insert_fset: "fold_fset f (insert_fset a A) = |
40962
069cd1c1f39b
more intimate definition of fold_list / fold_once in terms of fold
haftmann
parents:
40961
diff
changeset
|
892 |
(if rsp_fold f then if a |\<in>| A then fold_fset f A else fold_fset f A \<circ> f a else id)" |
069cd1c1f39b
more intimate definition of fold_list / fold_once in terms of fold
haftmann
parents:
40961
diff
changeset
|
893 |
by descending (simp add: fold_once_fold_remdups) |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
894 |
|
47092 | 895 |
lemma remdups_removeAll: |
896 |
"remdups (removeAll x xs) = remove1 x (remdups xs)" |
|
897 |
by (induct xs) auto |
|
898 |
||
899 |
lemma member_commute_fold_once: |
|
900 |
assumes "rsp_fold f" |
|
901 |
and "x \<in> set xs" |
|
902 |
shows "fold_once f xs = fold_once f (removeAll x xs) \<circ> f x" |
|
903 |
proof - |
|
904 |
from assms have "fold f (remdups xs) = fold f (remove1 x (remdups xs)) \<circ> f x" |
|
905 |
by (auto intro!: fold_remove1_split elim: rsp_foldE) |
|
63167 | 906 |
then show ?thesis using \<open>rsp_fold f\<close> by (simp add: fold_once_fold_remdups remdups_removeAll) |
47092 | 907 |
qed |
908 |
||
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
909 |
lemma in_commute_fold_fset: |
40962
069cd1c1f39b
more intimate definition of fold_list / fold_once in terms of fold
haftmann
parents:
40961
diff
changeset
|
910 |
"rsp_fold f \<Longrightarrow> h |\<in>| b \<Longrightarrow> fold_fset f b = fold_fset f (remove_fset h b) \<circ> f h" |
069cd1c1f39b
more intimate definition of fold_list / fold_once in terms of fold
haftmann
parents:
40961
diff
changeset
|
911 |
by descending (simp add: member_commute_fold_once) |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
912 |
|
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
913 |
|
63167 | 914 |
subsection \<open>Choice in fsets\<close> |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
915 |
|
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
916 |
lemma fset_choice: |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
917 |
assumes a: "\<forall>x. x |\<in>| A \<longrightarrow> (\<exists>y. P x y)" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
918 |
shows "\<exists>f. \<forall>x. x |\<in>| A \<longrightarrow> P x (f x)" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
919 |
using a |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
920 |
apply(descending) |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
921 |
using finite_set_choice |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
922 |
by (auto simp add: Ball_def) |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
923 |
|
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
924 |
|
63167 | 925 |
section \<open>Induction and Cases rules for fsets\<close> |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
926 |
|
41070 | 927 |
lemma fset_exhaust [case_names empty insert, cases type: fset]: |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
928 |
assumes empty_fset_case: "S = {||} \<Longrightarrow> P" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
929 |
and insert_fset_case: "\<And>x S'. S = insert_fset x S' \<Longrightarrow> P" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
930 |
shows "P" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
931 |
using assms by (lifting list.exhaust) |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
932 |
|
41070 | 933 |
lemma fset_induct [case_names empty insert]: |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
934 |
assumes empty_fset_case: "P {||}" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
935 |
and insert_fset_case: "\<And>x S. P S \<Longrightarrow> P (insert_fset x S)" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
936 |
shows "P S" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
937 |
using assms |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
938 |
by (descending) (blast intro: list.induct) |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
939 |
|
41070 | 940 |
lemma fset_induct_stronger [case_names empty insert, induct type: fset]: |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
941 |
assumes empty_fset_case: "P {||}" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
942 |
and insert_fset_case: "\<And>x S. \<lbrakk>x |\<notin>| S; P S\<rbrakk> \<Longrightarrow> P (insert_fset x S)" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
943 |
shows "P S" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
944 |
proof(induct S rule: fset_induct) |
41070 | 945 |
case empty |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
946 |
show "P {||}" using empty_fset_case by simp |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
947 |
next |
41070 | 948 |
case (insert x S) |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
949 |
have "P S" by fact |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
950 |
then show "P (insert_fset x S)" using insert_fset_case |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
951 |
by (cases "x |\<in>| S") (simp_all) |
36280 | 952 |
qed |
953 |
||
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
954 |
lemma fset_card_induct: |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
955 |
assumes empty_fset_case: "P {||}" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
956 |
and card_fset_Suc_case: "\<And>S T. Suc (card_fset S) = (card_fset T) \<Longrightarrow> P S \<Longrightarrow> P T" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
957 |
shows "P S" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
958 |
proof (induct S) |
41070 | 959 |
case empty |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
960 |
show "P {||}" by (rule empty_fset_case) |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
961 |
next |
41070 | 962 |
case (insert x S) |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
963 |
have h: "P S" by fact |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
964 |
have "x |\<notin>| S" by fact |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
965 |
then have "Suc (card_fset S) = card_fset (insert_fset x S)" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
966 |
using card_fset_Suc by auto |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
967 |
then show "P (insert_fset x S)" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
968 |
using h card_fset_Suc_case by simp |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
969 |
qed |
36280 | 970 |
|
971 |
lemma fset_raw_strong_cases: |
|
36465 | 972 |
obtains "xs = []" |
40953 | 973 |
| ys x where "\<not> List.member ys x" and "xs \<approx> x # ys" |
45129
1fce03e3e8ad
tuned proofs -- eliminated vacuous "induct arbitrary: ..." situations;
wenzelm
parents:
44512
diff
changeset
|
974 |
proof (induct xs) |
36465 | 975 |
case Nil |
976 |
then show thesis by simp |
|
977 |
next |
|
978 |
case (Cons a xs) |
|
40953 | 979 |
have a: "\<lbrakk>xs = [] \<Longrightarrow> thesis; \<And>x ys. \<lbrakk>\<not> List.member ys x; xs \<approx> x # ys\<rbrakk> \<Longrightarrow> thesis\<rbrakk> \<Longrightarrow> thesis" |
980 |
by (rule Cons(1)) |
|
981 |
have b: "\<And>x' ys'. \<lbrakk>\<not> List.member ys' x'; a # xs \<approx> x' # ys'\<rbrakk> \<Longrightarrow> thesis" by fact |
|
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
982 |
have c: "xs = [] \<Longrightarrow> thesis" using b |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
983 |
apply(simp) |
57816
d8bbb97689d3
no need for 'set_simps' now that 'datatype_new' generates the desired 'set' property
blanchet
parents:
55945
diff
changeset
|
984 |
by (metis list.set(1) emptyE empty_subsetI) |
40953 | 985 |
have "\<And>x ys. \<lbrakk>\<not> List.member ys x; xs \<approx> x # ys\<rbrakk> \<Longrightarrow> thesis" |
36465 | 986 |
proof - |
987 |
fix x :: 'a |
|
988 |
fix ys :: "'a list" |
|
40953 | 989 |
assume d:"\<not> List.member ys x" |
36465 | 990 |
assume e:"xs \<approx> x # ys" |
991 |
show thesis |
|
992 |
proof (cases "x = a") |
|
993 |
assume h: "x = a" |
|
40953 | 994 |
then have f: "\<not> List.member ys a" using d by simp |
36465 | 995 |
have g: "a # xs \<approx> a # ys" using e h by auto |
996 |
show thesis using b f g by simp |
|
997 |
next |
|
998 |
assume h: "x \<noteq> a" |
|
40953 | 999 |
then have f: "\<not> List.member (a # ys) x" using d by auto |
36465 | 1000 |
have g: "a # xs \<approx> x # (a # ys)" using e h by auto |
40953 | 1001 |
show thesis using b f g by (simp del: List.member_def) |
36465 | 1002 |
qed |
1003 |
qed |
|
1004 |
then show thesis using a c by blast |
|
1005 |
qed |
|
36280 | 1006 |
|
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
1007 |
|
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
1008 |
lemma fset_strong_cases: |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
1009 |
obtains "xs = {||}" |
40953 | 1010 |
| ys x where "x |\<notin>| ys" and "xs = insert_fset x ys" |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
1011 |
by (lifting fset_raw_strong_cases) |
36280 | 1012 |
|
39996
c02078ff8691
FSet: definition changes propagated from Nominal and more use of 'descending' tactic
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
39995
diff
changeset
|
1013 |
|
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
1014 |
lemma fset_induct2: |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
1015 |
"P {||} {||} \<Longrightarrow> |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
1016 |
(\<And>x xs. x |\<notin>| xs \<Longrightarrow> P (insert_fset x xs) {||}) \<Longrightarrow> |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
1017 |
(\<And>y ys. y |\<notin>| ys \<Longrightarrow> P {||} (insert_fset y ys)) \<Longrightarrow> |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
1018 |
(\<And>x xs y ys. \<lbrakk>P xs ys; x |\<notin>| xs; y |\<notin>| ys\<rbrakk> \<Longrightarrow> P (insert_fset x xs) (insert_fset y ys)) \<Longrightarrow> |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
1019 |
P xsa ysa" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
1020 |
apply (induct xsa arbitrary: ysa) |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
1021 |
apply (induct_tac x rule: fset_induct_stronger) |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
1022 |
apply simp_all |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
1023 |
apply (induct_tac xa rule: fset_induct_stronger) |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
1024 |
apply simp_all |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
1025 |
done |
36280 | 1026 |
|
63167 | 1027 |
text \<open>Extensionality\<close> |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
1028 |
|
41070 | 1029 |
lemma fset_eqI: |
1030 |
assumes "\<And>x. x \<in> fset A \<longleftrightarrow> x \<in> fset B" |
|
1031 |
shows "A = B" |
|
1032 |
using assms proof (induct A arbitrary: B) |
|
1033 |
case empty then show ?case |
|
1034 |
by (auto simp add: in_fset none_in_empty_fset [symmetric] sym) |
|
1035 |
next |
|
1036 |
case (insert x A) |
|
1037 |
from insert.prems insert.hyps(1) have "\<And>z. z \<in> fset A \<longleftrightarrow> z \<in> fset (B - {|x|})" |
|
1038 |
by (auto simp add: in_fset) |
|
53374
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents:
51371
diff
changeset
|
1039 |
then have A: "A = B - {|x|}" by (rule insert.hyps(2)) |
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents:
51371
diff
changeset
|
1040 |
with insert.prems [symmetric, of x] have "x |\<in>| B" by (simp add: in_fset) |
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents:
51371
diff
changeset
|
1041 |
with A show ?case by (metis in_fset_mdef) |
41070 | 1042 |
qed |
36280 | 1043 |
|
63167 | 1044 |
subsection \<open>alternate formulation with a different decomposition principle |
1045 |
and a proof of equivalence\<close> |
|
36280 | 1046 |
|
1047 |
inductive |
|
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80786
diff
changeset
|
1048 |
list_eq2 :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool" (\<open>_ \<approx>2 _\<close>) |
36280 | 1049 |
where |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
1050 |
"(a # b # xs) \<approx>2 (b # a # xs)" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
1051 |
| "[] \<approx>2 []" |
40952 | 1052 |
| "xs \<approx>2 ys \<Longrightarrow> ys \<approx>2 xs" |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
1053 |
| "(a # a # xs) \<approx>2 (a # xs)" |
40952 | 1054 |
| "xs \<approx>2 ys \<Longrightarrow> (a # xs) \<approx>2 (a # ys)" |
1055 |
| "xs1 \<approx>2 xs2 \<Longrightarrow> xs2 \<approx>2 xs3 \<Longrightarrow> xs1 \<approx>2 xs3" |
|
36280 | 1056 |
|
1057 |
lemma list_eq2_refl: |
|
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
1058 |
shows "xs \<approx>2 xs" |
36280 | 1059 |
by (induct xs) (auto intro: list_eq2.intros) |
1060 |
||
1061 |
lemma cons_delete_list_eq2: |
|
40953 | 1062 |
shows "(a # (removeAll a A)) \<approx>2 (if List.member A a then A else a # A)" |
36280 | 1063 |
apply (induct A) |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
1064 |
apply (simp add: list_eq2_refl) |
40953 | 1065 |
apply (case_tac "List.member (aa # A) a") |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
1066 |
apply (simp_all) |
36280 | 1067 |
apply (case_tac [!] "a = aa") |
1068 |
apply (simp_all) |
|
40953 | 1069 |
apply (case_tac "List.member A a") |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
1070 |
apply (auto)[2] |
36280 | 1071 |
apply (metis list_eq2.intros(3) list_eq2.intros(4) list_eq2.intros(5) list_eq2.intros(6)) |
1072 |
apply (metis list_eq2.intros(1) list_eq2.intros(5) list_eq2.intros(6)) |
|
40953 | 1073 |
apply (auto simp add: list_eq2_refl) |
36280 | 1074 |
done |
1075 |
||
40953 | 1076 |
lemma member_delete_list_eq2: |
1077 |
assumes a: "List.member r e" |
|
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
1078 |
shows "(e # removeAll e r) \<approx>2 r" |
36280 | 1079 |
using a cons_delete_list_eq2[of e r] |
1080 |
by simp |
|
1081 |
||
1082 |
lemma list_eq2_equiv: |
|
1083 |
"(l \<approx> r) \<longleftrightarrow> (list_eq2 l r)" |
|
1084 |
proof |
|
1085 |
show "list_eq2 l r \<Longrightarrow> l \<approx> r" by (induct rule: list_eq2.induct) auto |
|
1086 |
next |
|
1087 |
{ |
|
1088 |
fix n |
|
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
1089 |
assume a: "card_list l = n" and b: "l \<approx> r" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
1090 |
have "l \<approx>2 r" |
36280 | 1091 |
using a b |
1092 |
proof (induct n arbitrary: l r) |
|
1093 |
case 0 |
|
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
1094 |
have "card_list l = 0" by fact |
40953 | 1095 |
then have "\<forall>x. \<not> List.member l x" by auto |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
1096 |
then have z: "l = []" by auto |
63167 | 1097 |
then have "r = []" using \<open>l \<approx> r\<close> by simp |
36280 | 1098 |
then show ?case using z list_eq2_refl by simp |
1099 |
next |
|
1100 |
case (Suc m) |
|
1101 |
have b: "l \<approx> r" by fact |
|
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
1102 |
have d: "card_list l = Suc m" by fact |
40953 | 1103 |
then have "\<exists>a. List.member l a" |
41067 | 1104 |
apply(simp) |
1105 |
apply(drule card_eq_SucD) |
|
1106 |
apply(blast) |
|
1107 |
done |
|
40953 | 1108 |
then obtain a where e: "List.member l a" by auto |
1109 |
then have e': "List.member r a" using list_eq_def [simplified List.member_def [symmetric], of l r] b |
|
41067 | 1110 |
by auto |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
1111 |
have f: "card_list (removeAll a l) = m" using e d by (simp) |
47198
cfd8ff62eab1
use qualified names for rsp and rep_eq theorems in quotient_def
kuncar
parents:
47092
diff
changeset
|
1112 |
have g: "removeAll a l \<approx> removeAll a r" using remove_fset.rsp b by simp |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
1113 |
have "(removeAll a l) \<approx>2 (removeAll a r)" by (rule Suc.hyps[OF f g]) |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
1114 |
then have h: "(a # removeAll a l) \<approx>2 (a # removeAll a r)" by (rule list_eq2.intros(5)) |
41067 | 1115 |
have i: "l \<approx>2 (a # removeAll a l)" |
40953 | 1116 |
by (rule list_eq2.intros(3)[OF member_delete_list_eq2[OF e]]) |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
1117 |
have "l \<approx>2 (a # removeAll a r)" by (rule list_eq2.intros(6)[OF i h]) |
40953 | 1118 |
then show ?case using list_eq2.intros(6)[OF _ member_delete_list_eq2[OF e']] by simp |
36280 | 1119 |
qed |
1120 |
} |
|
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
1121 |
then show "l \<approx> r \<Longrightarrow> l \<approx>2 r" by blast |
36280 | 1122 |
qed |
1123 |
||
1124 |
||
1125 |
(* We cannot write it as "assumes .. shows" since Isabelle changes |
|
1126 |
the quantifiers to schematic variables and reintroduces them in |
|
1127 |
a different order *) |
|
1128 |
lemma fset_eq_cases: |
|
1129 |
"\<lbrakk>a1 = a2; |
|
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
1130 |
\<And>a b xs. \<lbrakk>a1 = insert_fset a (insert_fset b xs); a2 = insert_fset b (insert_fset a xs)\<rbrakk> \<Longrightarrow> P; |
36280 | 1131 |
\<lbrakk>a1 = {||}; a2 = {||}\<rbrakk> \<Longrightarrow> P; \<And>xs ys. \<lbrakk>a1 = ys; a2 = xs; xs = ys\<rbrakk> \<Longrightarrow> P; |
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
1132 |
\<And>a xs. \<lbrakk>a1 = insert_fset a (insert_fset a xs); a2 = insert_fset a xs\<rbrakk> \<Longrightarrow> P; |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
1133 |
\<And>xs ys a. \<lbrakk>a1 = insert_fset a xs; a2 = insert_fset a ys; xs = ys\<rbrakk> \<Longrightarrow> P; |
36280 | 1134 |
\<And>xs1 xs2 xs3. \<lbrakk>a1 = xs1; a2 = xs3; xs1 = xs2; xs2 = xs3\<rbrakk> \<Longrightarrow> P\<rbrakk> |
1135 |
\<Longrightarrow> P" |
|
1136 |
by (lifting list_eq2.cases[simplified list_eq2_equiv[symmetric]]) |
|
1137 |
||
1138 |
lemma fset_eq_induct: |
|
1139 |
assumes "x1 = x2" |
|
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
1140 |
and "\<And>a b xs. P (insert_fset a (insert_fset b xs)) (insert_fset b (insert_fset a xs))" |
36280 | 1141 |
and "P {||} {||}" |
1142 |
and "\<And>xs ys. \<lbrakk>xs = ys; P xs ys\<rbrakk> \<Longrightarrow> P ys xs" |
|
40030
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
1143 |
and "\<And>a xs. P (insert_fset a (insert_fset a xs)) (insert_fset a xs)" |
9f8dcf6ef563
reorganisation of the FSet theory (changed the primary naming scheme of constants and theorems to *_fset)
Christian Urban <urbanc@in.tum.de>
parents:
39996
diff
changeset
|
1144 |
and "\<And>xs ys a. \<lbrakk>xs = ys; P xs ys\<rbrakk> \<Longrightarrow> P (insert_fset a xs) (insert_fset a ys)" |
36280 | 1145 |
and "\<And>xs1 xs2 xs3. \<lbrakk>xs1 = xs2; P xs1 xs2; xs2 = xs3; P xs2 xs3\<rbrakk> \<Longrightarrow> P xs1 xs3" |
1146 |
shows "P x1 x2" |
|
1147 |
using assms |
|
1148 |
by (lifting list_eq2.induct[simplified list_eq2_equiv[symmetric]]) |
|
1149 |
||
63167 | 1150 |
ML \<open> |
80695 | 1151 |
fun dest_fsetT \<^Type>\<open>fset T\<close> = T |
36280 | 1152 |
| dest_fsetT T = raise TYPE ("dest_fsetT: fset type expected", [T], []); |
63167 | 1153 |
\<close> |
36280 | 1154 |
|
1155 |
no_notation |
|
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80786
diff
changeset
|
1156 |
list_eq (infix \<open>\<approx>\<close> 50) and |
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80786
diff
changeset
|
1157 |
list_eq2 (infix \<open>\<approx>2\<close> 50) |
36280 | 1158 |
|
1159 |
end |