]> matita.cs.unibo.it Git - helm.git/blob - matita/library/decidable_kit/list_aux.ma
...
[helm.git] / matita / library / decidable_kit / list_aux.ma
1 (**************************************************************************)
2 (*       ___                                                              *)
3 (*      ||M||                                                             *)
4 (*      ||A||       A project by Andrea Asperti                           *)
5 (*      ||T||                                                             *)
6 (*      ||I||       Developers:                                           *)
7 (*      ||T||         The HELM team.                                      *)
8 (*      ||A||         http://helm.cs.unibo.it                             *)
9 (*      \   /                                                             *)
10 (*       \ /        This file is distributed under the terms of the       *)
11 (*        v         GNU General Public License Version 2                  *)
12 (*                                                                        *)
13 (**************************************************************************)
14
15 set "baseuri" "cic:/matita/decidable_kit/list_aux/".
16
17 include "list/list.ma".
18 include "decidable_kit/eqtype.ma".
19 include "nat/plus.ma".
20
21 (* ### some functions on lists (some can be moved to list.ma ### *)
22
23 let rec foldr (A,B:Type) (f : A → B → B) (b : B) (l : list A) on l : B := 
24   match l with [ nil ⇒ b | (cons a l) ⇒ f a (foldr ? ? f b l)].
25    
26 definition length ≝ λT:Type.λl:list T.foldr T nat (λx,c.S c) O l.
27
28 definition count : ∀T : eqType.∀f : T → bool.∀l : list T. nat :=
29   λT:eqType.λf,l. 
30   foldr T nat (λx,acc. match (f x) with [ true ⇒ S acc | false ⇒ acc ]) O l.
31      
32 let rec mem (d : eqType) (x : d) (l : list d) on l : bool ≝
33   match l with
34   [ nil ⇒ false 
35   | cons y tl ⇒
36      match (cmp d x y) with [ true ⇒ true | false ⇒ mem d x tl]].
37
38 definition iota : nat → nat → list nat ≝
39   λn,m. nat_rect (λ_.list ?) (nil ?) (λx,acc.cons ? (n+x) acc) m.
40   
41 let rec map (A,B:Type) (f: A → B) (l : list A) on l : list B ≝
42   match l with [ nil ⇒ nil ? | cons x tl ⇒ f x :: (map A B f tl)].
43
44 (* ### induction principle for functions visiting 2 lists in parallel *)
45 lemma list_ind2 : 
46   ∀T1,T2:Type.∀l1:list T1.∀l2:list T2.∀P:list T1 → list T2 → Prop.
47   length ? l1 = length ? l2 →
48   (P (nil ?) (nil ?)) → 
49   (∀tl1,tl2,hd1,hd2. P tl1 tl2 → P (hd1::tl1) (hd2::tl2)) → 
50   P l1 l2.
51 intros (T1 T2 l1 l2 P Hl Pnil Pcons);
52 generalize in match Hl; clear Hl;
53 generalize in match l2; clear l2;
54 elim l1 1 (l2 x1); [ cases l2; intros (Hl); [assumption| destruct Hl]
55 | intros 3 (tl1 IH l2); cases l2; [1: simplify; intros 1 (Hl); destruct Hl] 
56   intros 1 (Hl); apply Pcons;
57   apply IH; destruct Hl; 
58   (* XXX simplify in Hl non folda length *) 
59   assumption;]
60 qed.
61   
62  
63 lemma eq_map : ∀A,B,f,g,l. (∀x.f x = g x) → map A B f l = map A B g l.
64 intros (A B f g l Efg); elim l; simplify; [1: reflexivity ];
65 rewrite > (Efg t); rewrite > H; reflexivity;  
66 qed.
67
68 (* ### eqtype for lists ### *)
69 let rec lcmp (d : eqType) (l1,l2 : list d) on l1 : bool ≝
70   match l1 with
71   [ nil ⇒ match l2 with [ nil ⇒ true | (cons _ _) ⇒ false]
72   | (cons x1 tl1) ⇒ match l2 with 
73       [ nil ⇒ false | (cons x2 tl2) ⇒ andb (cmp d x1 x2) (lcmp d tl1 tl2)]].
74            
75 lemma lcmp_length : 
76   ∀d:eqType.∀l1,l2:list d.
77  lcmp ? l1 l2 = true → length ? l1 = length ? l2.
78 intros 2 (d l1); elim l1 1 (l2 x1);
79 [ cases l2; simplify; intros;[reflexivity|destruct H] 
80 | intros 3 (tl1 IH l2); cases (l2); intros; [1:destruct H]
81   simplify; (* XXX la apply non fa simplify? *) 
82   apply congr_S; apply (IH l);
83   (* XXX qualcosa di enorme e' rotto! la regola di convertibilita?! *)
84   simplify in H;
85   letin H' ≝ (b2pT ? ? (andbP ? ?) H); decompose; assumption]
86 qed.
87
88 lemma lcmpP : ∀d:eqType.∀l1,l2:list d. eq_compatible (list d) l1 l2 (lcmp d l1 l2).
89 intros (d l1 l2);
90 generalize in match (refl_eq ? (lcmp d l1 l2));
91 generalize in match (lcmp d l1 l2) in ⊢ (? ? ? % → %); intros 1 (c);
92 cases c; intros (H); [ apply reflect_true | apply reflect_false ]
93 [ letin Hl≝ (lcmp_length ? ? ? H); clearbody Hl;
94   generalize in match H; clear H;
95   apply (list_ind2 ? ? ? ? ? Hl); [1: intros; reflexivity]
96   simplify; intros (tl1 tl2 hd1 hd2 IH H);
97   letin H' ≝ (b2pT ? ? (andbP ? ?) H); decompose;
98   rewrite > (IH H2); rewrite > (b2pT ? ? (eqP d ? ?) H1); reflexivity
99 | generalize in match H; clear H; generalize in match l2; clear l2;
100   elim l1 1 (l1 x1);
101    [ cases l1; [intros; destruct H]
102      unfold Not; intros; destruct H1;
103    | intros 3 (tl1 IH l2); cases l2;
104      [ unfold Not; intros; destruct H1;
105      | simplify;  intros;
106        letin H' ≝ (p2bT ? ? (negbP ?) H); clearbody H'; clear H;
107        letin H'' ≝ (b2pT ? ? (andbPF ? ?) H'); clearbody H''; clear H';
108        cases H''; clear H'';
109        [ intros; 
110          letin H' ≝ (b2pF ? ? (eqP d ? ?) H); clearbody H'; clear H;
111          (* XXX destruct H1; *)
112          change in H' with (Not ((match (x1::tl1) with [nil⇒x1|(cons x1 _) ⇒ x1]) = s));
113          rewrite > H1 in H'; simplify in H'; apply H'; reflexivity;
114        | intros; 
115          letin H' ≝ (IH ? H); clearbody H';
116          (* XXX destruct H1 *)
117          change in H' with (Not ((match (x1::tl1) with [nil⇒tl1|(cons _ l) ⇒ l]) = l)); 
118          rewrite > H1 in H'; simplify in H'; apply H'; reflexivity;
119        ]]]]
120 qed.
121     
122 definition list_eqType : eqType → eqType ≝ λd:eqType.mk_eqType ? ? (lcmpP d).  
123