]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambda_delta/Basic_2/grammar/aarity.ma
- the definition of the framework for strong normalization continues ...
[helm.git] / matita / matita / contribs / lambda_delta / Basic_2 / grammar / aarity.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 include "Ground_2/list.ma".
16 include "Basic_2/notation.ma".
17
18 (* ATOMIC ARITY *************************************************************)
19
20 inductive aarity: Type[0] ≝
21   | AAtom: aarity                   (* atomic aarity construction *)
22   | APair: aarity → aarity → aarity (* binary aarity construction *)
23 .
24
25 interpretation "aarity construction (atomic)" 'SItem = AAtom.
26
27 interpretation "aarity construction (binary)" 'SItem A1 A2 = (APair A1 A2).
28
29 (* Basic inversion lemmas ***************************************************)
30
31 lemma discr_apair_xy_x: ∀A,B. 𝕔 B. A = B → False.
32 #A #B elim B -B
33 [ #H destruct
34 | #Y #X #IHY #_ #H destruct
35   -H >e0 in e1; normalize (**) (* destruct: one quality is not simplified, the destucted equality is not erased *)
36   /2 width=1/
37 ]
38 qed-.
39
40 lemma discr_tpair_xy_y: ∀B,A. 𝕔 B. A = A → False.
41 #B #A elim A -A
42 [ #H destruct
43 | #Y #X #_ #IHX #H destruct
44   -H (**) (* destruct: the destucted equality is not erased *)
45   /2 width=1/
46 ]
47 qed-.
48
49 (* Basic properties *********************************************************)
50
51 lemma aarity_eq_dec: ∀A1,A2:aarity. Decidable (A1 = A2).
52 #A1 elim A1 -A1
53 [ #A2 elim A2 -A2 /2 width=1/
54   #B2 #A2 #_ #_ @or_intror #H destruct
55 | #B1 #A1 #IHB1 #IHA1 #A2 elim A2 -A2
56   [ -IHB1 -IHA1 @or_intror #H destruct
57   | #B2 #A2 #_ #_ elim (IHB1 B2) -IHB1
58     [ #H destruct elim (IHA1 A2) -IHA1
59       [ #H destruct /2 width=1/
60       | #HA12 @or_intror #H destruct /2 width=1/
61       ]
62     | -IHA1 #HB12 @or_intror #H destruct /2 width=1/
63     ]
64   ]
65 ]
66 qed-.