author | immler |
Mon, 26 Feb 2018 09:58:47 +0100 | |
changeset 67728 | d97a28a006f9 |
parent 67406 | 23307fd33906 |
child 67929 | 30486b96274d |
permissions | -rw-r--r-- |
61640 | 1 |
(* Author: Tobias Nipkow *) |
2 |
||
67406 | 3 |
section \<open>List Insertion and Deletion\<close> |
61640 | 4 |
|
5 |
theory List_Ins_Del |
|
6 |
imports Sorted_Less |
|
7 |
begin |
|
8 |
||
9 |
subsection \<open>Elements in a list\<close> |
|
10 |
||
11 |
fun elems :: "'a list \<Rightarrow> 'a set" where |
|
12 |
"elems [] = {}" | |
|
13 |
"elems (x#xs) = Set.insert x (elems xs)" |
|
14 |
||
15 |
lemma elems_app: "elems (xs @ ys) = (elems xs \<union> elems ys)" |
|
16 |
by (induction xs) auto |
|
17 |
||
18 |
lemma elems_eq_set: "elems xs = set xs" |
|
19 |
by (induction xs) auto |
|
20 |
||
21 |
lemma sorted_Cons_iff: |
|
66441 | 22 |
"sorted(x # xs) = ((\<forall>y \<in> elems xs. x < y) \<and> sorted xs)" |
23 |
by(simp add: elems_eq_set sorted_wrt_Cons) |
|
61640 | 24 |
|
25 |
lemma sorted_snoc_iff: |
|
26 |
"sorted(xs @ [x]) = (sorted xs \<and> (\<forall>y \<in> elems xs. y < x))" |
|
66441 | 27 |
by(simp add: elems_eq_set sorted_wrt_append) |
61640 | 28 |
|
67406 | 29 |
text\<open>The above two rules introduce quantifiers. It turns out |
61640 | 30 |
that in practice this is not a problem because of the simplicity of |
31 |
the "isin" functions that implement @{const elems}. Nevertheless |
|
67406 | 32 |
it is possible to avoid the quantifiers with the help of some rewrite rules:\<close> |
61640 | 33 |
|
61692
cb595e12451d
removed lemmas that were only needed for old version of isin.
nipkow
parents:
61642
diff
changeset
|
34 |
lemma sorted_ConsD: "sorted (y # xs) \<Longrightarrow> x \<le> y \<Longrightarrow> x \<notin> elems xs" |
cb595e12451d
removed lemmas that were only needed for old version of isin.
nipkow
parents:
61642
diff
changeset
|
35 |
by (auto simp: sorted_Cons_iff) |
61640 | 36 |
|
61692
cb595e12451d
removed lemmas that were only needed for old version of isin.
nipkow
parents:
61642
diff
changeset
|
37 |
lemma sorted_snocD: "sorted (xs @ [y]) \<Longrightarrow> y \<le> x \<Longrightarrow> x \<notin> elems xs" |
cb595e12451d
removed lemmas that were only needed for old version of isin.
nipkow
parents:
61642
diff
changeset
|
38 |
by (auto simp: sorted_snoc_iff) |
61640 | 39 |
|
40 |
lemmas elems_simps = sorted_lems elems_app |
|
41 |
lemmas elems_simps1 = elems_simps sorted_Cons_iff sorted_snoc_iff |
|
61692
cb595e12451d
removed lemmas that were only needed for old version of isin.
nipkow
parents:
61642
diff
changeset
|
42 |
lemmas elems_simps2 = elems_simps sorted_ConsD sorted_snocD |
61640 | 43 |
|
44 |
||
45 |
subsection \<open>Inserting into an ordered list without duplicates:\<close> |
|
46 |
||
47 |
fun ins_list :: "'a::linorder \<Rightarrow> 'a list \<Rightarrow> 'a list" where |
|
48 |
"ins_list x [] = [x]" | |
|
49 |
"ins_list x (a#xs) = |
|
50 |
(if x < a then x#a#xs else if x=a then a#xs else a # ins_list x xs)" |
|
51 |
||
52 |
lemma set_ins_list: "elems (ins_list x xs) = insert x (elems xs)" |
|
53 |
by(induction xs) auto |
|
54 |
||
55 |
lemma distinct_if_sorted: "sorted xs \<Longrightarrow> distinct xs" |
|
67168 | 56 |
apply(induction xs rule: induct_list012) |
61640 | 57 |
apply auto |
58 |
by (metis in_set_conv_decomp_first less_imp_not_less sorted_mid_iff2) |
|
59 |
||
60 |
lemma sorted_ins_list: "sorted xs \<Longrightarrow> sorted(ins_list x xs)" |
|
67168 | 61 |
by(induction xs rule: induct_list012) auto |
61640 | 62 |
|
63 |
lemma ins_list_sorted: "sorted (xs @ [a]) \<Longrightarrow> |
|
64 |
ins_list x (xs @ a # ys) = |
|
61642 | 65 |
(if x < a then ins_list x xs @ (a#ys) else xs @ ins_list x (a#ys))" |
61640 | 66 |
by(induction xs) (auto simp: sorted_lems) |
67 |
||
68 |
text\<open>In principle, @{thm ins_list_sorted} suffices, but the following two |
|
69 |
corollaries speed up proofs.\<close> |
|
70 |
||
71 |
corollary ins_list_sorted1: "sorted (xs @ [a]) \<Longrightarrow> a \<le> x \<Longrightarrow> |
|
72 |
ins_list x (xs @ a # ys) = xs @ ins_list x (a#ys)" |
|
61642 | 73 |
by(auto simp add: ins_list_sorted) |
61640 | 74 |
|
75 |
corollary ins_list_sorted2: "sorted (xs @ [a]) \<Longrightarrow> x < a \<Longrightarrow> |
|
76 |
ins_list x (xs @ a # ys) = ins_list x xs @ (a#ys)" |
|
77 |
by(auto simp: ins_list_sorted) |
|
78 |
||
79 |
lemmas ins_list_simps = sorted_lems ins_list_sorted1 ins_list_sorted2 |
|
80 |
||
61696 | 81 |
text\<open>Splay trees need two additional @{const ins_list} lemmas:\<close> |
82 |
||
83 |
lemma ins_list_Cons: "sorted (x # xs) \<Longrightarrow> ins_list x xs = x # xs" |
|
84 |
by (induction xs) auto |
|
85 |
||
86 |
lemma ins_list_snoc: "sorted (xs @ [x]) \<Longrightarrow> ins_list x xs = xs @ [x]" |
|
87 |
by(induction xs) (auto simp add: sorted_mid_iff2) |
|
88 |
||
61640 | 89 |
|
90 |
subsection \<open>Delete one occurrence of an element from a list:\<close> |
|
91 |
||
92 |
fun del_list :: "'a \<Rightarrow> 'a list \<Rightarrow> 'a list" where |
|
93 |
"del_list x [] = []" | |
|
94 |
"del_list x (a#xs) = (if x=a then xs else a # del_list x xs)" |
|
95 |
||
96 |
lemma del_list_idem: "x \<notin> elems xs \<Longrightarrow> del_list x xs = xs" |
|
97 |
by (induct xs) simp_all |
|
98 |
||
99 |
lemma elems_del_list_eq: |
|
100 |
"distinct xs \<Longrightarrow> elems (del_list x xs) = elems xs - {x}" |
|
101 |
apply(induct xs) |
|
102 |
apply simp |
|
103 |
apply (simp add: elems_eq_set) |
|
104 |
apply blast |
|
105 |
done |
|
106 |
||
107 |
lemma sorted_del_list: "sorted xs \<Longrightarrow> sorted(del_list x xs)" |
|
67168 | 108 |
apply(induction xs rule: induct_list012) |
61640 | 109 |
apply auto |
110 |
by (meson order.strict_trans sorted_Cons_iff) |
|
111 |
||
112 |
lemma del_list_sorted: "sorted (xs @ a # ys) \<Longrightarrow> |
|
113 |
del_list x (xs @ a # ys) = (if x < a then del_list x xs @ a # ys else xs @ del_list x (a # ys))" |
|
114 |
by(induction xs) |
|
115 |
(fastforce simp: sorted_lems sorted_Cons_iff elems_eq_set intro!: del_list_idem)+ |
|
116 |
||
117 |
text\<open>In principle, @{thm del_list_sorted} suffices, but the following |
|
118 |
corollaries speed up proofs.\<close> |
|
119 |
||
120 |
corollary del_list_sorted1: "sorted (xs @ a # ys) \<Longrightarrow> a \<le> x \<Longrightarrow> |
|
121 |
del_list x (xs @ a # ys) = xs @ del_list x (a # ys)" |
|
122 |
by (auto simp: del_list_sorted) |
|
123 |
||
124 |
corollary del_list_sorted2: "sorted (xs @ a # ys) \<Longrightarrow> x < a \<Longrightarrow> |
|
125 |
del_list x (xs @ a # ys) = del_list x xs @ a # ys" |
|
126 |
by (auto simp: del_list_sorted) |
|
127 |
||
128 |
corollary del_list_sorted3: |
|
129 |
"sorted (xs @ a # ys @ b # zs) \<Longrightarrow> x < b \<Longrightarrow> |
|
130 |
del_list x (xs @ a # ys @ b # zs) = del_list x (xs @ a # ys) @ b # zs" |
|
131 |
by (auto simp: del_list_sorted sorted_lems) |
|
132 |
||
133 |
corollary del_list_sorted4: |
|
134 |
"sorted (xs @ a # ys @ b # zs @ c # us) \<Longrightarrow> x < c \<Longrightarrow> |
|
135 |
del_list x (xs @ a # ys @ b # zs @ c # us) = del_list x (xs @ a # ys @ b # zs) @ c # us" |
|
136 |
by (auto simp: del_list_sorted sorted_lems) |
|
137 |
||
138 |
corollary del_list_sorted5: |
|
139 |
"sorted (xs @ a # ys @ b # zs @ c # us @ d # vs) \<Longrightarrow> x < d \<Longrightarrow> |
|
140 |
del_list x (xs @ a # ys @ b # zs @ c # us @ d # vs) = |
|
141 |
del_list x (xs @ a # ys @ b # zs @ c # us) @ d # vs" |
|
142 |
by (auto simp: del_list_sorted sorted_lems) |
|
143 |
||
144 |
lemmas del_list_simps = sorted_lems |
|
145 |
del_list_sorted1 |
|
146 |
del_list_sorted2 |
|
147 |
del_list_sorted3 |
|
148 |
del_list_sorted4 |
|
149 |
del_list_sorted5 |
|
150 |
||
61696 | 151 |
text\<open>Splay trees need two additional @{const del_list} lemmas:\<close> |
152 |
||
153 |
lemma del_list_notin_Cons: "sorted (x # xs) \<Longrightarrow> del_list x xs = xs" |
|
66441 | 154 |
by(induction xs)(fastforce simp: sorted_Cons_iff)+ |
61696 | 155 |
|
156 |
lemma del_list_sorted_app: |
|
157 |
"sorted(xs @ [x]) \<Longrightarrow> del_list x (xs @ ys) = xs @ del_list x ys" |
|
158 |
by (induction xs) (auto simp: sorted_mid_iff2) |
|
159 |
||
61640 | 160 |
end |