]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/legacy_1/coq/defs.ma
update in basic_2
[helm.git] / matita / matita / contribs / lambdadelta / legacy_1 / coq / defs.ma
index b059133d94c02c0fdc9f82938b05c76ccca1946d..497403d328dc5b76e0462a651e11c5cb41b4988e 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
@@ -68,13 +68,12 @@ definition pred:
 \def
  \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_2 \def (plus p m) in (S 
-TMP_2))]).
+rec definition plus (n: nat) on n: nat \to nat \def \lambda (m: nat).(match n 
+with [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 wit
-[O \Rightarrow O | (S k) \Rightarrow (match m with [O \Rightarrow (S k) | (S 
-l) \Rightarrow (minus k l)])]).
+rec definition minus (n: nat) on n: nat \to nat \def \lambda (m: nat).(matc
+n with [O \Rightarrow O | (S k) \Rightarrow (match m with [O \Rightarrow (S 
+k) | (S l) \Rightarrow (minus k l)])]).
 
 inductive Acc (A: Type[0]) (R: A \to (A \to Prop)): A \to Prop \def
 | Acc_intro: \forall (x: A).(((\forall (y: A).((R y x) \to (Acc A R y)))) \to 
@@ -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_4 \def (f a) in (let TMP_3 \def (f b) in (lt TMP_4 
-TMP_3)))))).
+(b: A).(lt (f a) (f b))))).