]> matita.cs.unibo.it Git - helm.git/blob - matita/tests/coercions_propagation.ma
fixed propagation under Fix/Lambda/Case of coercions, better names are
[helm.git] / matita / tests / coercions_propagation.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/test/coercions_propagation/".
16
17 include "logic/connectives.ma".
18 include "nat/orders.ma".
19 alias num (instance 0) = "natural number".
20
21 inductive sigma (A:Type) (P:A → Prop) : Type ≝
22  sigma_intro: ∀a:A. P a → sigma A P.
23
24 interpretation "sigma" 'exists \eta.x =
25   (cic:/matita/test/coercions_propagation/sigma.ind#xpointer(1/1) _ x).
26   
27 definition inject ≝ λP.λa:nat.λp:P a. sigma_intro ? P ? p.
28
29 coercion cic:/matita/test/coercions_propagation/inject.con 0 1.
30
31 definition eject ≝ λP.λc: ∃n:nat.P n. match c with [ sigma_intro w _ ⇒ w].
32
33 coercion cic:/matita/test/coercions_propagation/eject.con.
34
35 alias num (instance 0) = "natural number".
36
37 theorem test: ∃n. 0 ≤ n.
38  apply (S O : ∃n. 0 ≤ n).
39  autobatch.
40 qed.
41
42 theorem test2: nat → ∃n. 0 ≤ n.
43  apply ((λn:nat. 0) : nat → ∃n. 0 ≤ n);
44  autobatch.
45 qed.
46
47 theorem test3: (∃n. 0 ≤ n) → nat.
48  apply ((λn:nat.n) : (∃n. 0 ≤ n) → nat).
49 qed.
50
51 theorem test4: (∃n. 1 ≤ n) → ∃n. 0 < n.
52  apply ((λn:nat.n) : (∃n. 1 ≤ n) → ∃n. 0 < n);
53  cases s;
54  assumption.
55 qed.
56
57 theorem test5: nat → ∃n. 1 ≤ n.
58 apply (
59  let rec aux n : nat ≝
60   match n with
61    [ O ⇒ 1
62    | S m ⇒ S (aux m)
63    ]
64  in
65   aux
66 : nat → ∃n. 1 ≤ n);
67 [ cases (aux n1); simplify; ] autobatch;
68 qed.
69
70 inductive NN (A:Type) : nat -> Type ≝
71  | NO : NN A O
72  | NS : ∀n:nat. NN A n → NN A (S n).
73  
74 definition injectN ≝ λA,k.λP.λa:NN A k.λp:P a. sigma_intro ? P ? p.
75
76 coercion cic:/matita/test/coercions_propagation/injectN.con 0 1.
77
78 definition ejectN ≝ λA,k.λP.λc: ∃n:NN A k.P n. match c with [ sigma_intro w _ ⇒ w].
79
80 coercion cic:/matita/test/coercions_propagation/ejectN.con.
81
82 definition PN :=
83  λA,k.λx:NN A k. 1 <= k.
84
85 theorem test51_byhand: ∀A,k. NN A k → ∃n:NN A (S k). PN ? ? n.
86 intros 1;
87 apply (
88  let rec aux (w : nat) (n : NN A w) on n : ∃p:NN A (S w).PN ? ? p ≝
89   match n in NN return λx.λm:NN A x.∃p:NN A (S x).PN ? ? p with
90    [ NO ⇒ injectN ? ? ? (NS A ? (NO A)) ?
91    | NS x m ⇒ injectN ? ? ? (NS A (S x) (ejectN ? ? ? (aux ? m))) ?  
92    ]
93  in
94   aux
95 : ∀n:nat. NN A n → ∃m:NN A (S n). PN ? ? m);
96 [2: cases (aux x m); simplify; autobatch ] unfold PN; autobatch;
97 qed.
98
99 theorem f : nat -> nat -> ∃n:nat.O <= n.
100 apply (λx,y:nat.y : nat -> nat -> ∃n:nat. O <= n).
101 apply le_O_n; 
102 qed.
103
104 axiom F : nat -> nat -> nat.
105
106 theorem f1 : nat -> nat -> ∃n:nat.O <= n.
107 apply (F : nat -> nat -> ∃n:nat. O <= n).
108 apply le_O_n; 
109 qed.
110
111 theorem test51: ∀A,k. NN A k → ∃n:NN A (S k). PN ? ? n.
112 intros 1;
113 letin xxx ≝ (
114  let rec aux (w : nat) (n : NN A w) on n : NN A (S w) ≝
115   match n in NN return λx.λm:NN A x.NN A (S x) with
116    [ NO ⇒ NS A ? (NO A)
117    | NS x m ⇒ NS A (S x) (aux ? m)  
118    ]
119  in
120   aux
121 : ∀n:nat. NN A n → ∃m:NN A (S n). PN ? ? m); [3: apply xxx];
122 unfold PN in aux ⊢ %; [cases (aux n2 n3)] autobatch
123 qed.
124
125 (* guarded troppo debole 
126 theorem test5: nat → ∃n. 1 ≤ n.
127 apply (
128  let rec aux n : nat ≝
129   match n with
130    [ O ⇒ 1
131    | S m ⇒ S (aux m)
132    ]
133  in
134   aux
135 : nat → ∃n. 1 ≤ n);
136 cases name_con;
137 simplify;
138  [ autobatch
139  | cases (aux n);
140    simplify;
141    apply lt_O_S
142  ]
143 qed.
144 *)
145
146 (*
147 theorem test5: nat → ∃n. 1 ≤ n.
148 apply (
149  let rec aux (n : nat) : ∃n. 1 ≤ n ≝ 
150    match n with
151    [ O => (S O : ∃n. 1 ≤ n)
152    | S m => (S (aux m) : ∃n. 1 ≤ n)
153 (*      
154    inject ? (S (eject ? (aux m))) ? *)
155    ]
156  in
157   aux
158  : nat → ∃n. 1 ≤ n);
159  [ autobatch
160  | elim (aux m);
161    simplify;
162    autobatch
163  ]
164 qed.
165
166 Re1: nat => nat |- body[Rel1] : nat => nat
167 Re1: nat => X |- body[Rel1] : nat => nat : nat => X
168 Re1: Y => X |- body[Rel1] : nat => nat : Y => X
169
170 nat => nat
171 nat => X
172
173 theorem test5: (∃n. 2 ≤ n) → ∃n. 1 ≤ n.
174 apply (
175  λk: ∃n. 2 ≤ n.
176  let rec aux n : eject ? n = eject ? k → ∃n. 1 ≤ n ≝
177   match eject ? n return λx:nat. x = eject ? k → ∃n. 1 ≤ n with
178    [ O ⇒ λH: 0 = eject ? k.
179           sigma_intro ? ? 0 ?
180    | S m ⇒ λH: S m = eject ? k.
181           sigma_intro ? ? (S m) ?
182    ]
183  in
184   aux k (refl_eq ? (eject ? k)));
185
186
187 intro;
188 cases s; clear s;
189 generalize in match H; clear H;
190 elim a;
191  [ apply (sigma_intro ? ? 0);
192  | apply (sigma_intro ? ? (S n));
193  ].
194
195 apply (
196  λk.
197  let rec aux n : ∃n. 1 ≤ n ≝
198   inject ?
199    (match n with
200      [ O ⇒ O
201      | S m ⇒ S (eject ? (aux m))
202      ]) ?
203  in aux (eject ? k)).
204
205    
206 apply (
207  let rec aux n : nat ≝
208   match n with
209    [ O ⇒ O
210    | S m ⇒ S (aux m)
211    ]
212  in
213   aux
214 : (∃n. 2 ≤ n) → ∃n. 1 ≤ n);
215
216 qed.
217
218 (*
219 theorem test5: nat → ∃n. 0 ≤ n.
220  apply (λn:nat.?);
221  apply
222   (match n return λ_.∃n.0 ≤ n with [O ⇒ (0 : ∃n.0 ≤ n) | S n' ⇒ ex_intro ? ? n' ?]
223   : ∃n. 0 ≤ n);
224  autobatch.
225 qed.
226 *)
227 *)