]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/legacy_1/coq/defs.ma
legacy_1, ground_1, and basic_1 recommitted without anticipation
[helm.git] / matita / matita / contribs / lambdadelta / legacy_1 / coq / defs.ma
index d9e9f05e876e0a11068f1647e422c6513067c556..75e58b0364f93ac5077e8fb475a8ce0d70517b29 100644 (file)
@@ -55,7 +55,7 @@ inductive le (n: nat): nat \to Prop \def
 definition lt:
  nat \to (nat \to Prop)
 \def
- \lambda (n: nat).(\lambda (m: nat).(let TMP_1 \def (S n) in (le TMP_1 m))).
+ \lambda (n: nat).(\lambda (m: nat).(le (S n) m)).
 
 definition IsSucc:
  nat \to Prop
@@ -69,8 +69,7 @@ definition pred:
  \lambda (n: nat).(match n with [O \Rightarrow O | (S u) \Rightarrow u]).
 
 let rec plus (n: nat) on n: nat \to nat \def \lambda (m: nat).(match n with 
-[O \Rightarrow m | (S p) \Rightarrow (let TMP_1 \def (plus p m) in (S 
-TMP_1))]).
+[O \Rightarrow m | (S p) \Rightarrow (S (plus p m))]).
 
 let rec minus (n: nat) on n: nat \to nat \def \lambda (m: nat).(match n with 
 [O \Rightarrow O | (S k) \Rightarrow (match m with [O \Rightarrow (S k) | (S 
@@ -90,6 +89,5 @@ definition ltof:
  \forall (A: Type[0]).(((A \to nat)) \to (A \to (A \to Prop)))
 \def
  \lambda (A: Type[0]).(\lambda (f: ((A \to nat))).(\lambda (a: A).(\lambda 
-(b: A).(let TMP_1 \def (f a) in (let TMP_2 \def (f b) in (lt TMP_1 
-TMP_2)))))).
+(b: A).(lt (f a) (f b))))).