author | wenzelm |
Fri, 18 Aug 2017 20:47:47 +0200 | |
changeset 66453 | cc19f7ca2ed6 |
parent 63167 | 0909deb8059b |
permissions | -rw-r--r-- |
43766
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
1 |
(* Title: HOL/Quotient_Examples/DList.thy |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
2 |
Author: Cezary Kaliszyk, University of Tsukuba |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
3 |
|
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
4 |
Based on typedef version "Library/Dlist" by Florian Haftmann |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
5 |
and theory morphism version by Maksym Bortin |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
6 |
*) |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
7 |
|
63167 | 8 |
section \<open>Lists with distinct elements\<close> |
43766
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
9 |
|
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
10 |
theory DList |
66453
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
wenzelm
parents:
63167
diff
changeset
|
11 |
imports "HOL-Library.Quotient_List" |
43766
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
12 |
begin |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
13 |
|
63167 | 14 |
text \<open>Some facts about lists\<close> |
43766
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
15 |
|
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
16 |
lemma remdups_removeAll_commute[simp]: |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
17 |
"remdups (removeAll x l) = removeAll x (remdups l)" |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
18 |
by (induct l) auto |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
19 |
|
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
20 |
lemma removeAll_distinct[simp]: |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
21 |
assumes "distinct l" |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
22 |
shows "distinct (removeAll x l)" |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
23 |
using assms by (induct l) simp_all |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
24 |
|
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
25 |
lemma removeAll_commute: |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
26 |
"removeAll x (removeAll y l) = removeAll y (removeAll x l)" |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
27 |
by (induct l) auto |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
28 |
|
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
29 |
lemma remdups_hd_notin_tl: |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
30 |
"remdups dl = h # t \<Longrightarrow> h \<notin> set t" |
60668 | 31 |
proof (induct dl arbitrary: h t) |
32 |
case Nil |
|
33 |
then show ?case by simp |
|
34 |
next |
|
35 |
case (Cons a dl) |
|
36 |
then show ?case by (cases "a \<in> set dl") auto |
|
37 |
qed |
|
43766
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
38 |
|
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
39 |
lemma remdups_repeat: |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
40 |
"remdups dl = h # t \<Longrightarrow> t = remdups t" |
60668 | 41 |
proof (induct dl arbitrary: h t) |
42 |
case Nil |
|
43 |
then show ?case by simp |
|
44 |
next |
|
45 |
case (Cons a dl) |
|
46 |
then show ?case by (cases "a \<in> set dl") (auto simp: remdups_remdups) |
|
47 |
qed |
|
43766
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
48 |
|
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
49 |
lemma remdups_nil_noteq_cons: |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
50 |
"remdups (h # t) \<noteq> remdups []" |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
51 |
"remdups [] \<noteq> remdups (h # t)" |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
52 |
by auto |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
53 |
|
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
54 |
lemma remdups_eq_map_eq: |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
55 |
assumes "remdups xa = remdups ya" |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
56 |
shows "remdups (map f xa) = remdups (map f ya)" |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
57 |
using assms |
45129
1fce03e3e8ad
tuned proofs -- eliminated vacuous "induct arbitrary: ..." situations;
wenzelm
parents:
44263
diff
changeset
|
58 |
by (induct xa ya rule: list_induct2') |
43766
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
59 |
(metis (full_types) remdups_nil_noteq_cons(2) remdups_map_remdups)+ |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
60 |
|
44263
971d1be5d5ce
Quotient_Examples/DList: explicit proof of remdups_eq_member_eq needed for explicit set type.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
43766
diff
changeset
|
61 |
lemma remdups_eq_member_eq: |
971d1be5d5ce
Quotient_Examples/DList: explicit proof of remdups_eq_member_eq needed for explicit set type.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
43766
diff
changeset
|
62 |
assumes "remdups xa = remdups ya" |
971d1be5d5ce
Quotient_Examples/DList: explicit proof of remdups_eq_member_eq needed for explicit set type.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
43766
diff
changeset
|
63 |
shows "List.member xa = List.member ya" |
971d1be5d5ce
Quotient_Examples/DList: explicit proof of remdups_eq_member_eq needed for explicit set type.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
43766
diff
changeset
|
64 |
using assms |
971d1be5d5ce
Quotient_Examples/DList: explicit proof of remdups_eq_member_eq needed for explicit set type.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
43766
diff
changeset
|
65 |
unfolding fun_eq_iff List.member_def |
971d1be5d5ce
Quotient_Examples/DList: explicit proof of remdups_eq_member_eq needed for explicit set type.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
43766
diff
changeset
|
66 |
by (induct xa ya rule: list_induct2') |
971d1be5d5ce
Quotient_Examples/DList: explicit proof of remdups_eq_member_eq needed for explicit set type.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
43766
diff
changeset
|
67 |
(metis remdups_nil_noteq_cons set_remdups)+ |
971d1be5d5ce
Quotient_Examples/DList: explicit proof of remdups_eq_member_eq needed for explicit set type.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
43766
diff
changeset
|
68 |
|
63167 | 69 |
text \<open>Setting up the quotient type\<close> |
43766
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
70 |
|
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
71 |
definition |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
72 |
dlist_eq :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool" |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
73 |
where |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
74 |
[simp]: "dlist_eq xs ys \<longleftrightarrow> remdups xs = remdups ys" |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
75 |
|
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
76 |
lemma dlist_eq_reflp: |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
77 |
"reflp dlist_eq" |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
78 |
by (auto intro: reflpI) |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
79 |
|
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
80 |
lemma dlist_eq_symp: |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
81 |
"symp dlist_eq" |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
82 |
by (auto intro: sympI) |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
83 |
|
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
84 |
lemma dlist_eq_transp: |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
85 |
"transp dlist_eq" |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
86 |
by (auto intro: transpI) |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
87 |
|
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
88 |
lemma dlist_eq_equivp: |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
89 |
"equivp dlist_eq" |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
90 |
by (auto intro: equivpI dlist_eq_reflp dlist_eq_symp dlist_eq_transp) |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
91 |
|
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
92 |
quotient_type |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
93 |
'a dlist = "'a list" / "dlist_eq" |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
94 |
by (rule dlist_eq_equivp) |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
95 |
|
63167 | 96 |
text \<open>respectfulness and constant definitions\<close> |
43766
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
97 |
|
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
98 |
definition [simp]: "card_remdups = length \<circ> remdups" |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
99 |
definition [simp]: "foldr_remdups f xs e = foldr f (remdups xs) e" |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
100 |
|
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
101 |
quotient_definition empty where "empty :: 'a dlist" |
60668 | 102 |
is "Nil" . |
43766
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
103 |
|
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
104 |
quotient_definition insert where "insert :: 'a \<Rightarrow> 'a dlist \<Rightarrow> 'a dlist" |
47092 | 105 |
is "Cons" by (metis (mono_tags) List.insert_def dlist_eq_def remdups.simps(2) set_remdups) |
43766
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
106 |
|
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
107 |
quotient_definition "member :: 'a dlist \<Rightarrow> 'a \<Rightarrow> bool" |
47092 | 108 |
is "List.member" by (metis dlist_eq_def remdups_eq_member_eq) |
43766
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
109 |
|
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
110 |
quotient_definition foldr where "foldr :: ('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'a dlist \<Rightarrow> 'b \<Rightarrow> 'b" |
47092 | 111 |
is "foldr_remdups" by auto |
43766
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
112 |
|
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
113 |
quotient_definition "remove :: 'a \<Rightarrow> 'a dlist \<Rightarrow> 'a dlist" |
47092 | 114 |
is "removeAll" by force |
43766
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
115 |
|
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
116 |
quotient_definition card where "card :: 'a dlist \<Rightarrow> nat" |
47092 | 117 |
is "card_remdups" by fastforce |
43766
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
118 |
|
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
119 |
quotient_definition map where "map :: ('a \<Rightarrow> 'b) \<Rightarrow> 'a dlist \<Rightarrow> 'b dlist" |
47092 | 120 |
is "List.map :: ('a \<Rightarrow> 'b) \<Rightarrow> 'a list \<Rightarrow> 'b list" by (metis dlist_eq_def remdups_eq_map_eq) |
43766
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
121 |
|
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
122 |
quotient_definition filter where "filter :: ('a \<Rightarrow> bool) \<Rightarrow> 'a dlist \<Rightarrow> 'a dlist" |
47092 | 123 |
is "List.filter" by (metis dlist_eq_def remdups_filter) |
43766
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
124 |
|
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
125 |
quotient_definition "list_of_dlist :: 'a dlist \<Rightarrow> 'a list" |
47092 | 126 |
is "remdups" by simp |
43766
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
127 |
|
63167 | 128 |
text \<open>lifted theorems\<close> |
43766
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
129 |
|
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
130 |
lemma dlist_member_insert: |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
131 |
"member dl x \<Longrightarrow> insert x dl = dl" |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
132 |
by descending (simp add: List.member_def) |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
133 |
|
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
134 |
lemma dlist_member_insert_eq: |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
135 |
"member (insert y dl) x = (x = y \<or> member dl x)" |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
136 |
by descending (simp add: List.member_def) |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
137 |
|
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
138 |
lemma dlist_insert_idem: |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
139 |
"insert x (insert x dl) = insert x dl" |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
140 |
by descending simp |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
141 |
|
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
142 |
lemma dlist_insert_not_empty: |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
143 |
"insert x dl \<noteq> empty" |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
144 |
by descending auto |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
145 |
|
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
146 |
lemma not_dlist_member_empty: |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
147 |
"\<not> member empty x" |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
148 |
by descending (simp add: List.member_def) |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
149 |
|
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
150 |
lemma not_dlist_member_remove: |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
151 |
"\<not> member (remove x dl) x" |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
152 |
by descending (simp add: List.member_def) |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
153 |
|
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
154 |
lemma dlist_in_remove: |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
155 |
"a \<noteq> b \<Longrightarrow> member (remove b dl) a = member dl a" |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
156 |
by descending (simp add: List.member_def) |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
157 |
|
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
158 |
lemma dlist_not_memb_remove: |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
159 |
"\<not> member dl x \<Longrightarrow> remove x dl = dl" |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
160 |
by descending (simp add: List.member_def) |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
161 |
|
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
162 |
lemma dlist_no_memb_remove_insert: |
60668 | 163 |
"\<not> member dl x \<Longrightarrow> remove x (insert x dl) = dl" |
43766
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
164 |
by descending (simp add: List.member_def) |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
165 |
|
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
166 |
lemma dlist_remove_empty: |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
167 |
"remove x empty = empty" |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
168 |
by descending simp |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
169 |
|
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
170 |
lemma dlist_remove_insert_commute: |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
171 |
"a \<noteq> b \<Longrightarrow> remove a (insert b dl) = insert b (remove a dl)" |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
172 |
by descending simp |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
173 |
|
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
174 |
lemma dlist_remove_commute: |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
175 |
"remove a (remove b dl) = remove b (remove a dl)" |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
176 |
by (lifting removeAll_commute) |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
177 |
|
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
178 |
lemma dlist_foldr_empty: |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
179 |
"foldr f empty e = e" |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
180 |
by descending simp |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
181 |
|
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
182 |
lemma dlist_no_memb_foldr: |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
183 |
assumes "\<not> member dl x" |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
184 |
shows "foldr f (insert x dl) e = f x (foldr f dl e)" |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
185 |
using assms by descending (simp add: List.member_def) |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
186 |
|
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
187 |
lemma dlist_foldr_insert_not_memb: |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
188 |
assumes "\<not>member t h" |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
189 |
shows "foldr f (insert h t) e = f h (foldr f t e)" |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
190 |
using assms by descending (simp add: List.member_def) |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
191 |
|
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
192 |
lemma list_of_dlist_empty[simp]: |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
193 |
"list_of_dlist empty = []" |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
194 |
by descending simp |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
195 |
|
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
196 |
lemma list_of_dlist_insert[simp]: |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
197 |
"list_of_dlist (insert x xs) = |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
198 |
(if member xs x then (remdups (list_of_dlist xs)) |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
199 |
else x # (remdups (list_of_dlist xs)))" |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
200 |
by descending (simp add: List.member_def remdups_remdups) |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
201 |
|
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
202 |
lemma list_of_dlist_remove[simp]: |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
203 |
"list_of_dlist (remove x dxs) = remove1 x (list_of_dlist dxs)" |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
204 |
by descending (simp add: distinct_remove1_removeAll) |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
205 |
|
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
206 |
lemma list_of_dlist_map[simp]: |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
207 |
"list_of_dlist (map f dxs) = remdups (List.map f (list_of_dlist dxs))" |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
208 |
by descending (simp add: remdups_map_remdups) |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
209 |
|
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
210 |
lemma list_of_dlist_filter [simp]: |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
211 |
"list_of_dlist (filter P dxs) = List.filter P (list_of_dlist dxs)" |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
212 |
by descending (simp add: remdups_filter) |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
213 |
|
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
214 |
lemma dlist_map_empty: |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
215 |
"map f empty = empty" |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
216 |
by descending simp |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
217 |
|
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
218 |
lemma dlist_map_insert: |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
219 |
"map f (insert x xs) = insert (f x) (map f xs)" |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
220 |
by descending simp |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
221 |
|
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
222 |
lemma dlist_eq_iff: |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
223 |
"dxs = dys \<longleftrightarrow> list_of_dlist dxs = list_of_dlist dys" |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
224 |
by descending simp |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
225 |
|
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
226 |
lemma dlist_eqI: |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
227 |
"list_of_dlist dxs = list_of_dlist dys \<Longrightarrow> dxs = dys" |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
228 |
by (simp add: dlist_eq_iff) |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
229 |
|
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
230 |
abbreviation |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
231 |
"dlist xs \<equiv> abs_dlist xs" |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
232 |
|
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
233 |
lemma distinct_list_of_dlist [simp, intro]: |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
234 |
"distinct (list_of_dlist dxs)" |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
235 |
by descending simp |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
236 |
|
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
237 |
lemma list_of_dlist_dlist [simp]: |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
238 |
"list_of_dlist (dlist xs) = remdups xs" |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
239 |
unfolding list_of_dlist_def map_fun_apply id_def |
47308 | 240 |
by (metis Quotient3_rep_abs[OF Quotient3_dlist] dlist_eq_def) |
43766
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
241 |
|
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
242 |
lemma remdups_list_of_dlist [simp]: |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
243 |
"remdups (list_of_dlist dxs) = list_of_dlist dxs" |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
244 |
by simp |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
245 |
|
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
246 |
lemma dlist_list_of_dlist [simp, code abstype]: |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
247 |
"dlist (list_of_dlist dxs) = dxs" |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
248 |
by (simp add: list_of_dlist_def) |
47308 | 249 |
(metis Quotient3_def Quotient3_dlist dlist_eqI list_of_dlist_dlist remdups_list_of_dlist) |
43766
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
250 |
|
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
251 |
lemma dlist_filter_simps: |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
252 |
"filter P empty = empty" |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
253 |
"filter P (insert x xs) = (if P x then insert x (filter P xs) else filter P xs)" |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
254 |
by (lifting filter.simps) |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
255 |
|
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
256 |
lemma dlist_induct: |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
257 |
assumes "P empty" |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
258 |
and "\<And>a dl. P dl \<Longrightarrow> P (insert a dl)" |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
259 |
shows "P dl" |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
260 |
using assms by descending (drule list.induct, simp) |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
261 |
|
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
262 |
lemma dlist_induct_stronger: |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
263 |
assumes a1: "P empty" |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
264 |
and a2: "\<And>x dl. \<lbrakk>\<not>member dl x; P dl\<rbrakk> \<Longrightarrow> P (insert x dl)" |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
265 |
shows "P dl" |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
266 |
proof(induct dl rule: dlist_induct) |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
267 |
show "P empty" using a1 by simp |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
268 |
next |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
269 |
fix x dl |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
270 |
assume "P dl" |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
271 |
then show "P (insert x dl)" using a2 |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
272 |
by (cases "member dl x") (simp_all add: dlist_member_insert) |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
273 |
qed |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
274 |
|
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
275 |
lemma dlist_card_induct: |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
276 |
assumes "\<And>xs. (\<And>ys. card ys < card xs \<Longrightarrow> P ys) \<Longrightarrow> P xs" |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
277 |
shows "P xs" |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
278 |
using assms |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
279 |
by descending (rule measure_induct [of card_remdups], blast) |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
280 |
|
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
281 |
lemma dlist_cases: |
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
282 |
"dl = empty \<or> (\<exists>h t. dl = insert h t \<and> \<not> member t h)" |
60668 | 283 |
by descending |
284 |
(metis List.member_def dlist_eq_def hd_Cons_tl list_of_dlist_dlist remdups_hd_notin_tl remdups_list_of_dlist) |
|
43766
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
285 |
|
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
286 |
lemma dlist_exhaust: |
60668 | 287 |
obtains "y = empty" | a dl where "y = insert a dl" |
288 |
by (lifting list.exhaust) |
|
43766
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
289 |
|
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
290 |
lemma dlist_exhaust_stronger: |
60668 | 291 |
obtains "y = empty" | a dl where "y = insert a dl" "\<not> member dl a" |
292 |
by (metis dlist_cases) |
|
43766
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
293 |
|
9545bb3cefac
Quotient example: Lists with distinct elements
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
294 |
end |