]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_1/C/fwd.ma
update in basic_2
[helm.git] / matita / matita / contribs / lambdadelta / basic_1 / C / fwd.ma
index d2802e2ba7e225562aedffac3b7e2b13dacbb117..43d43fc7144563dd3d2d59b9aba809cd46301cf6 100644 (file)
 
 include "basic_1/C/defs.ma".
 
-let rec C_rect (P: (C \to Type[0])) (f: (\forall (n: nat).(P (CSort n)))) 
-(f0: (\forall (c: C).((P c) \to (\forall (k: K).(\forall (t: T).(P (CHead c k 
-t))))))) (c: C) on c: P c \def match c with [(CSort n) \Rightarrow (f n) | 
-(CHead c0 k t) \Rightarrow (f0 c0 ((C_rect P f f0) c0) k t)].
+implied rec lemma C_rect (P: (C \to Type[0])) (f: (\forall (n: nat).(P (CSort 
+n)))) (f0: (\forall (c: C).((P c) \to (\forall (k: K).(\forall (t: T).(P 
+(CHead c k t))))))) (c: C) on c: P c \def match c with [(CSort n) \Rightarrow 
+(f n) | (CHead c0 k t) \Rightarrow (f0 c0 ((C_rect P f f0) c0) k t)].
 
-theorem C_ind:
+implied lemma C_ind:
  \forall (P: ((C \to Prop))).(((\forall (n: nat).(P (CSort n)))) \to 
 (((\forall (c: C).((P c) \to (\forall (k: K).(\forall (t: T).(P (CHead c k 
 t))))))) \to (\forall (c: C).(P c))))
 \def
  \lambda (P: ((C \to Prop))).(C_rect P).
 
-theorem clt_wf__q_ind:
+fact clt_wf__q_ind:
  \forall (P: ((C \to Prop))).(((\forall (n: nat).((\lambda (P0: ((C \to 
 Prop))).(\lambda (n0: nat).(\forall (c: C).((eq nat (cweight c) n0) \to (P0 
 c))))) P n))) \to (\forall (c: C).(P c)))
@@ -39,7 +39,7 @@ Prop))).(\lambda (H: ((\forall (n: nat).(\forall (c: C).((eq nat (cweight c)
 n) \to (P c)))))).(\lambda (c: C).(H (cweight c) c (refl_equal nat (cweight 
 c)))))).
 
-theorem clt_wf_ind:
+lemma clt_wf_ind:
  \forall (P: ((C \to Prop))).(((\forall (c: C).(((\forall (d: C).((clt d c) 
 \to (P d)))) \to (P c)))) \to (\forall (c: C).(P c)))
 \def