author | wenzelm |
Mon, 11 Sep 2023 19:30:48 +0200 | |
changeset 78659 | b5f3d1051b13 |
parent 76217 | 8655344f1cf6 |
permissions | -rw-r--r-- |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
26060
diff
changeset
|
1 |
(* Title: ZF/UNITY/GenPrefix.thy |
12197 | 2 |
Author: Sidi O Ehmety, Cambridge University Computer Laboratory |
3 |
Copyright 2001 University of Cambridge |
|
4 |
||
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
5 |
\<langle>xs,ys\<rangle>:gen_prefix(r) |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
26060
diff
changeset
|
6 |
if ys = xs' @ zs where length(xs) = length(xs') |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
26060
diff
changeset
|
7 |
and corresponding elements of xs, xs' are pairwise related by r |
12197 | 8 |
|
9 |
Based on Lex/Prefix |
|
10 |
*) |
|
11 |
||
60770 | 12 |
section\<open>Charpentier's Generalized Prefix Relation\<close> |
15634 | 13 |
|
14 |
theory GenPrefix |
|
65449
c82e63b11b8b
clarified main ZF.thy / ZFC.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents:
63648
diff
changeset
|
15 |
imports ZF |
15634 | 16 |
begin |
12197 | 17 |
|
24893 | 18 |
definition (*really belongs in ZF/Trancl*) |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
19 |
part_order :: "[i, i] \<Rightarrow> o" where |
76214 | 20 |
"part_order(A, r) \<equiv> refl(A,r) \<and> trans[A](r) \<and> antisym(r)" |
14055 | 21 |
|
12197 | 22 |
consts |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
23 |
gen_prefix :: "[i, i] \<Rightarrow> i" |
24892 | 24 |
|
12197 | 25 |
inductive |
26 |
(* Parameter A is the domain of zs's elements *) |
|
24892 | 27 |
|
46823 | 28 |
domains "gen_prefix(A, r)" \<subseteq> "list(A)*list(A)" |
24892 | 29 |
|
15634 | 30 |
intros |
31 |
Nil: "<[],[]>:gen_prefix(A, r)" |
|
12197 | 32 |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
33 |
prepend: "\<lbrakk>\<langle>xs,ys\<rangle>:gen_prefix(A, r); \<langle>x,y\<rangle>:r; x \<in> A; y \<in> A\<rbrakk> |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
34 |
\<Longrightarrow> <Cons(x,xs), Cons(y,ys)>: gen_prefix(A, r)" |
12197 | 35 |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
36 |
append: "\<lbrakk>\<langle>xs,ys\<rangle>:gen_prefix(A, r); zs:list(A)\<rbrakk> |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
37 |
\<Longrightarrow> <xs, ys@zs>:gen_prefix(A, r)" |
15634 | 38 |
type_intros app_type list.Nil list.Cons |
12197 | 39 |
|
24893 | 40 |
definition |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
41 |
prefix :: "i\<Rightarrow>i" where |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
42 |
"prefix(A) \<equiv> gen_prefix(A, id(A))" |
12197 | 43 |
|
24893 | 44 |
definition |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
45 |
strict_prefix :: "i\<Rightarrow>i" where |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
46 |
"strict_prefix(A) \<equiv> prefix(A) - id(list(A))" |
12197 | 47 |
|
24892 | 48 |
|
49 |
(* less or equal and greater or equal over prefixes *) |
|
50 |
||
51 |
abbreviation |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
52 |
pfixLe :: "[i, i] \<Rightarrow> o" (infixl \<open>pfixLe\<close> 50) where |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
53 |
"xs pfixLe ys \<equiv> \<langle>xs, ys\<rangle>:gen_prefix(nat, Le)" |
12197 | 54 |
|
24892 | 55 |
abbreviation |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
56 |
pfixGe :: "[i, i] \<Rightarrow> o" (infixl \<open>pfixGe\<close> 50) where |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
57 |
"xs pfixGe ys \<equiv> \<langle>xs, ys\<rangle>:gen_prefix(nat, Ge)" |
12197 | 58 |
|
24892 | 59 |
|
60 |
lemma reflD: |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
61 |
"\<lbrakk>refl(A, r); x \<in> A\<rbrakk> \<Longrightarrow> \<langle>x,x\<rangle>:r" |
15634 | 62 |
apply (unfold refl_def, auto) |
63 |
done |
|
64 |
||
65 |
(*** preliminary lemmas ***) |
|
66 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
67 |
lemma Nil_gen_prefix: "xs \<in> list(A) \<Longrightarrow> <[], xs> \<in> gen_prefix(A, r)" |
15634 | 68 |
by (drule gen_prefix.append [OF gen_prefix.Nil], simp) |
69 |
declare Nil_gen_prefix [simp] |
|
70 |
||
71 |
||
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
72 |
lemma gen_prefix_length_le: "\<langle>xs,ys\<rangle> \<in> gen_prefix(A, r) \<Longrightarrow> length(xs) \<le> length(ys)" |
15634 | 73 |
apply (erule gen_prefix.induct) |
74 |
apply (subgoal_tac [3] "ys \<in> list (A) ") |
|
75 |
apply (auto dest: gen_prefix.dom_subset [THEN subsetD] |
|
76 |
intro: le_trans simp add: length_app) |
|
77 |
done |
|
78 |
||
79 |
||
80 |
lemma Cons_gen_prefix_aux: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
81 |
"\<lbrakk><xs', ys'> \<in> gen_prefix(A, r)\<rbrakk> |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
82 |
\<Longrightarrow> (\<forall>x xs. x \<in> A \<longrightarrow> xs'= Cons(x,xs) \<longrightarrow> |
76214 | 83 |
(\<exists>y ys. y \<in> A \<and> ys' = Cons(y,ys) \<and> |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
84 |
\<langle>x,y\<rangle>:r \<and> \<langle>xs, ys\<rangle> \<in> gen_prefix(A, r)))" |
15634 | 85 |
apply (erule gen_prefix.induct) |
24892 | 86 |
prefer 3 apply (force intro: gen_prefix.append, auto) |
15634 | 87 |
done |
88 |
||
89 |
lemma Cons_gen_prefixE: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
90 |
"\<lbrakk><Cons(x,xs), zs> \<in> gen_prefix(A, r); |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
91 |
\<And>y ys. \<lbrakk>zs = Cons(y, ys); y \<in> A; x \<in> A; \<langle>x,y\<rangle>:r; |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
92 |
\<langle>xs,ys\<rangle> \<in> gen_prefix(A, r)\<rbrakk> \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P" |
24892 | 93 |
apply (frule gen_prefix.dom_subset [THEN subsetD], auto) |
94 |
apply (blast dest: Cons_gen_prefix_aux) |
|
15634 | 95 |
done |
96 |
declare Cons_gen_prefixE [elim!] |
|
97 |
||
24892 | 98 |
lemma Cons_gen_prefix_Cons: |
99 |
"(<Cons(x,xs),Cons(y,ys)> \<in> gen_prefix(A, r)) |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
100 |
\<longleftrightarrow> (x \<in> A \<and> y \<in> A \<and> \<langle>x,y\<rangle>:r \<and> \<langle>xs,ys\<rangle> \<in> gen_prefix(A, r))" |
15634 | 101 |
apply (auto intro: gen_prefix.prepend) |
102 |
done |
|
103 |
declare Cons_gen_prefix_Cons [iff] |
|
104 |
||
105 |
(** Monotonicity of gen_prefix **) |
|
106 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
107 |
lemma gen_prefix_mono2: "r<=s \<Longrightarrow> gen_prefix(A, r) \<subseteq> gen_prefix(A, s)" |
15634 | 108 |
apply clarify |
109 |
apply (frule gen_prefix.dom_subset [THEN subsetD], clarify) |
|
110 |
apply (erule rev_mp) |
|
111 |
apply (erule gen_prefix.induct) |
|
112 |
apply (auto intro: gen_prefix.append) |
|
113 |
done |
|
114 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
115 |
lemma gen_prefix_mono1: "A<=B \<Longrightarrow>gen_prefix(A, r) \<subseteq> gen_prefix(B, r)" |
15634 | 116 |
apply clarify |
117 |
apply (frule gen_prefix.dom_subset [THEN subsetD], clarify) |
|
118 |
apply (erule rev_mp) |
|
119 |
apply (erule_tac P = "y \<in> list (A) " in rev_mp) |
|
120 |
apply (erule_tac P = "xa \<in> list (A) " in rev_mp) |
|
121 |
apply (erule gen_prefix.induct) |
|
122 |
apply (simp (no_asm_simp)) |
|
123 |
apply clarify |
|
124 |
apply (erule ConsE)+ |
|
125 |
apply (auto dest: gen_prefix.dom_subset [THEN subsetD] |
|
126 |
intro: gen_prefix.append list_mono [THEN subsetD]) |
|
127 |
done |
|
128 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
129 |
lemma gen_prefix_mono: "\<lbrakk>A \<subseteq> B; r \<subseteq> s\<rbrakk> \<Longrightarrow> gen_prefix(A, r) \<subseteq> gen_prefix(B, s)" |
15634 | 130 |
apply (rule subset_trans) |
131 |
apply (rule gen_prefix_mono1) |
|
132 |
apply (rule_tac [2] gen_prefix_mono2, auto) |
|
133 |
done |
|
134 |
||
135 |
(*** gen_prefix order ***) |
|
136 |
||
137 |
(* reflexivity *) |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
138 |
lemma refl_gen_prefix: "refl(A, r) \<Longrightarrow> refl(list(A), gen_prefix(A, r))" |
15634 | 139 |
apply (unfold refl_def, auto) |
140 |
apply (induct_tac "x", auto) |
|
141 |
done |
|
142 |
declare refl_gen_prefix [THEN reflD, simp] |
|
143 |
||
144 |
(* Transitivity *) |
|
145 |
(* A lemma for proving gen_prefix_trans_comp *) |
|
146 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
147 |
lemma append_gen_prefix [rule_format (no_asm)]: "xs \<in> list(A) \<Longrightarrow> |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
148 |
\<forall>zs. <xs @ ys, zs> \<in> gen_prefix(A, r) \<longrightarrow> \<langle>xs, zs\<rangle>: gen_prefix(A, r)" |
15634 | 149 |
apply (erule list.induct) |
150 |
apply (auto dest: gen_prefix.dom_subset [THEN subsetD]) |
|
151 |
done |
|
152 |
||
153 |
(* Lemma proving transitivity and more*) |
|
154 |
||
155 |
lemma gen_prefix_trans_comp [rule_format (no_asm)]: |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
156 |
"\<langle>x, y\<rangle>: gen_prefix(A, r) \<Longrightarrow> |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
157 |
(\<forall>z \<in> list(A). \<langle>y,z\<rangle> \<in> gen_prefix(A, s)\<longrightarrow>\<langle>x, z\<rangle> \<in> gen_prefix(A, s O r))" |
15634 | 158 |
apply (erule gen_prefix.induct) |
159 |
apply (auto elim: ConsE simp add: Nil_gen_prefix) |
|
160 |
apply (subgoal_tac "ys \<in> list (A) ") |
|
161 |
prefer 2 apply (blast dest: gen_prefix.dom_subset [THEN subsetD]) |
|
162 |
apply (drule_tac xs = ys and r = s in append_gen_prefix, auto) |
|
163 |
done |
|
164 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
165 |
lemma trans_comp_subset: "trans(r) \<Longrightarrow> r O r \<subseteq> r" |
15634 | 166 |
by (auto dest: transD) |
167 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
168 |
lemma trans_gen_prefix: "trans(r) \<Longrightarrow> trans(gen_prefix(A,r))" |
15634 | 169 |
apply (simp (no_asm) add: trans_def) |
170 |
apply clarify |
|
171 |
apply (rule trans_comp_subset [THEN gen_prefix_mono2, THEN subsetD], assumption) |
|
172 |
apply (rule gen_prefix_trans_comp) |
|
173 |
apply (auto dest: gen_prefix.dom_subset [THEN subsetD]) |
|
174 |
done |
|
175 |
||
24892 | 176 |
lemma trans_on_gen_prefix: |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
177 |
"trans(r) \<Longrightarrow> trans[list(A)](gen_prefix(A, r))" |
15634 | 178 |
apply (drule_tac A = A in trans_gen_prefix) |
179 |
apply (unfold trans_def trans_on_def, blast) |
|
180 |
done |
|
181 |
||
182 |
lemma prefix_gen_prefix_trans: |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
183 |
"\<lbrakk>\<langle>x,y\<rangle> \<in> prefix(A); \<langle>y, z\<rangle> \<in> gen_prefix(A, r); r<=A*A\<rbrakk> |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
184 |
\<Longrightarrow> \<langle>x, z\<rangle> \<in> gen_prefix(A, r)" |
76216
9fc34f76b4e8
getting rid of apply (unfold ...)
paulson <lp15@cam.ac.uk>
parents:
76215
diff
changeset
|
185 |
unfolding prefix_def |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
186 |
apply (rule_tac P = "\<lambda>r. \<langle>x,z\<rangle> \<in> gen_prefix (A, r) " in right_comp_id [THEN subst]) |
15634 | 187 |
apply (blast dest: gen_prefix_trans_comp gen_prefix.dom_subset [THEN subsetD])+ |
188 |
done |
|
189 |
||
190 |
||
24892 | 191 |
lemma gen_prefix_prefix_trans: |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
192 |
"\<lbrakk>\<langle>x,y\<rangle> \<in> gen_prefix(A,r); \<langle>y, z\<rangle> \<in> prefix(A); r<=A*A\<rbrakk> |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
193 |
\<Longrightarrow> \<langle>x, z\<rangle> \<in> gen_prefix(A, r)" |
76216
9fc34f76b4e8
getting rid of apply (unfold ...)
paulson <lp15@cam.ac.uk>
parents:
76215
diff
changeset
|
194 |
unfolding prefix_def |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
195 |
apply (rule_tac P = "\<lambda>r. \<langle>x,z\<rangle> \<in> gen_prefix (A, r) " in left_comp_id [THEN subst]) |
15634 | 196 |
apply (blast dest: gen_prefix_trans_comp gen_prefix.dom_subset [THEN subsetD])+ |
197 |
done |
|
198 |
||
199 |
(** Antisymmetry **) |
|
200 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
201 |
lemma nat_le_lemma [rule_format]: "n \<in> nat \<Longrightarrow> \<forall>b \<in> nat. n #+ b \<le> n \<longrightarrow> b = 0" |
15634 | 202 |
by (induct_tac "n", auto) |
203 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
204 |
lemma antisym_gen_prefix: "antisym(r) \<Longrightarrow> antisym(gen_prefix(A, r))" |
15634 | 205 |
apply (simp (no_asm) add: antisym_def) |
206 |
apply (rule impI [THEN allI, THEN allI]) |
|
24892 | 207 |
apply (erule gen_prefix.induct, blast) |
15634 | 208 |
apply (simp add: antisym_def, blast) |
60770 | 209 |
txt\<open>append case is hardest\<close> |
15634 | 210 |
apply clarify |
211 |
apply (subgoal_tac "length (zs) = 0") |
|
212 |
apply (subgoal_tac "ys \<in> list (A) ") |
|
213 |
prefer 2 apply (blast dest: gen_prefix.dom_subset [THEN subsetD]) |
|
214 |
apply (drule_tac psi = "<ys @ zs, xs> \<in> gen_prefix (A,r) " in asm_rl) |
|
215 |
apply simp |
|
76214 | 216 |
apply (subgoal_tac "length (ys @ zs) = length (ys) #+ length (zs) \<and>ys \<in> list (A) \<and>xs \<in> list (A) ") |
15634 | 217 |
prefer 2 apply (blast intro: length_app dest: gen_prefix.dom_subset [THEN subsetD]) |
218 |
apply (drule gen_prefix_length_le)+ |
|
219 |
apply clarify |
|
220 |
apply simp |
|
221 |
apply (drule_tac j = "length (xs) " in le_trans) |
|
222 |
apply blast |
|
223 |
apply (auto intro: nat_le_lemma) |
|
224 |
done |
|
225 |
||
226 |
(*** recursion equations ***) |
|
227 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
228 |
lemma gen_prefix_Nil: "xs \<in> list(A) \<Longrightarrow> <xs, []> \<in> gen_prefix(A,r) \<longleftrightarrow> (xs = [])" |
15634 | 229 |
by (induct_tac "xs", auto) |
230 |
declare gen_prefix_Nil [simp] |
|
231 |
||
24892 | 232 |
lemma same_gen_prefix_gen_prefix: |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
233 |
"\<lbrakk>refl(A, r); xs \<in> list(A)\<rbrakk> \<Longrightarrow> |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
234 |
<xs@ys, xs@zs>: gen_prefix(A, r) \<longleftrightarrow> \<langle>ys,zs\<rangle> \<in> gen_prefix(A, r)" |
76216
9fc34f76b4e8
getting rid of apply (unfold ...)
paulson <lp15@cam.ac.uk>
parents:
76215
diff
changeset
|
235 |
unfolding refl_def |
15634 | 236 |
apply (induct_tac "xs") |
237 |
apply (simp_all (no_asm_simp)) |
|
238 |
done |
|
239 |
declare same_gen_prefix_gen_prefix [simp] |
|
240 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
241 |
lemma gen_prefix_Cons: "\<lbrakk>xs \<in> list(A); ys \<in> list(A); y \<in> A\<rbrakk> \<Longrightarrow> |
46823 | 242 |
<xs, Cons(y,ys)> \<in> gen_prefix(A,r) \<longleftrightarrow> |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
243 |
(xs=[] | (\<exists>z zs. xs=Cons(z,zs) \<and> z \<in> A \<and> \<langle>z,y\<rangle>:r \<and> \<langle>zs,ys\<rangle> \<in> gen_prefix(A,r)))" |
15634 | 244 |
apply (induct_tac "xs", auto) |
245 |
done |
|
246 |
||
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
247 |
lemma gen_prefix_take_append: "\<lbrakk>refl(A,r); \<langle>xs,ys\<rangle> \<in> gen_prefix(A, r); zs \<in> list(A)\<rbrakk> |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
248 |
\<Longrightarrow> <xs@zs, take(length(xs), ys) @ zs> \<in> gen_prefix(A, r)" |
15634 | 249 |
apply (erule gen_prefix.induct) |
250 |
apply (simp (no_asm_simp)) |
|
251 |
apply (frule_tac [!] gen_prefix.dom_subset [THEN subsetD], auto) |
|
252 |
apply (frule gen_prefix_length_le) |
|
253 |
apply (subgoal_tac "take (length (xs), ys) \<in> list (A) ") |
|
254 |
apply (simp_all (no_asm_simp) add: diff_is_0_iff [THEN iffD2] take_type) |
|
255 |
done |
|
256 |
||
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
257 |
lemma gen_prefix_append_both: "\<lbrakk>refl(A, r); \<langle>xs,ys\<rangle> \<in> gen_prefix(A,r); |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
258 |
length(xs) = length(ys); zs \<in> list(A)\<rbrakk> |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
259 |
\<Longrightarrow> <xs@zs, ys @ zs> \<in> gen_prefix(A, r)" |
15634 | 260 |
apply (drule_tac zs = zs in gen_prefix_take_append, assumption+) |
261 |
apply (subgoal_tac "take (length (xs), ys) =ys") |
|
262 |
apply (auto intro!: take_all dest: gen_prefix.dom_subset [THEN subsetD]) |
|
263 |
done |
|
264 |
||
265 |
(*NOT suitable for rewriting since [y] has the form y#ys*) |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
266 |
lemma append_cons_conv: "xs \<in> list(A) \<Longrightarrow> xs @ Cons(y, ys) = (xs @ [y]) @ ys" |
15634 | 267 |
by (auto simp add: app_assoc) |
268 |
||
269 |
lemma append_one_gen_prefix_lemma [rule_format]: |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
270 |
"\<lbrakk>\<langle>xs,ys\<rangle> \<in> gen_prefix(A, r); refl(A, r)\<rbrakk> |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
271 |
\<Longrightarrow> length(xs) < length(ys) \<longrightarrow> |
15634 | 272 |
<xs @ [nth(length(xs), ys)], ys> \<in> gen_prefix(A, r)" |
273 |
apply (erule gen_prefix.induct, blast) |
|
274 |
apply (frule gen_prefix.dom_subset [THEN subsetD], clarify) |
|
275 |
apply (simp_all add: length_type) |
|
276 |
(* Append case is hardest *) |
|
277 |
apply (frule gen_prefix_length_le [THEN le_iff [THEN iffD1]]) |
|
278 |
apply (frule gen_prefix.dom_subset [THEN subsetD], clarify) |
|
76214 | 279 |
apply (subgoal_tac "length (xs) :nat\<and>length (ys) :nat \<and>length (zs) :nat") |
15634 | 280 |
prefer 2 apply (blast intro: length_type, clarify) |
281 |
apply (simp_all add: nth_append length_type length_app) |
|
282 |
apply (rule conjI) |
|
283 |
apply (blast intro: gen_prefix.append) |
|
59788 | 284 |
apply (erule_tac V = "length (xs) < length (ys) \<longrightarrow>u" for u in thin_rl) |
15634 | 285 |
apply (erule_tac a = zs in list.cases, auto) |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
286 |
apply (rule_tac P1 = "\<lambda>x. <u(x), v>:w" for u v w in nat_diff_split [THEN iffD2]) |
15634 | 287 |
apply auto |
288 |
apply (simplesubst append_cons_conv) |
|
289 |
apply (rule_tac [2] gen_prefix.append) |
|
290 |
apply (auto elim: ConsE simp add: gen_prefix_append_both) |
|
24892 | 291 |
done |
15634 | 292 |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
293 |
lemma append_one_gen_prefix: "\<lbrakk>\<langle>xs,ys\<rangle>: gen_prefix(A, r); length(xs) < length(ys); refl(A, r)\<rbrakk> |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
294 |
\<Longrightarrow> <xs @ [nth(length(xs), ys)], ys> \<in> gen_prefix(A, r)" |
15634 | 295 |
apply (blast intro: append_one_gen_prefix_lemma) |
296 |
done |
|
297 |
||
298 |
||
299 |
(** Proving the equivalence with Charpentier's definition **) |
|
300 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
301 |
lemma gen_prefix_imp_nth_lemma [rule_format]: "xs \<in> list(A) \<Longrightarrow> |
24892 | 302 |
\<forall>ys \<in> list(A). \<forall>i \<in> nat. i < length(xs) |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
303 |
\<longrightarrow> \<langle>xs, ys\<rangle>: gen_prefix(A, r) \<longrightarrow> <nth(i, xs), nth(i, ys)>:r" |
24892 | 304 |
apply (induct_tac "xs", simp, clarify) |
305 |
apply simp |
|
306 |
apply (erule natE, auto) |
|
15634 | 307 |
done |
308 |
||
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
309 |
lemma gen_prefix_imp_nth: "\<lbrakk>\<langle>xs,ys\<rangle> \<in> gen_prefix(A,r); i < length(xs)\<rbrakk> |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
310 |
\<Longrightarrow> <nth(i, xs), nth(i, ys)>:r" |
15634 | 311 |
apply (cut_tac A = A in gen_prefix.dom_subset) |
312 |
apply (rule gen_prefix_imp_nth_lemma) |
|
313 |
apply (auto simp add: lt_nat_in_nat) |
|
314 |
done |
|
315 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
316 |
lemma nth_imp_gen_prefix [rule_format]: "xs \<in> list(A) \<Longrightarrow> |
24892 | 317 |
\<forall>ys \<in> list(A). length(xs) \<le> length(ys) |
46823 | 318 |
\<longrightarrow> (\<forall>i. i < length(xs) \<longrightarrow> <nth(i, xs), nth(i,ys)>:r) |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
319 |
\<longrightarrow> \<langle>xs, ys\<rangle> \<in> gen_prefix(A, r)" |
15634 | 320 |
apply (induct_tac "xs") |
321 |
apply (simp_all (no_asm_simp)) |
|
322 |
apply clarify |
|
323 |
apply (erule_tac a = ys in list.cases, simp) |
|
324 |
apply (force intro!: nat_0_le simp add: lt_nat_in_nat) |
|
325 |
done |
|
326 |
||
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
327 |
lemma gen_prefix_iff_nth: "(\<langle>xs,ys\<rangle> \<in> gen_prefix(A,r)) \<longleftrightarrow> |
76214 | 328 |
(xs \<in> list(A) \<and> ys \<in> list(A) \<and> length(xs) \<le> length(ys) \<and> |
46823 | 329 |
(\<forall>i. i < length(xs) \<longrightarrow> <nth(i,xs), nth(i, ys)>: r))" |
15634 | 330 |
apply (rule iffI) |
331 |
apply (frule gen_prefix.dom_subset [THEN subsetD]) |
|
24892 | 332 |
apply (frule gen_prefix_length_le, auto) |
15634 | 333 |
apply (rule_tac [2] nth_imp_gen_prefix) |
334 |
apply (drule gen_prefix_imp_nth) |
|
335 |
apply (auto simp add: lt_nat_in_nat) |
|
336 |
done |
|
337 |
||
338 |
(** prefix is a partial order: **) |
|
339 |
||
340 |
lemma refl_prefix: "refl(list(A), prefix(A))" |
|
76216
9fc34f76b4e8
getting rid of apply (unfold ...)
paulson <lp15@cam.ac.uk>
parents:
76215
diff
changeset
|
341 |
unfolding prefix_def |
15634 | 342 |
apply (rule refl_gen_prefix) |
343 |
apply (auto simp add: refl_def) |
|
344 |
done |
|
345 |
declare refl_prefix [THEN reflD, simp] |
|
346 |
||
347 |
lemma trans_prefix: "trans(prefix(A))" |
|
76216
9fc34f76b4e8
getting rid of apply (unfold ...)
paulson <lp15@cam.ac.uk>
parents:
76215
diff
changeset
|
348 |
unfolding prefix_def |
15634 | 349 |
apply (rule trans_gen_prefix) |
350 |
apply (auto simp add: trans_def) |
|
351 |
done |
|
352 |
||
45602 | 353 |
lemmas prefix_trans = trans_prefix [THEN transD] |
15634 | 354 |
|
355 |
lemma trans_on_prefix: "trans[list(A)](prefix(A))" |
|
76216
9fc34f76b4e8
getting rid of apply (unfold ...)
paulson <lp15@cam.ac.uk>
parents:
76215
diff
changeset
|
356 |
unfolding prefix_def |
15634 | 357 |
apply (rule trans_on_gen_prefix) |
358 |
apply (auto simp add: trans_def) |
|
359 |
done |
|
360 |
||
45602 | 361 |
lemmas prefix_trans_on = trans_on_prefix [THEN trans_onD] |
15634 | 362 |
|
363 |
(* Monotonicity of "set" operator WRT prefix *) |
|
364 |
||
24892 | 365 |
lemma set_of_list_prefix_mono: |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
366 |
"\<langle>xs,ys\<rangle> \<in> prefix(A) \<Longrightarrow> set_of_list(xs) \<subseteq> set_of_list(ys)" |
15634 | 367 |
|
76216
9fc34f76b4e8
getting rid of apply (unfold ...)
paulson <lp15@cam.ac.uk>
parents:
76215
diff
changeset
|
368 |
unfolding prefix_def |
15634 | 369 |
apply (erule gen_prefix.induct) |
76214 | 370 |
apply (subgoal_tac [3] "xs \<in> list (A) \<and>ys \<in> list (A) ") |
15634 | 371 |
prefer 4 apply (blast dest: gen_prefix.dom_subset [THEN subsetD]) |
372 |
apply (auto simp add: set_of_list_append) |
|
373 |
done |
|
374 |
||
375 |
(** recursion equations **) |
|
376 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
377 |
lemma Nil_prefix: "xs \<in> list(A) \<Longrightarrow> <[],xs> \<in> prefix(A)" |
15634 | 378 |
|
76216
9fc34f76b4e8
getting rid of apply (unfold ...)
paulson <lp15@cam.ac.uk>
parents:
76215
diff
changeset
|
379 |
unfolding prefix_def |
15634 | 380 |
apply (simp (no_asm_simp) add: Nil_gen_prefix) |
381 |
done |
|
382 |
declare Nil_prefix [simp] |
|
383 |
||
384 |
||
46823 | 385 |
lemma prefix_Nil: "<xs, []> \<in> prefix(A) \<longleftrightarrow> (xs = [])" |
15634 | 386 |
|
387 |
apply (unfold prefix_def, auto) |
|
388 |
apply (frule gen_prefix.dom_subset [THEN subsetD]) |
|
389 |
apply (drule_tac psi = "<xs, []> \<in> gen_prefix (A, id (A))" in asm_rl) |
|
390 |
apply (simp add: gen_prefix_Nil) |
|
391 |
done |
|
392 |
declare prefix_Nil [iff] |
|
393 |
||
24892 | 394 |
lemma Cons_prefix_Cons: |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
395 |
"<Cons(x,xs), Cons(y,ys)> \<in> prefix(A) \<longleftrightarrow> (x=y \<and> \<langle>xs,ys\<rangle> \<in> prefix(A) \<and> y \<in> A)" |
15634 | 396 |
apply (unfold prefix_def, auto) |
397 |
done |
|
398 |
declare Cons_prefix_Cons [iff] |
|
399 |
||
24892 | 400 |
lemma same_prefix_prefix: |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
401 |
"xs \<in> list(A)\<Longrightarrow> <xs@ys,xs@zs> \<in> prefix(A) \<longleftrightarrow> (\<langle>ys,zs\<rangle> \<in> prefix(A))" |
76216
9fc34f76b4e8
getting rid of apply (unfold ...)
paulson <lp15@cam.ac.uk>
parents:
76215
diff
changeset
|
402 |
unfolding prefix_def |
15634 | 403 |
apply (subgoal_tac "refl (A,id (A))") |
404 |
apply (simp (no_asm_simp)) |
|
405 |
apply (auto simp add: refl_def) |
|
406 |
done |
|
407 |
declare same_prefix_prefix [simp] |
|
408 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
409 |
lemma same_prefix_prefix_Nil: "xs \<in> list(A) \<Longrightarrow> <xs@ys,xs> \<in> prefix(A) \<longleftrightarrow> (<ys,[]> \<in> prefix(A))" |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
410 |
apply (rule_tac P = "\<lambda>x. \<langle>u, x\<rangle>:v \<longleftrightarrow> w(x)" for u v w in app_right_Nil [THEN subst]) |
15634 | 411 |
apply (rule_tac [2] same_prefix_prefix, auto) |
412 |
done |
|
413 |
declare same_prefix_prefix_Nil [simp] |
|
414 |
||
24892 | 415 |
lemma prefix_appendI: |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
416 |
"\<lbrakk>\<langle>xs,ys\<rangle> \<in> prefix(A); zs \<in> list(A)\<rbrakk> \<Longrightarrow> <xs,ys@zs> \<in> prefix(A)" |
76216
9fc34f76b4e8
getting rid of apply (unfold ...)
paulson <lp15@cam.ac.uk>
parents:
76215
diff
changeset
|
417 |
unfolding prefix_def |
15634 | 418 |
apply (erule gen_prefix.append, assumption) |
419 |
done |
|
420 |
declare prefix_appendI [simp] |
|
421 |
||
24892 | 422 |
lemma prefix_Cons: |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
423 |
"\<lbrakk>xs \<in> list(A); ys \<in> list(A); y \<in> A\<rbrakk> \<Longrightarrow> |
46823 | 424 |
<xs,Cons(y,ys)> \<in> prefix(A) \<longleftrightarrow> |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
425 |
(xs=[] | (\<exists>zs. xs=Cons(y,zs) \<and> \<langle>zs,ys\<rangle> \<in> prefix(A)))" |
76216
9fc34f76b4e8
getting rid of apply (unfold ...)
paulson <lp15@cam.ac.uk>
parents:
76215
diff
changeset
|
426 |
unfolding prefix_def |
15634 | 427 |
apply (auto simp add: gen_prefix_Cons) |
428 |
done |
|
429 |
||
24892 | 430 |
lemma append_one_prefix: |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
431 |
"\<lbrakk>\<langle>xs,ys\<rangle> \<in> prefix(A); length(xs) < length(ys)\<rbrakk> |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
432 |
\<Longrightarrow> <xs @ [nth(length(xs),ys)], ys> \<in> prefix(A)" |
76216
9fc34f76b4e8
getting rid of apply (unfold ...)
paulson <lp15@cam.ac.uk>
parents:
76215
diff
changeset
|
433 |
unfolding prefix_def |
15634 | 434 |
apply (subgoal_tac "refl (A, id (A))") |
435 |
apply (simp (no_asm_simp) add: append_one_gen_prefix) |
|
436 |
apply (auto simp add: refl_def) |
|
437 |
done |
|
438 |
||
24892 | 439 |
lemma prefix_length_le: |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
440 |
"\<langle>xs,ys\<rangle> \<in> prefix(A) \<Longrightarrow> length(xs) \<le> length(ys)" |
76216
9fc34f76b4e8
getting rid of apply (unfold ...)
paulson <lp15@cam.ac.uk>
parents:
76215
diff
changeset
|
441 |
unfolding prefix_def |
15634 | 442 |
apply (blast dest: gen_prefix_length_le) |
443 |
done |
|
444 |
||
445 |
lemma prefix_type: "prefix(A)<=list(A)*list(A)" |
|
76216
9fc34f76b4e8
getting rid of apply (unfold ...)
paulson <lp15@cam.ac.uk>
parents:
76215
diff
changeset
|
446 |
unfolding prefix_def |
15634 | 447 |
apply (blast intro!: gen_prefix.dom_subset) |
448 |
done |
|
449 |
||
24892 | 450 |
lemma strict_prefix_type: |
46823 | 451 |
"strict_prefix(A) \<subseteq> list(A)*list(A)" |
76216
9fc34f76b4e8
getting rid of apply (unfold ...)
paulson <lp15@cam.ac.uk>
parents:
76215
diff
changeset
|
452 |
unfolding strict_prefix_def |
15634 | 453 |
apply (blast intro!: prefix_type [THEN subsetD]) |
454 |
done |
|
455 |
||
24892 | 456 |
lemma strict_prefix_length_lt_aux: |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
457 |
"\<langle>xs,ys\<rangle> \<in> prefix(A) \<Longrightarrow> xs\<noteq>ys \<longrightarrow> length(xs) < length(ys)" |
76216
9fc34f76b4e8
getting rid of apply (unfold ...)
paulson <lp15@cam.ac.uk>
parents:
76215
diff
changeset
|
458 |
unfolding prefix_def |
15634 | 459 |
apply (erule gen_prefix.induct, clarify) |
76214 | 460 |
apply (subgoal_tac [!] "ys \<in> list(A) \<and> xs \<in> list(A)") |
15634 | 461 |
apply (auto dest: gen_prefix.dom_subset [THEN subsetD] |
462 |
simp add: length_type) |
|
463 |
apply (subgoal_tac "length (zs) =0") |
|
464 |
apply (drule_tac [2] not_lt_imp_le) |
|
465 |
apply (rule_tac [5] j = "length (ys) " in lt_trans2) |
|
466 |
apply auto |
|
467 |
done |
|
468 |
||
24892 | 469 |
lemma strict_prefix_length_lt: |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
470 |
"\<langle>xs,ys\<rangle>:strict_prefix(A) \<Longrightarrow> length(xs) < length(ys)" |
76216
9fc34f76b4e8
getting rid of apply (unfold ...)
paulson <lp15@cam.ac.uk>
parents:
76215
diff
changeset
|
471 |
unfolding strict_prefix_def |
15634 | 472 |
apply (rule strict_prefix_length_lt_aux [THEN mp]) |
473 |
apply (auto dest: prefix_type [THEN subsetD]) |
|
474 |
done |
|
475 |
||
476 |
(*Equivalence to the definition used in Lex/Prefix.thy*) |
|
24892 | 477 |
lemma prefix_iff: |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
478 |
"\<langle>xs,zs\<rangle> \<in> prefix(A) \<longleftrightarrow> (\<exists>ys \<in> list(A). zs = xs@ys) \<and> xs \<in> list(A)" |
76216
9fc34f76b4e8
getting rid of apply (unfold ...)
paulson <lp15@cam.ac.uk>
parents:
76215
diff
changeset
|
479 |
unfolding prefix_def |
15634 | 480 |
apply (auto simp add: gen_prefix_iff_nth lt_nat_in_nat nth_append nth_type app_type length_app) |
481 |
apply (subgoal_tac "drop (length (xs), zs) \<in> list (A) ") |
|
482 |
apply (rule_tac x = "drop (length (xs), zs) " in bexI) |
|
483 |
apply safe |
|
484 |
prefer 2 apply (simp add: length_type drop_type) |
|
485 |
apply (rule nth_equalityI) |
|
486 |
apply (simp_all (no_asm_simp) add: nth_append app_type drop_type length_app length_drop) |
|
487 |
apply (rule nat_diff_split [THEN iffD2], simp_all, clarify) |
|
488 |
apply (drule_tac i = "length (zs) " in leI) |
|
489 |
apply (force simp add: le_subset_iff, safe) |
|
490 |
apply (subgoal_tac "length (xs) #+ (i #- length (xs)) = i") |
|
491 |
apply (subst nth_drop) |
|
63648 | 492 |
apply (simp_all (no_asm_simp) add: leI split: nat_diff_split) |
15634 | 493 |
done |
494 |
||
24892 | 495 |
lemma prefix_snoc: |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
496 |
"\<lbrakk>xs \<in> list(A); ys \<in> list(A); y \<in> A\<rbrakk> \<Longrightarrow> |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
497 |
<xs, ys@[y]> \<in> prefix(A) \<longleftrightarrow> (xs = ys@[y] | \<langle>xs,ys\<rangle> \<in> prefix(A))" |
15634 | 498 |
apply (simp (no_asm) add: prefix_iff) |
499 |
apply (rule iffI, clarify) |
|
500 |
apply (erule_tac xs = ysa in rev_list_elim, simp) |
|
501 |
apply (simp add: app_type app_assoc [symmetric]) |
|
502 |
apply (auto simp add: app_assoc app_type) |
|
503 |
done |
|
504 |
declare prefix_snoc [simp] |
|
505 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
506 |
lemma prefix_append_iff [rule_format]: "zs \<in> list(A) \<Longrightarrow> \<forall>xs \<in> list(A). \<forall>ys \<in> list(A). |
46823 | 507 |
(<xs, ys@zs> \<in> prefix(A)) \<longleftrightarrow> |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
508 |
(\<langle>xs,ys\<rangle> \<in> prefix(A) | (\<exists>us. xs = ys@us \<and> \<langle>us,zs\<rangle> \<in> prefix(A)))" |
24892 | 509 |
apply (erule list_append_induct, force, clarify) |
510 |
apply (rule iffI) |
|
15634 | 511 |
apply (simp add: add: app_assoc [symmetric]) |
24892 | 512 |
apply (erule disjE) |
513 |
apply (rule disjI2) |
|
514 |
apply (rule_tac x = "y @ [x]" in exI) |
|
15634 | 515 |
apply (simp add: add: app_assoc [symmetric], force+) |
516 |
done |
|
517 |
||
518 |
||
519 |
(*Although the prefix ordering is not linear, the prefixes of a list |
|
520 |
are linearly ordered.*) |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
521 |
lemma common_prefix_linear_lemma [rule_format]: "\<lbrakk>zs \<in> list(A); xs \<in> list(A); ys \<in> list(A)\<rbrakk> |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
522 |
\<Longrightarrow> \<langle>xs, zs\<rangle> \<in> prefix(A) \<longrightarrow> \<langle>ys,zs\<rangle> \<in> prefix(A) |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
523 |
\<longrightarrow>\<langle>xs,ys\<rangle> \<in> prefix(A) | \<langle>ys,xs\<rangle> \<in> prefix(A)" |
15634 | 524 |
apply (erule list_append_induct, auto) |
525 |
done |
|
526 |
||
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
527 |
lemma common_prefix_linear: "\<lbrakk>\<langle>xs, zs\<rangle> \<in> prefix(A); \<langle>ys,zs\<rangle> \<in> prefix(A)\<rbrakk> |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
528 |
\<Longrightarrow> \<langle>xs,ys\<rangle> \<in> prefix(A) | \<langle>ys,xs\<rangle> \<in> prefix(A)" |
15634 | 529 |
apply (cut_tac prefix_type) |
530 |
apply (blast del: disjCI intro: common_prefix_linear_lemma) |
|
531 |
done |
|
532 |
||
533 |
||
534 |
(*** pfixLe, pfixGe \<in> properties inherited from the translations ***) |
|
535 |
||
536 |
||
537 |
||
538 |
(** pfixLe **) |
|
539 |
||
540 |
lemma refl_Le: "refl(nat,Le)" |
|
541 |
||
542 |
apply (unfold refl_def, auto) |
|
543 |
done |
|
544 |
declare refl_Le [simp] |
|
545 |
||
546 |
lemma antisym_Le: "antisym(Le)" |
|
76216
9fc34f76b4e8
getting rid of apply (unfold ...)
paulson <lp15@cam.ac.uk>
parents:
76215
diff
changeset
|
547 |
unfolding antisym_def |
15634 | 548 |
apply (auto intro: le_anti_sym) |
549 |
done |
|
550 |
declare antisym_Le [simp] |
|
551 |
||
552 |
lemma trans_on_Le: "trans[nat](Le)" |
|
553 |
apply (unfold trans_on_def, auto) |
|
554 |
apply (blast intro: le_trans) |
|
555 |
done |
|
556 |
declare trans_on_Le [simp] |
|
557 |
||
558 |
lemma trans_Le: "trans(Le)" |
|
559 |
apply (unfold trans_def, auto) |
|
560 |
apply (blast intro: le_trans) |
|
561 |
done |
|
562 |
declare trans_Le [simp] |
|
563 |
||
564 |
lemma part_order_Le: "part_order(nat,Le)" |
|
565 |
by (unfold part_order_def, auto) |
|
566 |
declare part_order_Le [simp] |
|
567 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
568 |
lemma pfixLe_refl: "x \<in> list(nat) \<Longrightarrow> x pfixLe x" |
15634 | 569 |
by (blast intro: refl_gen_prefix [THEN reflD] refl_Le) |
570 |
declare pfixLe_refl [simp] |
|
571 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
572 |
lemma pfixLe_trans: "\<lbrakk>x pfixLe y; y pfixLe z\<rbrakk> \<Longrightarrow> x pfixLe z" |
15634 | 573 |
by (blast intro: trans_gen_prefix [THEN transD] trans_Le) |
574 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
575 |
lemma pfixLe_antisym: "\<lbrakk>x pfixLe y; y pfixLe x\<rbrakk> \<Longrightarrow> x = y" |
15634 | 576 |
by (blast intro: antisym_gen_prefix [THEN antisymE] antisym_Le) |
577 |
||
578 |
||
24892 | 579 |
lemma prefix_imp_pfixLe: |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
580 |
"\<langle>xs,ys\<rangle>:prefix(nat)\<Longrightarrow> xs pfixLe ys" |
15634 | 581 |
|
76216
9fc34f76b4e8
getting rid of apply (unfold ...)
paulson <lp15@cam.ac.uk>
parents:
76215
diff
changeset
|
582 |
unfolding prefix_def |
15634 | 583 |
apply (rule gen_prefix_mono [THEN subsetD], auto) |
584 |
done |
|
585 |
||
586 |
lemma refl_Ge: "refl(nat, Ge)" |
|
587 |
by (unfold refl_def Ge_def, auto) |
|
588 |
declare refl_Ge [iff] |
|
589 |
||
590 |
lemma antisym_Ge: "antisym(Ge)" |
|
76217 | 591 |
unfolding antisym_def Ge_def |
15634 | 592 |
apply (auto intro: le_anti_sym) |
593 |
done |
|
594 |
declare antisym_Ge [iff] |
|
595 |
||
596 |
lemma trans_Ge: "trans(Ge)" |
|
76217 | 597 |
unfolding trans_def Ge_def |
15634 | 598 |
apply (auto intro: le_trans) |
599 |
done |
|
600 |
declare trans_Ge [iff] |
|
601 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
602 |
lemma pfixGe_refl: "x \<in> list(nat) \<Longrightarrow> x pfixGe x" |
15634 | 603 |
by (blast intro: refl_gen_prefix [THEN reflD]) |
604 |
declare pfixGe_refl [simp] |
|
605 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
606 |
lemma pfixGe_trans: "\<lbrakk>x pfixGe y; y pfixGe z\<rbrakk> \<Longrightarrow> x pfixGe z" |
15634 | 607 |
by (blast intro: trans_gen_prefix [THEN transD]) |
608 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
609 |
lemma pfixGe_antisym: "\<lbrakk>x pfixGe y; y pfixGe x\<rbrakk> \<Longrightarrow> x = y" |
15634 | 610 |
by (blast intro: antisym_gen_prefix [THEN antisymE]) |
611 |
||
24892 | 612 |
lemma prefix_imp_pfixGe: |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
613 |
"\<langle>xs,ys\<rangle>:prefix(nat) \<Longrightarrow> xs pfixGe ys" |
76217 | 614 |
unfolding prefix_def Ge_def |
15634 | 615 |
apply (rule gen_prefix_mono [THEN subsetD], auto) |
616 |
done |
|
617 |
(* Added by Sidi \<in> prefix and take *) |
|
618 |
||
24892 | 619 |
lemma prefix_imp_take: |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
620 |
"\<langle>xs, ys\<rangle> \<in> prefix(A) \<Longrightarrow> xs = take(length(xs), ys)" |
15634 | 621 |
|
76216
9fc34f76b4e8
getting rid of apply (unfold ...)
paulson <lp15@cam.ac.uk>
parents:
76215
diff
changeset
|
622 |
unfolding prefix_def |
15634 | 623 |
apply (erule gen_prefix.induct) |
624 |
apply (subgoal_tac [3] "length (xs) :nat") |
|
625 |
apply (auto dest: gen_prefix.dom_subset [THEN subsetD] simp add: length_type) |
|
626 |
apply (frule gen_prefix.dom_subset [THEN subsetD]) |
|
627 |
apply (frule gen_prefix_length_le) |
|
628 |
apply (auto simp add: take_append) |
|
629 |
apply (subgoal_tac "length (xs) #- length (ys) =0") |
|
630 |
apply (simp_all (no_asm_simp) add: diff_is_0_iff) |
|
631 |
done |
|
632 |
||
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
633 |
lemma prefix_length_equal: "\<lbrakk>\<langle>xs,ys\<rangle> \<in> prefix(A); length(xs)=length(ys)\<rbrakk> \<Longrightarrow> xs = ys" |
15634 | 634 |
apply (cut_tac A = A in prefix_type) |
635 |
apply (drule subsetD, auto) |
|
636 |
apply (drule prefix_imp_take) |
|
637 |
apply (erule trans, simp) |
|
638 |
done |
|
639 |
||
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
640 |
lemma prefix_length_le_equal: "\<lbrakk>\<langle>xs,ys\<rangle> \<in> prefix(A); length(ys) \<le> length(xs)\<rbrakk> \<Longrightarrow> xs = ys" |
15634 | 641 |
by (blast intro: prefix_length_equal le_anti_sym prefix_length_le) |
642 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
643 |
lemma take_prefix [rule_format]: "xs \<in> list(A) \<Longrightarrow> \<forall>n \<in> nat. <take(n, xs), xs> \<in> prefix(A)" |
76216
9fc34f76b4e8
getting rid of apply (unfold ...)
paulson <lp15@cam.ac.uk>
parents:
76215
diff
changeset
|
644 |
unfolding prefix_def |
15634 | 645 |
apply (erule list.induct, simp, clarify) |
646 |
apply (erule natE, auto) |
|
647 |
done |
|
648 |
||
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
649 |
lemma prefix_take_iff: "\<langle>xs,ys\<rangle> \<in> prefix(A) \<longleftrightarrow> (xs=take(length(xs), ys) \<and> xs \<in> list(A) \<and> ys \<in> list(A))" |
15634 | 650 |
apply (rule iffI) |
651 |
apply (frule prefix_type [THEN subsetD]) |
|
652 |
apply (blast intro: prefix_imp_take, clarify) |
|
653 |
apply (erule ssubst) |
|
654 |
apply (blast intro: take_prefix length_type) |
|
655 |
done |
|
656 |
||
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
657 |
lemma prefix_imp_nth: "\<lbrakk>\<langle>xs,ys\<rangle> \<in> prefix(A); i < length(xs)\<rbrakk> \<Longrightarrow> nth(i,xs) = nth(i,ys)" |
15634 | 658 |
by (auto dest!: gen_prefix_imp_nth simp add: prefix_def) |
659 |
||
24892 | 660 |
lemma nth_imp_prefix: |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
661 |
"\<lbrakk>xs \<in> list(A); ys \<in> list(A); length(xs) \<le> length(ys); |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
662 |
\<And>i. i < length(xs) \<Longrightarrow> nth(i, xs) = nth(i,ys)\<rbrakk> |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
663 |
\<Longrightarrow> \<langle>xs,ys\<rangle> \<in> prefix(A)" |
15634 | 664 |
apply (auto simp add: prefix_def nth_imp_gen_prefix) |
665 |
apply (auto intro!: nth_imp_gen_prefix simp add: prefix_def) |
|
666 |
apply (blast intro: nth_type lt_trans2) |
|
667 |
done |
|
668 |
||
669 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
670 |
lemma length_le_prefix_imp_prefix: "\<lbrakk>length(xs) \<le> length(ys); |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
671 |
\<langle>xs,zs\<rangle> \<in> prefix(A); \<langle>ys,zs\<rangle> \<in> prefix(A)\<rbrakk> \<Longrightarrow> \<langle>xs,ys\<rangle> \<in> prefix(A)" |
15634 | 672 |
apply (cut_tac A = A in prefix_type) |
673 |
apply (rule nth_imp_prefix, blast, blast) |
|
674 |
apply assumption |
|
675 |
apply (rule_tac b = "nth (i,zs)" in trans) |
|
676 |
apply (blast intro: prefix_imp_nth) |
|
677 |
apply (blast intro: sym prefix_imp_nth prefix_length_le lt_trans2) |
|
678 |
done |
|
679 |
||
12197 | 680 |
end |