2 ||M|| This file is part of HELM, an Hypertextual, Electronic
3 ||A|| Library of Mathematics, developed at the Computer Science
4 ||T|| Department of the University of Bologna, Italy.
8 \ / This file is distributed under the terms of the
9 \ / GNU General Public License Version 2
10 V_____________________________________________________________*)
17 include "turing/universal/copy.ma".
23 if is_true(current) (* current state is final *)
35 if is_marked(current) = false (* match ok *)
48 case bit false: move_tape_l
49 case bit true: move_tape_r
50 case null: adv_to_grid_l; move_l; adv_to_grid_l;
57 definition init_match ≝
59 (seq ? (adv_to_mark_r ? (λc:STape.is_grid (\fst c)))
63 (adv_to_mark_l ? (is_marked ?)))))).
65 definition R_init_match ≝ λt1,t2.
66 ∀ls,l,rs,c,d. no_grids (〈c,false〉::l) → no_marks l →
67 t1 = midtape STape ls 〈c,false〉 (l@〈grid,false〉::〈d,false〉::rs) →
68 t2 = midtape STape ls 〈c,true〉 (l@〈grid,false〉::〈d,true〉::rs).
70 lemma sem_init_match : Realize ? init_match R_init_match.
72 cases (sem_seq ????? (sem_mark ?)
73 (sem_seq ????? (sem_adv_to_mark_r ? (λc:STape.is_grid (\fst c)))
74 (sem_seq ????? (sem_move_r ?)
75 (sem_seq ????? (sem_mark ?)
76 (sem_seq ????? (sem_move_l ?)
77 (sem_adv_to_mark_l ? (is_marked ?)))))) intape)
78 #k * #outc * #Hloop #HR
79 @(ex_intro ?? k) @(ex_intro ?? outc) % [@Hloop] -Hloop
80 #ls #l #rs #c #d #Hnogrids #Hnomarks #Hintape
82 #ta * whd in ⊢ (%→?); #Hta lapply (Hta … Hintape) -Hta -Hintape #Hta
83 * #tb * whd in ⊢ (%→?); #Htb cases (Htb … Hta) -Htb -Hta
84 [* #Hgridc @False_ind @(absurd … Hgridc) @eqnot_to_noteq
85 @(Hnogrids 〈c,false〉) @memb_hd ]
86 * #Hgrdic #Htb lapply (Htb l 〈grid,false〉 (〈d,false〉::rs) (refl …) (refl …) ?)
87 [#x #membl @Hnogrids @memb_cons @membl] -Htb #Htb
88 * #tc * whd in ⊢ (%→?); #Htc lapply (Htc … Htb) -Htc -Htb #Htc
89 * #td * whd in ⊢ (%→?); #Htd lapply (Htd … Htc) -Htd -Htc #Htd
90 * #te * whd in ⊢ (%→?); #Hte lapply (Hte … Htd) -Hte -Htd #Hte
91 whd in ⊢ (%→?); #Htf cases (Htf … Hte) -Htf -Hte
92 [* whd in ⊢ ((??%?)→?); #Habs destruct (Habs)]
93 * #_ #Htf lapply (Htf (reverse ? l) 〈c,true〉 ls (refl …) (refl …) ?)
94 [#x #membl @Hnomarks @daemon] -Htf #Htf >Htf >reverse_reverse %
100 init_current_on_match; (* no marks in current *)
107 definition init_copy ≝
108 seq ? init_current_on_match
110 (seq ? (adv_to_mark_r ? (is_marked ?))
113 definition R_init_copy ≝ λt1,t2.
115 no_marks l1 → no_grids l1 →
116 no_marks l2 → is_grid c = false →
117 t1 = midtape STape (l1@〈c,false〉::〈grid,false〉::ls) 〈grid,false〉 (l2@〈comma,true〉::〈d,false〉::rs) →
118 t2 = midtape STape (〈comma,false〉::(reverse ? l2)@〈grid,false〉::l1@〈c,true〉::〈grid,false〉::ls) 〈d,true〉 rs.
120 lemma list_last: ∀A.∀l:list A.
121 l = [ ] ∨ ∃a,l1. l = l1@[a].
122 #A #l <(reverse_reverse ? l) cases (reverse A l)
124 |#a #l1 %2 @(ex_intro ?? a) @(ex_intro ?? (reverse ? l1)) //
128 lemma sem_init_copy : Realize ? init_copy R_init_copy.
130 cases (sem_seq ????? sem_init_current_on_match
131 (sem_seq ????? (sem_move_r ?)
132 (sem_seq ????? (sem_adv_to_mark_r ? (is_marked ?))
133 (sem_adv_mark_r ?))) intape)
134 #k * #outc * #Hloop #HR
135 @(ex_intro ?? k) @(ex_intro ?? outc) % [@Hloop] -Hloop
136 #l1 #l2 #c #ls #d #rs #Hl1marks #Hl1grids #Hl2marks #Hc #Hintape
138 #ta * whd in ⊢ (%→?); #Hta lapply (Hta … Hl1grids Hc Hintape) -Hta -Hintape #Hta
139 * #tb * whd in ⊢ (%→?); #Htb lapply (Htb … Hta) -Htb -Hta
140 generalize in match Hl1marks; -Hl1marks cases (list_last ? l1)
141 [#eql1 >eql1 #Hl1marks whd in ⊢ ((???%)→?); whd in ⊢ ((???(????%))→?); #Htb
142 * #tc * whd in ⊢ (%→?); #Htc lapply (Htc … Htb) -Htc -Htb *
143 [* whd in ⊢ ((??%?)→?); #Htemp destruct (Htemp)]
144 * #_ #Htc lapply (Htc … (refl …) (refl …) ?)
145 [#x #membx @Hl2marks @membx]
146 #Htc whd in ⊢ (%→?); #Houtc lapply (Houtc … Htc) -Houtc -Htc #Houtc
148 |* #c1 * #tl #eql1 >eql1 #Hl1marks >reverse_append >reverse_single
149 whd in ⊢ ((???%)→?); whd in ⊢ ((???(????%))→?);
150 >associative_append whd in ⊢ ((???(????%))→?); #Htb
151 * #tc * whd in ⊢ (%→?); #Htc lapply (Htc … Htb) -Htc -Htb *
152 [* >Hl1marks [#Htemp destruct (Htemp)] @memb_append_l2 @memb_hd]
153 * #_ >append_cons <associative_append #Htc lapply (Htc … (refl …) (refl …) ?)
154 [#x #membx cases (memb_append … membx) -membx #membx
155 [cases (memb_append … membx) -membx #membx
156 [@Hl1marks @memb_append_l1 @daemon
157 |>(memb_single … membx) %
161 #Htc whd in ⊢ (%→?); #Houtc lapply (Houtc … Htc) -Houtc -Htc #Houtc
162 >Houtc >reverse_append >reverse_append >reverse_single
163 >reverse_reverse >associative_append >associative_append
164 >associative_append %
168 definition init_copy ≝
170 (seq ? init_current_on_match
172 (adv_to_mark_r ? (is_marked ?)))).
174 definition R_init_copy ≝ λt1,t2.
176 no_marks l1 → no_grids l1 →
177 no_marks l2 → no_grids l2 → is_grid c = false → is_grid d =false →
178 t1 = midtape STape (l1@〈grid,false〉::l2@〈c,false〉::〈grid,false〉::l3) 〈comma,true〉 (〈d,false〉::rs) →
179 t2 = midtape STape (〈comma,false〉::l1@〈grid,false〉::l2@〈c,true〉::〈grid,false〉::l3) 〈d,true〉 rs.
181 lemma list_last: ∀A.∀l:list A.
182 l = [ ] ∨ ∃a,l1. l = l1@[a].
183 #A #l <(reverse_reverse ? l) cases (reverse A l)
185 |#a #l1 %2 @(ex_intro ?? a) @(ex_intro ?? (reverse ? l1)) //
189 lemma sem_init_copy : Realize ? init_copy R_init_copy.
191 cases (sem_seq ????? (sem_adv_mark_r ?)
192 (sem_seq ????? sem_init_current_on_match
193 (sem_seq ????? (sem_move_r ?)
194 (sem_adv_to_mark_r ? (is_marked ?)))) intape)
195 #k * #outc * #Hloop #HR
196 @(ex_intro ?? k) @(ex_intro ?? outc) % [@Hloop] -Hloop
197 #l1 #l2 #c #l3 #d #rs #Hl1marks #Hl1grids #Hl2marks #Hl2grids #Hc #Hd #Hintape
199 #ta * whd in ⊢ (%→?); #Hta lapply (Hta … Hintape) -Hta -Hintape #Hta
200 * #tb * whd in ⊢ (%→?);
201 >append_cons #Htb lapply (Htb (〈comma,false〉::l1) l2 c … Hta)
203 |#x #membx cases (orb_true_l … membx) -membx #membx
204 [>(\P membx) // | @Hl1grids @membx]
206 * #tc * whd in ⊢ (%→?); #Htc lapply (Htc … Htb) -Htc -Htb
207 >reverse_append >reverse_cons cases (list_last ? l2)
208 [#Hl2 >Hl2 >associative_append whd in ⊢ ((???(??%%%))→?); #Htc
209 whd in ⊢ (%→?); #Htd cases (Htd … Htc) -Htd -Htc
210 [* whd in ⊢ ((??%?)→?); #Habs destruct (Habs)]
211 * #_ #Htf lapply (Htf … (refl …) (refl …) ?)
212 [#x >reverse_cons #membx cases (memb_append … membx) -membx #membx
213 [@Hl1marks @daemon |>(memb_single … membx) //]
215 |#Htf >Htf >reverse_reverse >associative_append %
217 |* #a * #l21 #Heq >Heq >reverse_append >reverse_single
218 >associative_append >associative_append >associative_append whd in ⊢ ((???(??%%%))→?); #Htc
219 whd in ⊢ (%→?); #Htd cases (Htd … Htc) -Htd -Htc
220 [* >Hl2marks [#Habs destruct (Habs) |>Heq @memb_append_l2 @memb_hd]]
221 * #_ <associative_append <associative_append #Htf lapply (Htf … (refl …) (refl …) ?)
222 [#x >reverse_cons #membx cases (memb_append … membx) -membx #membx
223 [cases (memb_append … membx) -membx #membx
224 [@Hl2marks >Heq @memb_append_l1 @daemon
225 |>(memb_single … membx) //]
226 |cases (memb_append … membx) -membx #membx
227 [@Hl1marks @daemon |>(memb_single … membx) //]
229 | #Htf >Htf >reverse_append >reverse_reverse
230 >reverse_append >reverse_reverse >associative_append
231 >reverse_single >associative_append >associative_append
232 >associative_append %
237 include "turing/universal/move_tape.ma".
239 definition exec_move ≝
242 (seq ? (move_r …) move_tape)).
244 definition map_move ≝
245 λc,mv.match c with [ null ⇒ None ? | _ ⇒ Some ? 〈c,false,move_of_unialpha mv〉 ].
247 (* - aggiungere a legal_tape le condizioni
248 only_bits ls, rs; bit_or_null c
249 - ci vuole un lemma che dimostri
250 bit_or_null c1 = true bit_or_null mv = true
251 mv ≠ null → c1 ≠ null
252 dal fatto che c1 e mv sono contenuti nella table
254 definition R_exec_move ≝ λt1,t2.
255 ∀n,curconfig,ls,rs,c0,c1,s0,s1,table1,newconfig,mv,table2.
256 table_TM n (table1@〈comma,false〉::〈s1,false〉::newconfig@〈c1,false〉::〈comma,false〉::〈mv,false〉::table2) →
257 no_marks curconfig → only_bits (curconfig@[〈s0,false〉]) → only_bits (〈s1,false〉::newconfig) →
258 |curconfig| = |newconfig| →
259 no_nulls ls → no_nulls rs → no_marks ls → no_marks rs →
260 legal_tape ls 〈c0,false〉 rs →
261 t1 = midtape STape (〈c0,false〉::curconfig@〈s0,false〉::〈grid,false〉::ls) 〈grid,false〉
262 (table1@〈comma,true〉::〈s1,false〉::newconfig@〈c1,false〉::〈comma,false〉::〈mv,false〉::table2@〈grid,false〉::rs) →
263 ∀t1'.t1' = lift_tape ls 〈c0,false〉 rs →
265 t2 = midtape STape ls1 〈grid,false〉
266 (〈s1,false〉::newconfig@〈c2,false〉::〈grid,false〉::
267 table1@〈comma,false〉::〈s1,false〉::newconfig@〈c1,false〉::〈comma,false〉::〈mv,false〉::table2@〈grid,false〉::rs1) ∧
268 lift_tape ls1 〈c2,false〉 rs1 =
269 tape_move STape t1' (map_move c1 mv) ∧ legal_tape ls1 〈c2,false〉 rs1.
271 (* move the following 2 lemmata to mono.ma *)
272 lemma tape_move_left_eq :
274 tape_move ? t (Some ? 〈c,L〉) =
275 tape_move_left ? (left ? t) c (right ? t).
279 lemma tape_move_right_eq :
281 tape_move ? t (Some ? 〈c,R〉) =
282 tape_move_right ? (left ? t) c (right ? t).
286 lemma lift_tape_not_null :
287 ∀ls,c,bc,rs.c ≠ null → lift_tape ls 〈c,bc〉 rs = midtape ? ls 〈c,bc〉 rs.
288 #ls #c #bc #rs cases c //
289 #Hfalse @False_ind /2/
292 lemma merge_char_not_null :
293 ∀c1,c2.c1 ≠ null → merge_char c1 c2 ≠ null.
294 #c1 #c2 @not_to_not cases c2
295 [ #c1' normalize #Hfalse destruct (Hfalse)
297 | *: normalize #Hfalse destruct (Hfalse)
301 lemma sem_exec_move : Realize ? exec_move R_exec_move.
303 cases (sem_seq … sem_init_copy
305 (sem_seq … (sem_move_r …) sem_move_tape )) intape)
306 #k * #outc * #Hloop #HR
307 @(ex_intro ?? k) @(ex_intro ?? outc) % [ @Hloop ] -Hloop
308 #n #curconfig #ls #rs #c0 #c1 #s0 #s1 #table1 #newconfig #mv #table2
309 #Htable #Hcurconfig1 #Hcurconfig2 #Hnewconfig #Hlen #Hls #Hrs #Hls1 #Hrs1 #Htape #Hintape #t1' #Ht1'
310 cases HR -HR #ta * whd in ⊢ (%→?); #Hta
311 lapply (Hta (〈c0,false〉::curconfig) table1 s0 ls s1
312 (newconfig@〈c1,false〉::〈comma,false〉::〈mv,false〉::table2@〈grid,false〉::rs) … Hintape) -Hta
313 [ (*Hcurconfig2*) @daemon
315 | (*bit_or_null c0 = true *) @daemon
316 | (*Hcurconfig1*) @daemon
317 | #Hta * #tb * whd in ⊢ (%→?); #Htb
318 lapply (Htb (〈grid,false〉::ls) s0 s1 c0 c1 (〈mv,false〉::table2@〈grid,false〉::rs) newconfig (〈comma,false〉::reverse ? table1) curconfig Hta ????????) -Htb
319 [9:|*:(* bit_or_null c0,c1; |curconfig| = |newconfig|*) @daemon ]
320 #Htb * #tc * whd in ⊢ (%→?); #Htc lapply (Htc … Htb) -Htc whd in ⊢(???(??%%%)→?);#Htc
321 whd in ⊢ (%→?); #Houtc whd in Htc:(???%); whd in Htc:(???(??%%%));
323 (〈comma,false〉::〈c1,false〉::reverse ? newconfig@〈s1,false〉::〈comma,false〉::reverse ? table1)
324 mv table2 (merge_char c0 c1) (reverse ? newconfig@[〈s1,false〉]) ls ?????
325 Hls Hrs Hls1 Hrs1 ???)
326 [3: cases (true_or_false (c0 == null)) #Hc0
327 [ >(\P Hc0) in Htape; #Htape cases (legal_tape_cases … Htape)
328 [ * #Hfalse @False_ind /2/
330 [ #Hlsnil @legal_tape_conditions % %2 //
331 | #Hrsnil @legal_tape_conditions %2 //
334 | @legal_tape_conditions % % @merge_char_not_null @(\Pf Hc0) ]
335 |4:>Htc @(eq_f3 … (midtape ?))
336 [ @eq_f @eq_f >associative_append >associative_append %
340 || >reverse_cons >reverse_cons >reverse_append >reverse_reverse
341 >reverse_cons >reverse_cons >reverse_reverse
342 >associative_append >associative_append >associative_append
343 >associative_append >associative_append
345 | (* Hnewconfig *) @daemon
346 | (* bit_or_null mv = true *) @daemon
347 | (* bit_or_null c1,c2 → bit_or_null (merge_char c1 c2) *) @daemon
348 | -Houtc * #ls1 * #rs1 * #newc * #Hnewtapelegal * #Houtc *
351 @(ex_intro ?? ls1) @(ex_intro ?? rs1) @(ex_intro ?? newc)
354 [ >Houtc -Houtc >reverse_append
355 >reverse_reverse >reverse_single @eq_f
356 >reverse_cons >reverse_cons >reverse_append >reverse_cons
357 >reverse_cons >reverse_reverse >reverse_reverse
358 >associative_append >associative_append
359 >associative_append >associative_append
360 >associative_append >associative_append %
361 | >Hmv >Ht1' >Htapemove
362 (* mv = bit false -→ c1 = bit ? *)
363 cut (∃c1'.c1 = bit c1') [ @daemon ] * #c1' #Hc1
364 >Hc1 >tape_move_left_eq cases (legal_tape_cases … Htape)
365 [ #Hc0 >lift_tape_not_null /2/
367 [ #Hls >Hc0 >Hls whd in ⊢ (??%%); cases rs;
369 | #Hrs >Hc0 >Hrs cases ls; [%| #l0 #ls0 %]
376 @(ex_intro ?? ls1) @(ex_intro ?? rs1) @(ex_intro ?? newc) %
378 [ >Houtc -Houtc >reverse_append
379 >reverse_reverse >reverse_single @eq_f
380 >reverse_cons >reverse_cons >reverse_append >reverse_cons
381 >reverse_cons >reverse_reverse >reverse_reverse
382 >associative_append >associative_append
383 >associative_append >associative_append
384 >associative_append >associative_append %
385 |>Hmv >Ht1' >Htapemove
386 cut (∃c1'.c1 = bit c1') [ @daemon ] * #c1' #Hc1
387 >Hc1 >tape_move_right_eq cases (legal_tape_cases … Htape)
388 [ #Hc0 >lift_tape_not_null /2/
390 [ #Hls >Hc0 >Hls whd in ⊢ (??%%); cases rs;
392 | #Hrs >Hc0 >Hrs cases ls; [%| #l0 #ls0 %]
399 | * * * #Hmv #Hlseq #Hrseq #Hnewc
400 @(ex_intro ?? ls1) @(ex_intro ?? rs1) @(ex_intro ?? newc) %
402 [ >Houtc -Houtc >reverse_append
403 >reverse_reverse >reverse_single @eq_f
404 >reverse_cons >reverse_cons >reverse_append >reverse_cons
405 >reverse_cons >reverse_reverse >reverse_reverse
406 >associative_append >associative_append
407 >associative_append >associative_append
408 >associative_append >associative_append %
409 |>Hmv >Ht1' cases c1 in Hnewc;
410 [ #c1' whd in ⊢ (??%?→?);#Hnewc <Hnewc
411 >Hlseq >Hrseq cases Htape * #Hleft #Hright #_
412 whd in ⊢ (??%%); >Hleft >Hright %
413 | whd in ⊢ (??%?→?); #Hnewc >Hnewc >Hlseq >Hrseq %
414 |*: whd in ⊢ (??%?→?);#Hnewc <Hnewc
415 >Hlseq >Hrseq cases Htape * #Hleft #Hright #_
416 whd in ⊢ (??%%); >Hleft >Hright % ]
425 if is_false(current) (* current state is not final *)
428 if is_marked(current) = false (* match ok *)
436 definition uni_step ≝
437 ifTM ? (test_char STape (λc.\fst c == bit false))
438 (single_finalTM ? (seq ? init_match
440 (ifTM ? (test_char ? (λc.¬is_marked ? c))
441 (seq ? exec_move (move_r …))
447 definition R_uni_step_true ≝ λt1,t2.
448 ∀n,t0,table,s0,s1,c0,c1,ls,rs,curconfig,newconfig,mv.
449 table_TM (S n) (〈t0,false〉::table) →
450 match_in_table (〈s0,false〉::curconfig@[〈c0,false〉])
451 (〈s1,false〉::newconfig@[〈c1,false〉]) mv (〈t0,false〉::table) →
452 legal_tape ls 〈c0,false〉 rs →
453 t1 = midtape STape (〈grid,false〉::ls) 〈s0,false〉
454 (curconfig@〈c0,false〉::〈grid,false〉::〈t0,false〉::table@〈grid,false〉::rs) →
455 ∀t1'.t1' = lift_tape ls 〈c0,false〉 rs →
458 (t2 = midtape STape (〈grid,false〉::ls1) 〈s1,false〉
459 (newconfig@〈c2,false〉::〈grid,false〉::〈t0,false〉::table@〈grid,false〉::rs1) ∧
460 lift_tape ls1 〈c2,false〉 rs1 =
461 tape_move STape t1' (map_move c1 mv) ∧ legal_tape ls1 〈c2,false〉 rs1).
463 definition R_uni_step_false ≝ λt1,t2.
464 ∀b. current STape t1 = Some ? 〈bit b,false〉 → b = true ∧ t2 = t1.
466 axiom sem_match_tuple : Realize ? match_tuple R_match_tuple.
469 accRealize ? uni_step (inr … (inl … (inr … 0)))
470 R_uni_step_true R_uni_step_false.
471 @(acc_sem_if_app … (sem_test_char ? (λc:STape.\fst c == bit false))
472 (sem_seq … sem_init_match
473 (sem_seq … sem_match_tuple
474 (sem_if … (sem_test_char ? (λc.¬is_marked ? c))
475 (sem_seq … sem_exec_move (sem_move_r …))
480 #ta whd in ⊢ (%→?); #Hta #HR
481 #n #t0 #table #s0 #s1 #c0 #c1 #ls #rs #curconfig #newconfig #mv
482 #Htable #Hmatch #Htape #Hintape #t1' #Ht1'
483 >Hintape in Hta; #Hta cases (Hta ? (refl ??)) -Hta
484 #Hs0 lapply (\P Hs0) -Hs0 #Hs0 #Hta % //
486 #tb * whd in ⊢ (%→?); #Htb
487 lapply (Htb (〈grid,false〉::ls) (curconfig@[〈c0,false〉]) (table@〈grid,false〉::rs) s0 t0 ???)
488 [ >Hta >associative_append %
489 | (* da decidere se aggiungere un'assunzione o utilizzare Hmatch *) @daemon
492 #tc * whd in ⊢ (%→?); #Htc cases (Htc … Htable … Htb) -Htb -Htc
493 [| * #Hcurrent #Hfalse @False_ind
494 (* absurd by Hmatch *) @daemon
496 | (* da decidere se aggiungere un'assunzione o utilizzare Hmatch *) @daemon
497 | (* Htable *) @daemon
498 | (* Htable, Hmatch → |config| = n *) @daemon ]
499 * #table1 * #newc * #mv1 * #table2 * #Htableeq #Htc *
500 [ * #td * whd in ⊢ (%→?); >Htc -Htc #Htd
501 cases (Htd ? (refl ??)) #_ -Htd
502 cut (newc = 〈s1,false〉::newconfig@[〈c1,false〉]) [@daemon] #Hnewc
504 cut (mv1 = 〈\fst mv1,false〉)
505 [ >(eq_pair_fst_snd … mv1) @eq_f (*Htable, Htableeq*) @daemon ] #Hmv1
506 * #te * whd in ⊢ (%→?); #Hte
507 (* mv1 ha tipo lista ma dovrebbe avere tipo unialpha *)
508 cut (td = midtape STape (〈c0,false〉::reverse STape curconfig@〈s0,false〉::〈grid,false〉::ls)
510 ((table1@〈s0,false〉::curconfig@[〈c0,false〉])@〈comma,true〉::〈s1,false〉::
511 newconfig@〈c1,false〉::〈comma,false〉::〈\fst mv1,false〉::table2@〈grid,false〉::rs))
513 [ >reverse_append >reverse_single %
514 | >associative_append >associative_append normalize
515 >associative_append >associative_append <Hmv1 %
518 -Htd #Htd lapply (Hte … (S n) … Htd … Ht1') -Htd -Hte
520 |2,3,4,5: (* dovrebbe scomparire (lo metteremo nel legal_tape) *) @daemon
521 | (*|curconfig| = |newconfig|*) @daemon
522 | (* only_bits (〈s1,false〉::newconfig) *) @daemon
523 | (* only_bits (curconfig@[〈s0,false〉]) *) @daemon
524 | (* no_marks (reverse ? curconfig) *) @daemon
525 | <Hmv1 >Hnewc in Htableeq;
526 >associative_append >associative_append normalize
527 >associative_append >associative_append
528 #Htableeq <Htableeq // ]
529 * #ls1 * #rs1 * #c2 * * #Hte #Hliftte #Hlegalte
530 whd in ⊢ (%→?); #Houttape lapply (Houttape … Hte) -Houttape #Houttape
531 whd in Houttape:(???%); whd in Houttape:(???(??%%%));
532 @ex_intro [| @(ex_intro ?? rs1) @ex_intro [| % [ %
533 [ >Houttape @eq_f @eq_f @eq_f @eq_f
534 change with ((〈t0,false〉::table)@?) in ⊢ (???%);
535 >Htableeq >associative_append >associative_append
536 >associative_append normalize >associative_append
537 >associative_append normalize >Hnewc <Hmv1
538 >associative_append normalize >associative_append %
539 | >(?: mv = \fst mv1) [| (*Hmatch, Htableeq*) @daemon ]
546 | * #td * whd in ⊢ (%→%→?); >Htc #Htd
547 cases (Htd ? (refl ??)) normalize in ⊢ (%→?);
548 #Hfalse destruct (Hfalse)
551 | #t1 #t2 #t3 whd in ⊢ (%→%→?); #Ht1 #Ht2
552 #b #Hb cases (Ht1 ? Hb) #Hb' #Ht3 >Ht2 % //
553 cases b in Hb'; normalize #H1 //