]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambda_delta/Basic_2/reducibility/tnf.ma
- main lemmas about abstract reducibility candidates closed
[helm.git] / matita / matita / contribs / lambda_delta / Basic_2 / reducibility / tnf.ma
index 5868a5c935777d578f0eae1fda58ab43fbb12439..3c0184e551a4a7a36aff4fae2264b928d0b0dab0 100644 (file)
@@ -18,8 +18,7 @@ include "Basic_2/reducibility/tpr.ma".
 
 (* CONTEXT-FREE NORMAL TERMS ************************************************)
 
-definition tnf: term → Prop ≝
-   NF … tpr (eq …).
+definition tnf: predicate term ≝ NF … tpr (eq …).
 
 interpretation
    "context-free normality (term)"
@@ -27,34 +26,44 @@ interpretation
 
 (* Basic inversion lemmas ***************************************************)
 
-lemma tnf_inv_abst: ∀V,T. ℕ[𝕔{Abst}V.T] → ℕ[V] ∧ ℕ[T].
+lemma tnf_inv_abst: ∀V,T. ℕ[V.T] → ℕ[V] ∧ ℕ[T].
 #V1 #T1 #HVT1 @conj
-[ #V2 #HV2 lapply (HVT1 (𝕔{Abst}V2.T1) ?) -HVT1 /2/ -HV2 #H destruct -V1 T1 //
-| #T2 #HT2 lapply (HVT1 (𝕔{Abst}V1.T2) ?) -HVT1 /2/ -HT2 #H destruct -V1 T1 //
+[ #V2 #HV2 lapply (HVT1 (ⓛV2.T1) ?) -HVT1 /2 width=1/ -HV2 #H destruct //
+| #T2 #HT2 lapply (HVT1 (ⓛV1.T2) ?) -HVT1 /2 width=1/ -HT2 #H destruct //
 ]
-qed.
+qed-.
 
-lemma tnf_inv_appl: ∀V,T. ℕ[𝕔{Appl}V.T] → ∧∧ ℕ[V] & ℕ[T] & 𝕊[T].
+lemma tnf_inv_appl: ∀V,T. ℕ[V.T] → ∧∧ ℕ[V] & ℕ[T] & 𝕊[T].
 #V1 #T1 #HVT1 @and3_intro
-[ #V2 #HV2 lapply (HVT1 (𝕔{Appl}V2.T1) ?) -HVT1 /2/ -HV2 #H destruct -V1 T1 //
-| #T2 #HT2 lapply (HVT1 (𝕔{Appl}V1.T2) ?) -HVT1 /2/ -HT2 #H destruct -V1 T1 //
-| generalize in match HVT1 -HVT1; elim T1 -T1 * // * #W1 #U1 #_ #_ #H
+[ #V2 #HV2 lapply (HVT1 (ⓐV2.T1) ?) -HVT1 /2 width=1/ -HV2 #H destruct //
+| #T2 #HT2 lapply (HVT1 (ⓐV1.T2) ?) -HVT1 /2 width=1/ -HT2 #H destruct //
+| generalize in match HVT1; -HVT1 elim T1 -T1 * // * #W1 #U1 #_ #_ #H
   [ elim (lift_total V1 0 1) #V2 #HV12
-    lapply (H (𝕔{Abbr}W1.𝕔{Appl}V2.U1) ?) -H /2/ -HV12 #H destruct
-  | lapply (H (𝕔{Abbr}V1.U1) ?) -H /2/ #H destruct
+    lapply (H (ⓓW1.ⓐV2.U1) ?) -H /2 width=3/ -HV12 #H destruct
+  | lapply (H (ⓓV1.U1) ?) -H /2 width=1/ #H destruct
+]
+qed-.
+
+lemma tnf_inv_abbr: ∀V,T. ℕ[ⓓV.T] → False.
+#V #T #H elim (is_lift_dec T 0 1)
+[ * #U #HTU
+  lapply (H U ?) -H /2 width=3/ #H destruct
+  elim (lift_inv_pair_xy_y … HTU)
+| #HT
+  elim (tps_full (⋆) V T (⋆. ⓓV) 0 ?) // #T2 #T1 #HT2 #HT12
+  lapply (H (ⓓV.T2) ?) -H /2 width=3/ -HT2 #H destruct /3 width=2/
 ]
 qed.
 
-axiom tnf_inv_abbr: ∀V,T. ℕ[𝕔{Abbr}V.T] → False.
-
-lemma tnf_inv_cast: ∀V,T. ℕ[𝕔{Cast}V.T] → False.
-#V #T #H lapply (H T ?) -H /2/
-qed.
+lemma tnf_inv_cast: ∀V,T. ℕ[ⓣV.T] → False.
+#V #T #H lapply (H T ?) -H /2 width=1/ #H
+@(discr_tpair_xy_y … H)
+qed-.
 
 (* Basic properties *********************************************************)
 
-lemma tpr_tif_eq: â\88\80T1,T2. T1 â\87\92 T2 →  𝕀[T1] → T1 = T2.
-#T1 #T2 #H elim H -T1 T2
+lemma tpr_tif_eq: â\88\80T1,T2. T1 â\9e¡ T2 →  𝕀[T1] → T1 = T2.
+#T1 #T2 #H elim H -T1 -T2
 [ //
 | * #V1 #V2 #T1 #T2 #_ #_ #IHV1 #IHT1 #H
   [ elim (tif_inv_appl … H) -H #HV1 #HT1 #_
@@ -65,7 +74,7 @@ lemma tpr_tif_eq: ∀T1,T2. T1 ⇒ T2 →  𝕀[T1] → T1 = T2.
   elim (tif_inv_appl … H) -H #_ #_ #H
   elim (simple_inv_bind … H)
 | * #V1 #V2 #T1 #T #T2 #_ #_ #HT2 #IHV1 #IHT1 #H
-  [ -HT2 IHV1 IHT1; elim (tif_inv_abbr … H)
+  [ -HT2 -IHV1 -IHT1 elim (tif_inv_abbr … H)
   | <(tps_inv_refl_SO2 … HT2 ?) -HT2 //
     elim (tif_inv_abst … H) -H #HV1 #HT1
     >IHV1 -IHV1 // -HV1 >IHT1 -IHT1 //
@@ -81,15 +90,15 @@ lemma tpr_tif_eq: ∀T1,T2. T1 ⇒ T2 →  𝕀[T1] → T1 = T2.
 qed.
 
 theorem tif_tnf: ∀T1.  𝕀[T1] → ℕ[T1].
-/2/ qed.
+/2 width=1/ qed.
 
 (* Note: this property is unusual *)
 theorem tnf_trf_false: ∀T1. ℝ[T1] → ℕ[T1] → False.
 #T1 #H elim H -T1
-[ #V #T #_ #IHV #H elim (tnf_inv_abst … H) -H /2/
-| #V #T #_ #IHT #H elim (tnf_inv_abst … H) -H /2/
-| #V #T #_ #IHV #H elim (tnf_inv_appl … H) -H /2/
-| #V #T #_ #IHV #H elim (tnf_inv_appl … H) -H /2/
+[ #V #T #_ #IHV #H elim (tnf_inv_abst … H) -H /2 width=1/
+| #V #T #_ #IHT #H elim (tnf_inv_abst … H) -H /2 width=1/
+| #V #T #_ #IHV #H elim (tnf_inv_appl … H) -H /2 width=1/
+| #V #T #_ #IHV #H elim (tnf_inv_appl … H) -H /2 width=1/
 | #V #T #H elim (tnf_inv_abbr … H)
 | #V #T #H elim (tnf_inv_cast … H)
 | #V #W #T #H elim (tnf_inv_appl … H) -H #_ #_ #H
@@ -98,10 +107,10 @@ theorem tnf_trf_false: ∀T1. ℝ[T1] → ℕ[T1] → False.
 qed.
 
 theorem tnf_tif: ∀T1. ℕ[T1] → 𝕀[T1].
-/2/ qed.
+/2 width=3/ qed.
 
-lemma tnf_abst: ∀V,T. ℕ[V] → ℕ[T] → ℕ[𝕔{Abst}V.T].
+lemma tnf_abst: ∀V,T. ℕ[V] → ℕ[T] → ℕ[V.T].
 /4 width=1/ qed.
 
-lemma tnf_appl: ∀V,T. ℕ[V] → ℕ[T] → 𝕊[T] → ℕ[𝕔{Appl}V.T].
+lemma tnf_appl: ∀V,T. ℕ[V] → ℕ[T] → 𝕊[T] → ℕ[V.T].
 /4 width=1/ qed.