]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/grammar/item.ma
ground_2 milestone: multiple relocation with lists of booleans
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / grammar / item.ma
index 323f22f882830ba1658eeea355aa1dfb8fc72ae3..cb69d3687b3476b60b390b8711f536a59763a442 100644 (file)
@@ -12,8 +12,8 @@
 (*                                                                        *)
 (**************************************************************************)
 
-include "ground_2/arith.ma".
-include "basic_2/notation.ma".
+include "ground_2/lib/bool.ma".
+include "ground_2/lib/arith.ma".
 
 (* ITEMS ********************************************************************)
 
@@ -42,18 +42,43 @@ inductive item2: Type[0] ≝
   | Flat2: flat2 → item2        (* non-binding item *)
 .
 
+(* Basic inversion lemmas ***************************************************)
+
+fact destruct_sort_sort_aux: ∀k1,k2. Sort k1 = Sort k2 → k1 = k2.
+#k1 #k2 #H destruct //
+qed-.
+
 (* Basic properties *********************************************************)
 
-axiom item0_eq_dec: ∀I1,I2:item0. Decidable (I1 = I2).
+lemma eq_item0_dec: ∀I1,I2:item0. Decidable (I1 = I2).
+* #i1 * #i2 [2,3,4,6,7,8: @or_intror #H destruct ]
+[2: elim (eq_nat_dec i1 i2) |1,3: elim (eq_nat_dec i1 i2) ] /2 width=1 by or_introl/
+#Hni12 @or_intror #H destruct /2 width=1 by/
+qed-.
 
 (* Basic_1: was: bind_dec *)
-axiom bind2_eq_dec: ∀I1,I2:bind2. Decidable (I1 = I2).
+lemma eq_bind2_dec: ∀I1,I2:bind2. Decidable (I1 = I2).
+* * /2 width=1 by or_introl/
+@or_intror #H destruct
+qed-.
 
 (* Basic_1: was: flat_dec *)
-axiom flat2_eq_dec: ∀I1,I2:flat2. Decidable (I1 = I2).
+lemma eq_flat2_dec: ∀I1,I2:flat2. Decidable (I1 = I2).
+* * /2 width=1 by or_introl/
+@or_intror #H destruct
+qed-.
 
 (* Basic_1: was: kind_dec *)
-axiom item2_eq_dec: ∀I1,I2:item2. Decidable (I1 = I2).
+lemma eq_item2_dec: ∀I1,I2:item2. Decidable (I1 = I2).
+* [ #a1 ] #I1 * [1,3: #a2 ] #I2
+[2,3: @or_intror #H destruct
+| elim (eq_bool_dec a1 a2) #Ha
+  [ elim (eq_bind2_dec I1 I2) /2 width=1 by or_introl/ #HI ]
+  @or_intror #H destruct /2 width=1 by/
+| elim (eq_flat2_dec I1 I2) /2 width=1 by or_introl/ #HI
+  @or_intror #H destruct /2 width=1 by/
+]
+qed-.
 
 (* Basic_1: removed theorems 21:
             s_S s_plus s_plus_sym s_minus minus_s_s s_le s_lt s_inj s_inc