]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/lib/lambda/subst.ma
- rc_sat.ma: we changed the notation for extensional equality. we now
[helm.git] / matita / matita / lib / lambda / subst.ma
index 85bef55e155fa0e4a01a83de2fcaeff28d56fc47..bd8c5b71353226e7f78f5b5257ead19727351876 100644 (file)
 include "arithmetics/nat.ma".
 
 inductive T : Type[0] ≝
-  | Sort: nat → T
-  | Rel: nat → T 
-  | App: T → T → T 
+  | Sort: nat → T     (* starts from 0 *)
+  | Rel: nat → T      (* starts from ... ? *)
+  | App: T → T → T    (* function, argument *)
   | Lambda: T → T → T (* type, body *)
-  | Prod: T → T → T (* type, body *)
-  | D: T →T
+  | Prod: T → T → T   (* type, body *)
+  | D: T → T          (* dummifier *)
 .
 
+(* arguments: k is the depth (starts from 0), p is the height (starts from 0) *)
 let rec lift t k p ≝
   match t with 
     [ Sort n ⇒ Sort n
@@ -227,8 +228,3 @@ lemma subst_lemma: ∀A,B,C.∀k,i.
     ]
   ] 
 qed.
-
-  
-
-
-