]> matita.cs.unibo.it Git - helm.git/commitdiff
commit by user andrea
authormatitaweb <claudio.sacerdoticoen@unibo.it>
Tue, 6 Mar 2012 18:09:27 +0000 (18:09 +0000)
committermatitaweb <claudio.sacerdoticoen@unibo.it>
Tue, 6 Mar 2012 18:09:27 +0000 (18:09 +0000)
weblib/basics/list.ma

index e699b626a4f2f02dab5c38a939871afab0a988c4..39fdb1bd5ca9f550d6f3b1d37e72395af2c8546d 100644 (file)
@@ -9,7 +9,7 @@
      \ /   GNU General Public License Version 2   
       V_______________________________________________________________ *)
 
-include "arithmetics/nat.ma".
+include "arithmetics/nat.ma".\ 5span class="error" title="disambiguation error"\ 6\ 5/span\ 6
 
 inductive list (A:Type[0]) : Type[0] :=
   | nil: list A
@@ -124,6 +124,46 @@ match l1 with
   | cons a tl ⇒ (map ??(dp ?? a) (g a)) @ dprodl A f tl g
   ]. *)
 
+(**************************** reverse *****************************)
+let rec rev_append S (l1,l2:\ 5a href="cic:/matita/basics/list/list.ind(1,0,1)"\ 6list\ 5/a\ 6 S) on l1 ≝
+  match l1 with 
+  [ nil ⇒ l2
+  | cons a tl ⇒ rev_append S tl (a\ 5a title="cons" href="cic:/fakeuri.def(1)"\ 6:\ 5/a\ 6:l2)
+  ]
+.
+
+definition reverse ≝λS.λl.\ 5a href="cic:/matita/basics/list/rev_append.fix(0,1,1)"\ 6rev_append\ 5/a\ 6 S l \ 5a title="nil" href="cic:/fakeuri.def(1)"\ 6[\ 5/a\ 6].
+
+lemma reverse_single : ∀S,a. \ 5a href="cic:/matita/basics/list/reverse.def(2)"\ 6reverse\ 5/a\ 6 S \ 5a title="cons" href="cic:/fakeuri.def(1)"\ 6[\ 5/a\ 6a] \ 5a title="leibnitz's equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6 \ 5a title="cons" href="cic:/fakeuri.def(1)"\ 6[\ 5/a\ 6a]. 
+// qed.
+
+lemma rev_append_def : ∀S,l1,l2. 
+  \ 5a href="cic:/matita/basics/list/rev_append.fix(0,1,1)"\ 6rev_append\ 5/a\ 6 S l1 l2 \ 5a title="leibnitz's equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6 (\ 5a href="cic:/matita/basics/list/reverse.def(2)"\ 6reverse\ 5/a\ 6 S l1) \ 5a title="append" href="cic:/fakeuri.def(1)"\ 6@\ 5/a\ 6 l2 .
+#S #l1 elim l1 normalize // 
+qed.
+
+lemma reverse_cons : ∀S,a,l. \ 5a href="cic:/matita/basics/list/reverse.def(2)"\ 6reverse\ 5/a\ 6 S (a\ 5a title="cons" href="cic:/fakeuri.def(1)"\ 6:\ 5/a\ 6:l) \ 5a title="leibnitz's equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6 (\ 5a href="cic:/matita/basics/list/reverse.def(2)"\ 6reverse\ 5/a\ 6 S l)\ 5a title="append" href="cic:/fakeuri.def(1)"\ 6@\ 5/a\ 6\ 5a title="cons" href="cic:/fakeuri.def(1)"\ 6[\ 5/a\ 6a].
+#S #a #l whd in ⊢ (??%?); // 
+qed.
+
+lemma reverse_append: ∀S,l1,l2. 
+  \ 5a href="cic:/matita/basics/list/reverse.def(2)"\ 6reverse\ 5/a\ 6 S (l1 \ 5a title="append" href="cic:/fakeuri.def(1)"\ 6@\ 5/a\ 6 l2) \ 5a title="leibnitz's equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6 (\ 5a href="cic:/matita/basics/list/reverse.def(2)"\ 6reverse\ 5/a\ 6 S l2)\ 5a title="append" href="cic:/fakeuri.def(1)"\ 6@\ 5/a\ 6(\ 5a href="cic:/matita/basics/list/reverse.def(2)"\ 6reverse\ 5/a\ 6 S l1).
+#S #l1 elim l1 [normalize // | #a #tl #Hind #l2 >\ 5a href="cic:/matita/basics/list/reverse_cons.def(7)"\ 6reverse_cons\ 5/a\ 6
+>\ 5a href="cic:/matita/basics/list/reverse_cons.def(7)"\ 6reverse_cons\ 5/a\ 6 // qed.
+
+lemma reverse_reverse : ∀S,l. \ 5a href="cic:/matita/basics/list/reverse.def(2)"\ 6reverse\ 5/a\ 6 S (\ 5a href="cic:/matita/basics/list/reverse.def(2)"\ 6reverse\ 5/a\ 6 S l) \ 5a title="leibnitz's equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6 l.
+#S #l elim l // #a #tl #Hind >\ 5a href="cic:/matita/basics/list/reverse_cons.def(7)"\ 6reverse_cons\ 5/a\ 6 >\ 5a href="cic:/matita/basics/list/reverse_append.def(8)"\ 6reverse_append\ 5/a\ 6 
+normalize // qed.
+
+(* an elimination principle for lists working on the tail;
+useful for strings *)
+lemma list_elim_left: ∀S.∀P:\ 5a href="cic:/matita/basics/list/list.ind(1,0,1)"\ 6list\ 5/a\ 6 S → Prop. P (\ 5a href="cic:/matita/basics/list/list.con(0,1,1)"\ 6nil\ 5/a\ 6 S) →
+(∀a.∀tl.P tl → P (tl\ 5a title="append" href="cic:/fakeuri.def(1)"\ 6@\ 5/a\ 6\ 5a title="cons" href="cic:/fakeuri.def(1)"\ 6[\ 5/a\ 6a])) → ∀l. P l.
+#S #P #Pnil #Pstep #l <(\ 5a href="cic:/matita/basics/list/reverse_reverse.def(9)"\ 6reverse_reverse\ 5/a\ 6 … l) 
+generalize in match (\ 5a href="cic:/matita/basics/list/reverse.def(2)"\ 6reverse\ 5/a\ 6 S l); #l elim l //
+#a #tl #H >\ 5a href="cic:/matita/basics/list/reverse_cons.def(7)"\ 6reverse_cons\ 5/a\ 6 @Pstep //
+qed.
+
 (**************************** length *******************************)
 
 let rec length (A:Type[0]) (l:\ 5a href="cic:/matita/basics/list/list.ind(1,0,1)"\ 6list\ 5/a\ 6 A) on l ≝