]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambda_delta/Basic_2/reducibility/twhnf.ma
component "reducibility" updated to new syntax!
[helm.git] / matita / matita / contribs / lambda_delta / Basic_2 / reducibility / twhnf.ma
index ccfaa6aa69c8b0e106ea2b7ed1501023899850d3..fe6b66447ae96eac4bf52faa30219a704d5902b3 100644 (file)
@@ -26,22 +26,22 @@ interpretation
 (* Basic inversion lemmas ***************************************************)
 
 lemma twhnf_inv_thom: ∀T. 𝕎ℍℕ[T] → T ≈ T.
-normalize /2 depth=1/
+normalize /2 width=1/
 qed-.
 
 (* Basic properties *********************************************************)
 
 lemma tpr_thom: ∀T1,T2. T1 ⇒ T2 → T1 ≈ T1 → T1 ≈ T2.
-#T1 #T2 #H elim H -T1 T2 //
+#T1 #T2 #H elim H -T1 -T2 //
 [ #I #V1 #V2 #T1 #T2 #_ #_ #_ #IHT12 #H
-  elim (thom_inv_flat1 … H) -H #W2 #U2 #HT1U2 #HT1 #_ #H1 #H2 destruct -I T1 V1;
-  lapply (IHT12 HT1U2) -IHT12 HT1U2 #HUT2
+  elim (thom_inv_flat1 … H) -H #W2 #U2 #HT1U2 #HT1 #_ #H1 #H2 destruct
+  lapply (IHT12 HT1U2) -IHT12 -HT1U2 #HUT2
   lapply (simple_thom_repl_dx … HUT2 HT1) /2 width=1/
 | #V1 #V2 #W #T1 #T2 #_ #_ #_ #_ #H
   elim (thom_inv_flat1 … H) -H #W2 #U2 #_ #H
   elim (simple_inv_bind … H)
 | #I #V1 #V2 #T1 #T #T2 #_ #_ #_ #_ #_ #H
-  elim (thom_inv_bind1 … H) -H #W2 #U2 #H destruct -I //
+  elim (thom_inv_bind1 … H) -H #W2 #U2 #H destruct //
 | #V2 #V1 #V #W1 #W2 #T1 #T2 #_ #_ #_ #_ #_ #_ #_ #H
   elim (thom_inv_flat1 … H) -H #U1 #U2 #_ #H
   elim (simple_inv_bind … H)
@@ -53,4 +53,4 @@ lemma tpr_thom: ∀T1,T2. T1 ⇒ T2 → T1 ≈ T1 → T1 ≈ T2.
 qed.
 
 lemma twhnf_thom: ∀T. T ≈ T → 𝕎ℍℕ[T].
-/2/ qed.
+/2 width=1/ qed.