1 (**************************************************************************)
4 (* ||A|| A project by Andrea Asperti *)
6 (* ||I|| Developers: *)
7 (* ||T|| A.Asperti, C.Sacerdoti Coen, *)
8 (* ||A|| E.Tassi, S.Zacchiroli *)
10 (* \ / This file is distributed under the terms of the *)
11 (* v GNU Lesser General Public License Version 2.1 *)
13 (**************************************************************************)
15 set "baseuri" "cic:/matita/nat/permutation".
17 include "nat/compare.ma".
18 include "nat/sigma_and_pi.ma".
20 definition injn: (nat \to nat) \to nat \to Prop \def
21 \lambda f:nat \to nat.\lambda n:nat.\forall i,j:nat.
22 i \le n \to j \le n \to f i = f j \to i = j.
24 theorem injn_Sn_n: \forall f:nat \to nat. \forall n:nat.
25 injn f (S n) \to injn f n.unfold injn.
27 apply le_S.assumption.
28 apply le_S.assumption.
32 theorem injective_to_injn: \forall f:nat \to nat. \forall n:nat.
33 injective nat nat f \to injn f n.
34 unfold injective.unfold injn.intros.apply H.assumption.
37 definition permut : (nat \to nat) \to nat \to Prop
38 \def \lambda f:nat \to nat. \lambda m:nat.
39 (\forall i:nat. i \le m \to f i \le m )\land injn f m.
41 theorem permut_O_to_eq_O: \forall h:nat \to nat.
42 permut h O \to (h O) = O.
43 intros.unfold permut in H.
44 elim H.apply sym_eq.apply le_n_O_to_eq.
48 theorem permut_S_to_permut: \forall f:nat \to nat. \forall m:nat.
49 permut f (S m) \to f (S m) = (S m) \to permut f m.
53 cut (f i < S m \lor f i = S m).
55 apply le_S_S_to_le.assumption.
57 apply (not_le_Sn_n m).
59 rewrite > Hcut1.assumption.
60 apply H3.apply le_n.apply le_S.assumption.
61 rewrite > H5.assumption.
62 apply le_to_or_lt_eq.apply H2.apply le_S.assumption.
63 apply (injn_Sn_n f m H3).
68 definition transpose : nat \to nat \to nat \to nat \def
75 | false \Rightarrow n]].
78 right associative with precedence 71
79 for @{ 'transposition $i $j $n}.
81 notation < "(❲i \atop j❳)n"
82 right associative with precedence 71
83 for @{ 'transposition $i $j $n}.
85 interpretation "natural transposition" 'transposition i j n =
86 (cic:/matita/nat/permutation/transpose.con i j n).
88 lemma transpose_i_j_i: \forall i,j:nat. transpose i j i = j.
89 intros.unfold transpose.
90 rewrite > (eqb_n_n i).simplify. reflexivity.
93 lemma transpose_i_j_j: \forall i,j:nat. transpose i j j = i.
94 intros.unfold transpose.
95 apply (eqb_elim j i).simplify.intro.assumption.
96 rewrite > (eqb_n_n j).simplify.
100 theorem transpose_i_i: \forall i,n:nat. (transpose i i n) = n.
101 intros.unfold transpose.
102 apply (eqb_elim n i).
103 intro.simplify.apply sym_eq. assumption.
104 intro.simplify.reflexivity.
107 theorem transpose_i_j_j_i: \forall i,j,n:nat.
108 transpose i j n = transpose j i n.
109 intros.unfold transpose.
110 apply (eqb_elim n i).
111 apply (eqb_elim n j).
112 intros. simplify.rewrite < H. rewrite < H1.
114 intros.simplify.reflexivity.
115 apply (eqb_elim n j).
116 intros.simplify.reflexivity.
117 intros.simplify.reflexivity.
120 theorem transpose_transpose: \forall i,j,n:nat.
121 (transpose i j (transpose i j n)) = n.
122 intros.unfold transpose. unfold transpose.
123 apply (eqb_elim n i).simplify.
125 apply (eqb_elim j i).
126 simplify.intros.rewrite > H. rewrite > H1.reflexivity.
127 rewrite > (eqb_n_n j).simplify.intros.
130 apply (eqb_elim n j).simplify.
131 rewrite > (eqb_n_n i).intros.simplify.
132 apply sym_eq. assumption.
134 rewrite > (not_eq_to_eqb_false n i H1).
135 rewrite > (not_eq_to_eqb_false n j H).
136 simplify.reflexivity.
139 theorem injective_transpose : \forall i,j:nat.
140 injective nat nat (transpose i j).
143 rewrite < (transpose_transpose i j x).
144 rewrite < (transpose_transpose i j y).
145 apply eq_f.assumption.
148 variant inj_transpose: \forall i,j,n,m:nat.
149 transpose i j n = transpose i j m \to n = m \def
152 theorem permut_transpose: \forall i,j,n:nat. i \le n \to j \le n \to
153 permut (transpose i j) n.
154 unfold permut.intros.
155 split.unfold transpose.
157 elim (eqb i1 i).simplify.assumption.
158 elim (eqb i1 j).simplify.assumption.
160 apply (injective_to_injn (transpose i j) n).
161 apply injective_transpose.
164 theorem permut_fg: \forall f,g:nat \to nat. \forall n:nat.
165 permut f n \to permut g n \to permut (\lambda m.(f(g m))) n.
166 unfold permut. intros.
168 split.intros.simplify.apply H2.
171 apply H5.assumption.assumption.
172 apply H3.apply H4.assumption.apply H4.assumption.
176 theorem permut_transpose_l:
177 \forall f:nat \to nat. \forall m,i,j:nat.
178 i \le m \to j \le m \to permut f m \to permut (\lambda n.transpose i j (f n)) m.
179 intros.apply (permut_fg (transpose i j) f m ? ?).
180 apply permut_transpose.assumption.assumption.
184 theorem permut_transpose_r:
185 \forall f:nat \to nat. \forall m,i,j:nat.
186 i \le m \to j \le m \to permut f m \to permut (\lambda n.f (transpose i j n)) m.
187 intros.apply (permut_fg f (transpose i j) m ? ?).
188 assumption.apply permut_transpose.assumption.assumption.
191 theorem eq_transpose : \forall i,j,k,n:nat. \lnot j=i \to
192 \lnot i=k \to \lnot j=k \to
193 transpose i j n = transpose i k (transpose k j (transpose i k n)).
194 (* uffa: triplo unfold? *)
195 intros.unfold transpose.unfold transpose.unfold transpose.
196 apply (eqb_elim n i).intro.
197 simplify.rewrite > (eqb_n_n k).
198 simplify.rewrite > (not_eq_to_eqb_false j i H).
199 rewrite > (not_eq_to_eqb_false j k H2).
201 intro.apply (eqb_elim n j).
205 rewrite > (not_eq_to_eqb_false n k Hcut).
207 rewrite > (not_eq_to_eqb_false n k Hcut).
208 rewrite > (eq_to_eqb_true n j H4).
210 rewrite > (not_eq_to_eqb_false k i).
211 rewrite > (eqb_n_n k).
212 simplify.reflexivity.
213 unfold Not.intro.apply H1.apply sym_eq.assumption.
215 unfold Not.intro.apply H2.apply (trans_eq ? ? n).
216 apply sym_eq.assumption.assumption.
217 intro.apply (eqb_elim n k).intro.
219 rewrite > (not_eq_to_eqb_false i k H1).
220 rewrite > (not_eq_to_eqb_false i j).
222 rewrite > (eqb_n_n i).
224 unfold Not.intro.apply H.apply sym_eq.assumption.
226 rewrite > (not_eq_to_eqb_false n k H5).
227 rewrite > (not_eq_to_eqb_false n j H4).
229 rewrite > (not_eq_to_eqb_false n i H3).
230 rewrite > (not_eq_to_eqb_false n k H5).
231 simplify.reflexivity.
234 theorem permut_S_to_permut_transpose: \forall f:nat \to nat.
235 \forall m:nat. permut f (S m) \to permut (\lambda n.transpose (f (S m)) (S m)
237 unfold permut.intros.
239 split.intros.simplify.unfold transpose.
240 apply (eqb_elim (f i) (f (S m))).
241 intro.apply False_ind.
243 apply (not_le_Sn_n m).
244 rewrite < Hcut.assumption.
245 apply H2.apply le_S.assumption.apply le_n.assumption.
247 apply (eqb_elim (f i) (S m)).
249 cut (f (S m) \lt (S m) \lor f (S m) = (S m)).
250 elim Hcut.apply le_S_S_to_le.assumption.
251 apply False_ind.apply H4.rewrite > H6.assumption.
252 apply le_to_or_lt_eq.apply H1.apply le_n.
254 cut (f i \lt (S m) \lor f i = (S m)).
255 elim Hcut.apply le_S_S_to_le.assumption.
256 apply False_ind.apply H5.assumption.
257 apply le_to_or_lt_eq.apply H1.apply le_S.assumption.
259 apply H2.apply le_S.assumption.apply le_S.assumption.
260 apply (inj_transpose (f (S m)) (S m)).
264 (* bounded bijectivity *)
266 definition bijn : (nat \to nat) \to nat \to Prop \def
267 \lambda f:nat \to nat. \lambda n. \forall m:nat. m \le n \to
268 ex nat (\lambda p. p \le n \land f p = m).
270 theorem eq_to_bijn: \forall f,g:nat\to nat. \forall n:nat.
271 (\forall i:nat. i \le n \to (f i) = (g i)) \to
272 bijn f n \to bijn g n.
273 intros 4.unfold bijn.
275 apply (ex_intro ? ? a).
276 rewrite < (H a).assumption.
277 elim H3.assumption.assumption.
280 theorem bijn_Sn_n: \forall f:nat \to nat. \forall n:nat.
281 bijn f (S n) \to f (S n) = (S n) \to bijn f n.
282 unfold bijn.intros.elim (H m).
284 apply (ex_intro ? ? a).split.
285 cut (a < S n \lor a = S n).
286 elim Hcut.apply le_S_S_to_le.assumption.
288 apply (not_le_Sn_n n).
289 rewrite < H1.rewrite < H6.rewrite > H5.assumption.
290 apply le_to_or_lt_eq.assumption.assumption.
291 apply le_S.assumption.
294 theorem bijn_n_Sn: \forall f:nat \to nat. \forall n:nat.
295 bijn f n \to f (S n) = (S n) \to bijn f (S n).
297 cut (m < S n \lor m = S n).
301 apply (ex_intro ? ? a).split.
302 apply le_S.assumption.assumption.
303 apply le_S_S_to_le.assumption.
304 apply (ex_intro ? ? (S n)).
306 rewrite > H3.assumption.
307 apply le_to_or_lt_eq.assumption.
310 theorem bijn_fg: \forall f,g:nat\to nat. \forall n:nat.
311 bijn f n \to bijn g n \to bijn (\lambda p.f(g p)) n.
316 apply (ex_intro ? ? a1).
318 rewrite > H8.assumption.
319 assumption.assumption.
322 theorem bijn_transpose : \forall n,i,j. i \le n \to j \le n \to
323 bijn (transpose i j) n.
324 intros.unfold bijn.unfold transpose.intros.
325 cut (m = i \lor \lnot m = i).
327 apply (ex_intro ? ? j).
329 apply (eqb_elim j i).
330 intro.simplify.rewrite > H3.rewrite > H4.reflexivity.
331 rewrite > (eqb_n_n j).simplify.
332 intros. apply sym_eq.assumption.
333 cut (m = j \lor \lnot m = j).
335 apply (ex_intro ? ? i).
337 rewrite > (eqb_n_n i).simplify.
338 apply sym_eq. assumption.
339 apply (ex_intro ? ? m).
341 rewrite > (not_eq_to_eqb_false m i).
342 rewrite > (not_eq_to_eqb_false m j).
343 simplify. reflexivity.
346 apply (decidable_eq_nat m j).
347 apply (decidable_eq_nat m i).
350 theorem bijn_transpose_r: \forall f:nat\to nat.\forall n,i,j. i \le n \to j \le n \to
351 bijn f n \to bijn (\lambda p.f (transpose i j p)) n.
353 apply (bijn_fg f ?).assumption.
354 apply (bijn_transpose n i j).assumption.assumption.
357 theorem bijn_transpose_l: \forall f:nat\to nat.\forall n,i,j. i \le n \to j \le n \to
358 bijn f n \to bijn (\lambda p.transpose i j (f p)) n.
361 apply (bijn_transpose n i j).assumption.assumption.
365 theorem permut_to_bijn: \forall n:nat.\forall f:nat\to nat.
366 permut f n \to bijn f n.
368 elim n.unfold bijn.intros.
369 apply (ex_intro ? ? m).
371 apply (le_n_O_elim m ? (\lambda p. f p = p)).
372 assumption.unfold permut in H.
373 elim H.apply sym_eq. apply le_n_O_to_eq.apply H2.apply le_n.
374 apply (eq_to_bijn (\lambda p.
375 (transpose (f (S n1)) (S n1)) (transpose (f (S n1)) (S n1) (f p))) f).
376 intros.apply transpose_transpose.
377 apply (bijn_fg (transpose (f (S n1)) (S n1))).
378 apply bijn_transpose.
380 elim H1.apply H2.apply le_n.apply le_n.
383 apply permut_S_to_permut_transpose.
384 assumption.unfold transpose.
385 rewrite > (eqb_n_n (f (S n1))).simplify.reflexivity.
388 let rec invert_permut n f m \def
389 match eqb m (f n) with
394 |(S p) \Rightarrow invert_permut p f m]].
396 theorem invert_permut_f: \forall f:nat \to nat. \forall n,m:nat.
397 m \le n \to injn f n\to invert_permut n f (f m) = m.
402 rewrite > (eqb_n_n (f O)).simplify.reflexivity.
404 rewrite > (eqb_n_n (f (S m1))).simplify.reflexivity.
406 rewrite > (not_eq_to_eqb_false (f m) (f (S n1))).
408 apply injn_Sn_n. assumption.
409 unfold Not.intro.absurd (m = S n1).
410 apply H3.apply le_S.assumption.apply le_n.assumption.
412 apply (not_le_Sn_n n1).rewrite < H5.assumption.
415 theorem injective_invert_permut: \forall f:nat \to nat. \forall n:nat.
416 permut f n \to injn (invert_permut n f) n.
421 generalize in match (Hcut i H1).intro.
422 generalize in match (Hcut j H2).intro.
428 rewrite < (invert_permut_f f n a).
429 rewrite < (invert_permut_f f n a1).
432 assumption.assumption.
433 unfold permut in H.elim H. assumption.
435 unfold permut in H.elim H. assumption.
436 apply permut_to_bijn.assumption.
439 theorem permut_invert_permut: \forall f:nat \to nat. \forall n:nat.
440 permut f n \to permut (invert_permut n f) n.
441 intros.unfold permut.split.
442 intros.simplify.elim n.
443 simplify.elim (eqb i (f O)).simplify.apply le_n.simplify.apply le_n.
444 simplify.elim (eqb i (f (S n1))).simplify.apply le_n.
445 simplify.apply le_S. assumption.
446 apply injective_invert_permut.assumption.
449 theorem f_invert_permut: \forall f:nat \to nat. \forall n,m:nat.
450 m \le n \to permut f n\to f (invert_permut n f m) = m.
452 apply (injective_invert_permut f n H1).
453 unfold permut in H1.elim H1.
455 cut (permut (invert_permut n f) n).unfold permut in Hcut.
456 elim Hcut.apply H4.assumption.
457 apply permut_invert_permut.assumption.assumption.
458 apply invert_permut_f.
459 cut (permut (invert_permut n f) n).unfold permut in Hcut.
460 elim Hcut.apply H2.assumption.
461 apply permut_invert_permut.assumption.
462 unfold permut in H1.elim H1.assumption.
465 theorem permut_n_to_eq_n: \forall h:nat \to nat.\forall n:nat.
466 permut h n \to (\forall m:nat. m < n \to h m = m) \to h n = n.
467 intros.unfold permut in H.elim H.
468 cut (invert_permut n h n < n \lor invert_permut n h n = n).
470 rewrite < (f_invert_permut h n n) in \vdash (? ? ? %).
472 rewrite < (f_invert_permut h n n) in \vdash (? ? % ?).
473 apply H1.assumption.apply le_n.assumption.apply le_n.assumption.
474 rewrite < H4 in \vdash (? ? % ?).
475 apply (f_invert_permut h).apply le_n.assumption.
476 apply le_to_or_lt_eq.
477 cut (permut (invert_permut n h) n).
478 unfold permut in Hcut.elim Hcut.
480 apply permut_invert_permut.assumption.
483 theorem permut_n_to_le: \forall h:nat \to nat.\forall k,n:nat.
484 k \le n \to permut h n \to (\forall m:nat. m < k \to h m = m) \to
485 \forall j. k \le j \to j \le n \to k \le h j.
486 intros.unfold permut in H1.elim H1.
487 cut (h j < k \lor \not(h j < k)).
488 elim Hcut.absurd (k \le j).assumption.
490 cut (h j = j).rewrite < Hcut1.assumption.
491 apply H6.apply H5.assumption.assumption.
493 apply not_lt_to_le.assumption.
494 apply (decidable_lt (h j) k).
499 let rec map_iter_i k (g:nat \to nat) f (i:nat) \def
502 | (S k) \Rightarrow f (g (S (k+i))) (map_iter_i k g f i)].
504 theorem eq_map_iter_i: \forall g1,g2:nat \to nat.
505 \forall f:nat \to nat \to nat. \forall n,i:nat.
506 (\forall m:nat. i\le m \to m \le n+i \to g1 m = g2 m) \to
507 map_iter_i n g1 f i = map_iter_i n g2 f i.
508 intros 5.elim n.simplify.apply H.apply le_n.
509 apply le_n.simplify.apply eq_f2.apply H1.simplify.
510 apply le_S.apply le_plus_n.simplify.apply le_n.
511 apply H.intros.apply H1.assumption.simplify.apply le_S.assumption.
514 (* map_iter examples *)
516 theorem eq_map_iter_i_sigma: \forall g:nat \to nat. \forall n,m:nat.
517 map_iter_i n g plus m = sigma n g m.
518 intros.elim n.simplify.reflexivity.
520 apply eq_f.assumption.
523 theorem eq_map_iter_i_pi: \forall g:nat \to nat. \forall n,m:nat.
524 map_iter_i n g times m = pi n g m.
525 intros.elim n.simplify.reflexivity.
527 apply eq_f.assumption.
530 theorem eq_map_iter_i_fact: \forall n:nat.
531 map_iter_i n (\lambda m.m) times (S O) = (S n)!.
533 simplify.reflexivity.
535 (((S n1)+(S O))*(map_iter_i n1 (\lambda m.m) times (S O)) = (S(S n1))*(S n1)!).
536 rewrite < plus_n_Sm.rewrite < plus_n_O.
537 apply eq_f.assumption.
540 theorem eq_map_iter_i_transpose_l : \forall f:nat\to nat \to nat.associative nat f \to
541 symmetric2 nat nat f \to \forall g:nat \to nat. \forall n,k:nat.
542 map_iter_i (S k) g f n = map_iter_i (S k) (\lambda m. g (transpose (k+n) (S k+n) m)) f n.
543 intros.apply (nat_case1 k).
544 intros.simplify.fold simplify (transpose n (S n) (S n)).
545 rewrite > transpose_i_j_i.
546 rewrite > transpose_i_j_j.
550 (f (g (S (S (m+n)))) (f (g (S (m+n))) (map_iter_i m g f n)) =
551 f (g (transpose (S m + n) (S (S m) + n) (S (S m)+n)))
552 (f (g (transpose (S m + n) (S (S m) + n) (S m+n)))
553 (map_iter_i m (\lambda m1. g (transpose (S m+n) (S (S m)+n) m1)) f n))).
554 rewrite > transpose_i_j_i.
555 rewrite > transpose_i_j_j.
558 rewrite < (H1 (g (S m + n))).
561 intros.simplify.unfold transpose.
562 rewrite > (not_eq_to_eqb_false m1 (S m+n)).
563 rewrite > (not_eq_to_eqb_false m1 (S (S m)+n)).
566 apply (lt_to_not_eq m1 (S ((S m)+n))).
567 unfold lt.apply le_S_S.change with (m1 \leq S (m+n)).apply le_S.assumption.
568 apply (lt_to_not_eq m1 (S m+n)).
569 simplify.unfold lt.apply le_S_S.assumption.
572 theorem eq_map_iter_i_transpose_i_Si : \forall f:nat\to nat \to nat.associative nat f \to
573 symmetric2 nat nat f \to \forall g:nat \to nat. \forall n,k,i:nat. n \le i \to i \le k+n \to
574 map_iter_i (S k) g f n = map_iter_i (S k) (\lambda m. g (transpose i (S i) m)) f n.
575 intros 6.elim k.cut (i=n).
577 apply (eq_map_iter_i_transpose_l f H H1 g n O).
578 apply antisymmetric_le.assumption.assumption.
579 cut (i < S n1 + n \lor i = S n1 + n).
582 (f (g (S (S n1)+n)) (map_iter_i (S n1) g f n) =
583 f (g (transpose i (S i) (S (S n1)+n))) (map_iter_i (S n1) (\lambda m. g (transpose i (S i) m)) f n)).
584 apply eq_f2.unfold transpose.
585 rewrite > (not_eq_to_eqb_false (S (S n1)+n) i).
586 rewrite > (not_eq_to_eqb_false (S (S n1)+n) (S i)).
587 simplify.reflexivity.
588 simplify.unfold Not.intro.
589 apply (lt_to_not_eq i (S n1+n)).assumption.
590 apply inj_S.apply sym_eq. assumption.
591 simplify.unfold Not.intro.
592 apply (lt_to_not_eq i (S (S n1+n))).simplify.unfold lt.
593 apply le_S_S.assumption.
594 apply sym_eq. assumption.
595 apply H2.assumption.apply le_S_S_to_le.
598 apply (eq_map_iter_i_transpose_l f H H1 g n (S n1)).
599 apply le_to_or_lt_eq.assumption.
602 theorem eq_map_iter_i_transpose:
603 \forall f:nat\to nat \to nat.
604 associative nat f \to symmetric2 nat nat f \to \forall n,k,o:nat.
605 \forall g:nat \to nat. \forall i:nat. n \le i \to S (o + i) \le S k+n \to
606 map_iter_i (S k) g f n = map_iter_i (S k) (\lambda m. g (transpose i (S(o + i)) m)) f n.
610 apply (nat_case m ?).
612 apply (eq_map_iter_i_transpose_i_Si ? H H1).
613 exact H3.apply le_S_S_to_le.assumption.
615 apply (trans_eq ? ? (map_iter_i (S k) (\lambda m. g (transpose i (S(m1 + i)) m)) f n)).
617 unfold lt. apply le_n.assumption.
618 apply (trans_le ? (S(S (m1+i)))).
619 apply le_S.apply le_n.assumption.
620 apply (trans_eq ? ? (map_iter_i (S k) (\lambda m. g
621 (transpose i (S(m1 + i)) (transpose (S(m1 + i)) (S(S(m1 + i))) m))) f n)).
622 apply (H2 O ? ? (S(m1+i))).
623 unfold lt.apply le_S_S.apply le_O_n.
624 apply (trans_le ? i).assumption.
625 change with (i \le (S m1)+i).apply le_plus_n.
627 apply (trans_eq ? ? (map_iter_i (S k) (\lambda m. g
628 (transpose i (S(m1 + i))
629 (transpose (S(m1 + i)) (S(S(m1 + i)))
630 (transpose i (S(m1 + i)) m)))) f n)).
632 unfold lt. apply le_n.assumption.
633 apply (trans_le ? (S(S (m1+i)))).
634 apply le_S.apply le_n.assumption.
637 apply sym_eq. apply eq_transpose.
639 apply (not_le_Sn_n i).
640 rewrite < H7 in \vdash (? ? %).
641 apply le_S_S.apply le_S.
644 apply (not_le_Sn_n i).
645 rewrite > H7 in \vdash (? ? %).
649 apply (not_eq_n_Sn (S m1+i)).
650 apply sym_eq.assumption.
653 theorem eq_map_iter_i_transpose1: \forall f:nat\to nat \to nat.associative nat f \to
654 symmetric2 nat nat f \to \forall n,k,i,j:nat.
655 \forall g:nat \to nat. n \le i \to i < j \to j \le S k+n \to
656 map_iter_i (S k) g f n = map_iter_i (S k) (\lambda m. g (transpose i j m)) f n.
659 cut ((S i) < j \lor (S i) = j).
661 cut (j = S ((j - (S i)) + i)).
663 apply (eq_map_iter_i_transpose f H H1 n k (j - (S i)) g i).
665 rewrite < Hcut1.assumption.
667 apply plus_minus_m_m.apply lt_to_le.assumption.
669 apply (eq_map_iter_i_transpose_i_Si f H H1 g).
671 assumption.apply le_S_S_to_le.
672 apply (trans_le ? j).assumption.assumption.
673 apply le_to_or_lt_eq.assumption.
676 theorem eq_map_iter_i_transpose2: \forall f:nat\to nat \to nat.associative nat f \to
677 symmetric2 nat nat f \to \forall n,k,i,j:nat.
678 \forall g:nat \to nat. n \le i \to i \le (S k+n) \to n \le j \to j \le (S k+n) \to
679 map_iter_i (S k) g f n = map_iter_i (S k) (\lambda m. g (transpose i j m)) f n.
681 apply (nat_compare_elim i j).
682 intro.apply (eq_map_iter_i_transpose1 f H H1 n k i j g H2 H6 H5).
684 apply eq_map_iter_i.intros.
685 rewrite > (transpose_i_i j).reflexivity.
687 apply (trans_eq ? ? (map_iter_i (S k) (\lambda m:nat.g (transpose j i m)) f n)).
688 apply (eq_map_iter_i_transpose1 f H H1 n k j i g H4 H6 H3).
690 intros.apply eq_f.apply transpose_i_j_j_i.
693 theorem permut_to_eq_map_iter_i:\forall f:nat\to nat \to nat.associative nat f \to
694 symmetric2 nat nat f \to \forall k,n:nat.\forall g,h:nat \to nat.
695 permut h (k+n) \to (\forall m:nat. m \lt n \to h m = m) \to
696 map_iter_i k g f n = map_iter_i k (\lambda m.g(h m)) f n.
698 simplify.rewrite > (permut_n_to_eq_n h).reflexivity.assumption.assumption.
699 apply (trans_eq ? ? (map_iter_i (S n) (\lambda m.g ((transpose (h (S n+n1)) (S n+n1)) m)) f n1)).
702 apply (eq_map_iter_i_transpose2 f H H1 n1 n ? ? g).
703 apply (permut_n_to_le h n1 (S n+n1)).
704 apply le_plus_n.assumption.assumption.apply le_plus_n.apply le_n.
705 apply H5.apply le_n.apply le_plus_n.apply le_n.
706 apply (trans_eq ? ? (map_iter_i (S n) (\lambda m.
707 (g(transpose (h (S n+n1)) (S n+n1)
708 (transpose (h (S n+n1)) (S n+n1) (h m)))) )f n1)).
709 simplify.fold simplify (transpose (h (S n+n1)) (S n+n1) (S n+n1)).
710 apply eq_f2.apply eq_f.
711 rewrite > transpose_i_j_j.
712 rewrite > transpose_i_j_i.
713 rewrite > transpose_i_j_j.reflexivity.
714 apply (H2 n1 (\lambda m.(g(transpose (h (S n+n1)) (S n+n1) m)))).
715 apply permut_S_to_permut_transpose.
719 rewrite > (not_eq_to_eqb_false (h m) (h (S n+n1))).
720 rewrite > (not_eq_to_eqb_false (h m) (S n+n1)).
721 simplify.apply H4.assumption.
723 apply lt_to_not_eq.apply (trans_lt ? n1).assumption.
724 simplify.unfold lt.apply le_S_S.apply le_plus_n.assumption.
725 unfold permut in H3.elim H3.
726 simplify.unfold Not.intro.
727 apply (lt_to_not_eq m (S n+n1)).apply (trans_lt ? n1).assumption.
728 simplify.unfold lt.apply le_S_S.apply le_plus_n.
730 apply (H7 m (S n+n1)).apply (trans_le ? n1).
731 apply lt_to_le.assumption.apply le_plus_n.apply le_n.
733 apply eq_map_iter_i.intros.
734 rewrite > transpose_transpose.reflexivity.