1 (**************************************************************************)
4 (* ||A|| A project by Andrea Asperti *)
6 (* ||I|| Developers: *)
7 (* ||T|| The HELM team. *)
8 (* ||A|| http://helm.cs.unibo.it *)
10 (* \ / This file is distributed under the terms of the *)
11 (* v GNU General Public License Version 2 *)
13 (**************************************************************************)
15 (* THE FORMAL SYSTEM λδ: MATITA SOURCE FILES
16 * Initial invocation: - Patience on me to gain peace and perfection! -
19 include "ground/lib/relations.ma".
20 include "static_2/notation/functions/item0_0.ma".
21 include "static_2/notation/functions/snitem2_2.ma".
23 (* ATOMIC ARITY *************************************************************)
25 inductive aarity: Type[0] ≝
26 | AAtom: aarity (* atomic aarity construction *)
27 | APair: aarity → aarity → aarity (* binary aarity construction *)
30 interpretation "atomic arity construction (atomic)"
33 interpretation "atomic arity construction (binary)"
34 'SnItem2 A1 A2 = (APair A1 A2).
36 (* Basic inversion lemmas ***************************************************)
38 fact destruct_apair_apair_aux: ∀A1,A2,B1,B2. ②B1.A1 = ②B2.A2 → B1 = B2 ∧ A1 = A2.
39 #A1 #A2 #B1 #B2 #H destruct /2 width=1 by conj/
42 lemma discr_apair_xy_x: ∀A,B. ②B.A = B → ⊥.
45 | #Y #X #IHY #_ #H elim (destruct_apair_apair_aux … H) -H /2 width=1 by/ (**) (* destruct lemma needed *)
49 (* Basic_2A1: was: discr_tpair_xy_y *)
50 lemma discr_apair_xy_y: ∀B,A. ②B. A = A → ⊥.
53 | #Y #X #_ #IHX #H elim (destruct_apair_apair_aux … H) -H /2 width=1 by/ (**) (* destruct lemma needed *)
57 (* Basic properties *********************************************************)
59 lemma eq_aarity_dec: ∀A1,A2:aarity. Decidable (A1 = A2).
61 [ #A2 elim A2 -A2 /2 width=1 by or_introl/
62 #B2 #A2 #_ #_ @or_intror #H destruct
63 | #B1 #A1 #IHB1 #IHA1 #A2 elim A2 -A2
64 [ -IHB1 -IHA1 @or_intror #H destruct
65 | #B2 #A2 #_ #_ elim (IHB1 B2) -IHB1
66 [ #H destruct elim (IHA1 A2) -IHA1
67 [ #H destruct /2 width=1 by or_introl/
68 | #HA12 @or_intror #H destruct /2 width=1 by/
70 | -IHA1 #HB12 @or_intror #H destruct /2 width=1 by/
76 lemma is_apear_dec (B) (X): Decidable (∃A. ②B.A = X).
78 [| elim (eq_aarity_dec X B) #HX ]
79 [| /3 width=2 by ex_intro, or_introl/ ]
80 @or_intror * #A #H destruct