X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_1%2FT%2Fprops.ma;h=7b62a5a15ed03f87a33fd56e0968900bddbc60b8;hb=57ae1762497a5f3ea75740e2908e04adb8642cc2;hp=cce311989ed62153b122f9e4bb934af1ca43e409;hpb=14a8276e6d877c2281a1fda452ed3e4c150f5d39;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/basic_1/T/props.ma b/matita/matita/contribs/lambdadelta/basic_1/T/props.ma index cce311989..7b62a5a15 100644 --- a/matita/matita/contribs/lambdadelta/basic_1/T/props.ma +++ b/matita/matita/contribs/lambdadelta/basic_1/T/props.ma @@ -16,49 +16,49 @@ include "basic_1/T/fwd.ma". -theorem not_abbr_abst: +lemma not_abbr_abst: not (eq B Abbr Abst) \def - \lambda (H: (eq B Abbr Abst)).(let TMP_1 \def (\lambda (ee: B).(match ee -with [Abbr \Rightarrow True | Abst \Rightarrow False | Void \Rightarrow -False])) in (let H0 \def (eq_ind B Abbr TMP_1 I Abst H) in (False_ind False -H0))). + \lambda (H: (eq B Abbr Abst)).(let H0 \def (eq_ind B Abbr (\lambda (ee: +B).(match ee with [Abbr \Rightarrow True | Abst \Rightarrow False | Void +\Rightarrow False])) I Abst H) in (False_ind False H0)). -theorem not_void_abst: +lemma not_void_abst: not (eq B Void Abst) \def - \lambda (H: (eq B Void Abst)).(let TMP_1 \def (\lambda (ee: B).(match ee -with [Abbr \Rightarrow False | Abst \Rightarrow False | Void \Rightarrow -True])) in (let H0 \def (eq_ind B Void TMP_1 I Abst H) in (False_ind False -H0))). + \lambda (H: (eq B Void Abst)).(let H0 \def (eq_ind B Void (\lambda (ee: +B).(match ee with [Abbr \Rightarrow False | Abst \Rightarrow False | Void +\Rightarrow True])) I Abst H) in (False_ind False H0)). -theorem not_abbr_void: +lemma not_abbr_void: not (eq B Abbr Void) \def - \lambda (H: (eq B Abbr Void)).(let TMP_1 \def (\lambda (ee: B).(match ee -with [Abbr \Rightarrow True | Abst \Rightarrow False | Void \Rightarrow -False])) in (let H0 \def (eq_ind B Abbr TMP_1 I Void H) in (False_ind False -H0))). + \lambda (H: (eq B Abbr Void)).(let H0 \def (eq_ind B Abbr (\lambda (ee: +B).(match ee with [Abbr \Rightarrow True | Abst \Rightarrow False | Void +\Rightarrow False])) I Void H) in (False_ind False H0)). -theorem not_abst_void: +lemma not_abst_void: not (eq B Abst Void) \def - \lambda (H: (eq B Abst Void)).(let TMP_1 \def (\lambda (ee: B).(match ee -with [Abbr \Rightarrow False | Abst \Rightarrow True | Void \Rightarrow -False])) in (let H0 \def (eq_ind B Abst TMP_1 I Void H) in (False_ind False -H0))). + \lambda (H: (eq B Abst Void)).(let H0 \def (eq_ind B Abst (\lambda (ee: +B).(match ee with [Abbr \Rightarrow False | Abst \Rightarrow True | Void +\Rightarrow False])) I Void H) in (False_ind False H0)). -theorem tweight_lt: +lemma tweight_lt: \forall (t: T).(lt O (tweight t)) \def - \lambda (t: T).(let TMP_2 \def (\lambda (t0: T).(let TMP_1 \def (tweight t0) -in (lt O TMP_1))) in (let TMP_4 \def (\lambda (_: nat).(let TMP_3 \def (S O) -in (le_n TMP_3))) in (let TMP_6 \def (\lambda (_: nat).(let TMP_5 \def (S O) -in (le_n TMP_5))) in (let TMP_15 \def (\lambda (_: K).(\lambda (t0: -T).(\lambda (H: (lt O (tweight t0))).(\lambda (t1: T).(\lambda (_: (lt O -(tweight t1))).(let TMP_7 \def (S O) in (let TMP_8 \def (tweight t0) in (let -TMP_9 \def (tweight t1) in (let TMP_10 \def (plus TMP_8 TMP_9) in (let TMP_11 -\def (S O) in (let TMP_12 \def (tweight t0) in (let TMP_13 \def (tweight t1) -in (let TMP_14 \def (le_plus_trans TMP_11 TMP_12 TMP_13 H) in (le_S TMP_7 -TMP_10 TMP_14)))))))))))))) in (T_ind TMP_2 TMP_4 TMP_6 TMP_15 t))))). + \lambda (t: T).(T_ind (\lambda (t0: T).(lt O (tweight t0))) (\lambda (_: +nat).(le_n (S O))) (\lambda (_: nat).(le_n (S O))) (\lambda (_: K).(\lambda +(t0: T).(\lambda (H: (lt O (tweight t0))).(\lambda (t1: T).(\lambda (_: (lt O +(tweight t1))).(le_S (S O) (plus (tweight t0) (tweight t1)) (le_plus_trans (S +O) (tweight t0) (tweight t1) H))))))) t). + +lemma tle_r: + \forall (t: T).(tle t t) +\def + \lambda (t: T).(T_ind (\lambda (t0: T).(le (tweight t0) (tweight t0))) +(\lambda (_: nat).(le_n (S O))) (\lambda (_: nat).(le_n (S O))) (\lambda (_: +K).(\lambda (t0: T).(\lambda (_: (le (tweight t0) (tweight t0))).(\lambda +(t1: T).(\lambda (_: (le (tweight t1) (tweight t1))).(le_n (S (plus (tweight +t0) (tweight t1))))))))) t).