]> matita.cs.unibo.it Git - helm.git/commitdiff
- lib: one lemma about equality was missing
authorFerruccio Guidi <ferruccio.guidi@unibo.it>
Fri, 4 Nov 2011 15:48:04 +0000 (15:48 +0000)
committerFerruccio Guidi <ferruccio.guidi@unibo.it>
Fri, 4 Nov 2011 15:48:04 +0000 (15:48 +0000)
- lambda_delta: one inversion lemma closed

matita/matita/contribs/lambda_delta/Basic_2/grammar/term.ma
matita/matita/contribs/lambda_delta/Basic_2/reducibility/tnf.ma
matita/matita/lib/basics/logic.ma

index ce8357d21bb0d283609706d73aeb79e751011630..f119e6de0a99876111ee8ab2ad42134fbb81d4f2 100644 (file)
@@ -41,7 +41,10 @@ interpretation "term flat construction (binary)" 'SFlat I T1 T2 = (TPair (Flat I
 lemma discr_tpair_xy_x: ∀I,T,V. 𝕔{I} V. T = V → False.
 #I #T #V elim V -V
 [ #J #H destruct
-| #J #W #U #IHW #_ #H destruct -I /2/ (**) (* improve context after destruct *)
+| #J #W #U #IHW #_ #H destruct
+(*
+ (generalize in match e1) -e1 >e0 normalize
+*) -I /2/ (**) (* destruct: one quality is not simplified, the destucted equality is not erased *)
 ]
 qed.
 
@@ -49,7 +52,7 @@ qed.
 lemma discr_tpair_xy_y: ∀I,V,T. 𝕔{I} V. T = T → False.
 #I #V #T elim T -T
 [ #J #H destruct
-| #J #W #U #_ #IHU #H destruct -I V /2/ (**) (* improve context after destruct *)
+| #J #W #U #_ #IHU #H destruct -I V /2/ (**) (* destruct: the destucted equality is not erased *)
 ]
 qed.
 
index 5d55e2ae6441e86298baa558051bb16ced41e6e6..5868a5c935777d578f0eae1fda58ab43fbb12439 100644 (file)
@@ -47,7 +47,9 @@ qed.
 
 axiom tnf_inv_abbr: ∀V,T. ℕ[𝕔{Abbr}V.T] → False.
 
-axiom tnf_inv_cast: ∀V,T. ℕ[𝕔{Cast}V.T] → False.
+lemma tnf_inv_cast: ∀V,T. ℕ[𝕔{Cast}V.T] → False.
+#V #T #H lapply (H T ?) -H /2/
+qed.
 
 (* Basic properties *********************************************************)
 
index 2fbe45f370cb71a589d8bff1fbe1029be2ef84bc..d1dbfad775aaebc0d1388255fb32890675c9ac49 100644 (file)
@@ -32,7 +32,12 @@ lemma eq_rect_Type0_r:
   ∀A.∀a.∀P: ∀x:A. eq ? x a → Type[0]. P a (refl A a) → ∀x.∀p:eq ? x a.P x p.
   #A #a #P #H #x #p (generalize in match H) (generalize in match P)
   cases p; //; qed.
-  
+
+lemma eq_rect_Type1_r:
+  ∀A.∀a.∀P: ∀x:A. eq ? x a → Type[1]. P a (refl A a) → ∀x.∀p:eq ? x a.P x p.
+  #A #a #P #H #x #p (generalize in match H) (generalize in match P)
+  cases p; //; qed.
+
 lemma eq_rect_Type2_r:
   ∀A.∀a.∀P: ∀x:A. eq ? x a → Type[2]. P a (refl A a) → ∀x.∀p:eq ? x a.P x p.
   #A #a #P #H #x #p (generalize in match H) (generalize in match P)