aboutsummaryrefslogtreecommitdiffhomepage
path: root/kernel/term.ml
diff options
context:
space:
mode:
Diffstat (limited to 'kernel/term.ml')
-rw-r--r--kernel/term.ml112
1 files changed, 62 insertions, 50 deletions
diff --git a/kernel/term.ml b/kernel/term.ml
index ad8ae3be7..9ba45f540 100644
--- a/kernel/term.ml
+++ b/kernel/term.ml
@@ -10,7 +10,6 @@ open Util
open Pp
open Errors
open Names
-open Context
open Vars
(**********************************************************************)
@@ -471,26 +470,58 @@ let rec to_prod n lam =
| Cast (c,_,_) -> to_prod n c
| _ -> errorlabstrm "to_prod" (mt ())
-(* pseudo-reduction rule:
- * [prod_app s (Prod(_,B)) N --> B[N]
- * with an strip_outer_cast on the first argument to produce a product *)
+let it_mkProd_or_LetIn = List.fold_left (fun c d -> mkProd_or_LetIn d c)
+let it_mkLambda_or_LetIn = List.fold_left (fun c d -> mkLambda_or_LetIn d c)
-let prod_app t n =
- match kind_of_term (strip_outer_cast t) with
- | Prod (_,_,b) -> subst1 n b
- | _ ->
- errorlabstrm "prod_app"
- (str"Needed a product, but didn't find one" ++ fnl ())
+(* Application with expected on-the-fly reduction *)
+let lambda_applist c l =
+ let rec app subst c l =
+ match kind_of_term c, l with
+ | Lambda(_,_,c), arg::l -> app (arg::subst) c l
+ | _, [] -> substl subst c
+ | _ -> anomaly (Pp.str "Not enough lambda's") in
+ app [] c l
-(* prod_appvect T [| a1 ; ... ; an |] -> (T a1 ... an) *)
-let prod_appvect t nL = Array.fold_left prod_app t nL
+let lambda_appvect c v = lambda_applist c (Array.to_list v)
+
+let lambda_applist_assum n c l =
+ let rec app n subst t l =
+ if Int.equal n 0 then
+ if l == [] then substl subst t
+ else anomaly (Pp.str "Not enough arguments")
+ else match kind_of_term t, l with
+ | Lambda(_,_,c), arg::l -> app (n-1) (arg::subst) c l
+ | LetIn(_,b,_,c), _ -> app (n-1) (substl subst b::subst) c l
+ | _ -> anomaly (Pp.str "Not enough lambda/let's") in
+ app n [] c l
+
+let lambda_appvect_assum n c v = lambda_applist_assum n c (Array.to_list v)
(* prod_applist T [ a1 ; ... ; an ] -> (T a1 ... an) *)
-let prod_applist t nL = List.fold_left prod_app t nL
+let prod_applist c l =
+ let rec app subst c l =
+ match kind_of_term c, l with
+ | Prod(_,_,c), arg::l -> app (arg::subst) c l
+ | _, [] -> substl subst c
+ | _ -> anomaly (Pp.str "Not enough prod's") in
+ app [] c l
-let it_mkProd_or_LetIn = List.fold_left (fun c d -> mkProd_or_LetIn d c)
-let it_mkLambda_or_LetIn = List.fold_left (fun c d -> mkLambda_or_LetIn d c)
+(* prod_appvect T [| a1 ; ... ; an |] -> (T a1 ... an) *)
+let prod_appvect c v = prod_applist c (Array.to_list v)
+
+let prod_applist_assum n c l =
+ let rec app n subst t l =
+ if Int.equal n 0 then
+ if l == [] then substl subst t
+ else anomaly (Pp.str "Not enough arguments")
+ else match kind_of_term t, l with
+ | Prod(_,_,c), arg::l -> app (n-1) (arg::subst) c l
+ | LetIn(_,b,_,c), _ -> app (n-1) (substl subst b::subst) c l
+ | _ -> anomaly (Pp.str "Not enough prod/let's") in
+ app n [] c l
+
+let prod_appvect_assum n c v = prod_applist_assum n c (Array.to_list v)
(*********************************)
(* Other term destructors *)
@@ -547,24 +578,24 @@ let decompose_lam_n n =
let decompose_prod_assum =
let rec prodec_rec l c =
match kind_of_term c with
- | Prod (x,t,c) -> prodec_rec (add_rel_decl (x,None,t) l) c
- | LetIn (x,b,t,c) -> prodec_rec (add_rel_decl (x,Some b,t) l) c
+ | Prod (x,t,c) -> prodec_rec (Context.Rel.add (x,None,t) l) c
+ | LetIn (x,b,t,c) -> prodec_rec (Context.Rel.add (x,Some b,t) l) c
| Cast (c,_,_) -> prodec_rec l c
| _ -> l,c
in
- prodec_rec empty_rel_context
+ prodec_rec Context.Rel.empty
(* Transforms a lambda term [x1:T1]..[xn:Tn]T into the pair
([(xn,Tn);...;(x1,T1)],T), where T is not a lambda *)
let decompose_lam_assum =
let rec lamdec_rec l c =
match kind_of_term c with
- | Lambda (x,t,c) -> lamdec_rec (add_rel_decl (x,None,t) l) c
- | LetIn (x,b,t,c) -> lamdec_rec (add_rel_decl (x,Some b,t) l) c
+ | Lambda (x,t,c) -> lamdec_rec (Context.Rel.add (x,None,t) l) c
+ | LetIn (x,b,t,c) -> lamdec_rec (Context.Rel.add (x,Some b,t) l) c
| Cast (c,_,_) -> lamdec_rec l c
| _ -> l,c
in
- lamdec_rec empty_rel_context
+ lamdec_rec Context.Rel.empty
(* Given a positive integer n, decompose a product or let-in term
of the form [forall (x1:T1)..(xi:=ci:Ti)..(xn:Tn), T] into the pair
@@ -576,12 +607,12 @@ let decompose_prod_n_assum n =
let rec prodec_rec l n c =
if Int.equal n 0 then l,c
else match kind_of_term c with
- | Prod (x,t,c) -> prodec_rec (add_rel_decl (x,None,t) l) (n-1) c
- | LetIn (x,b,t,c) -> prodec_rec (add_rel_decl (x,Some b,t) l) (n-1) c
+ | Prod (x,t,c) -> prodec_rec (Context.Rel.add (x,None,t) l) (n-1) c
+ | LetIn (x,b,t,c) -> prodec_rec (Context.Rel.add (x,Some b,t) l) (n-1) c
| Cast (c,_,_) -> prodec_rec l n c
| c -> error "decompose_prod_n_assum: not enough assumptions"
in
- prodec_rec empty_rel_context n
+ prodec_rec Context.Rel.empty n
(* Given a positive integer n, decompose a lambda or let-in term [fun
(x1:T1)..(xi:=ci:Ti)..(xn:Tn) => T] into the pair of the abstracted
@@ -595,12 +626,12 @@ let decompose_lam_n_assum n =
let rec lamdec_rec l n c =
if Int.equal n 0 then l,c
else match kind_of_term c with
- | Lambda (x,t,c) -> lamdec_rec (add_rel_decl (x,None,t) l) (n-1) c
- | LetIn (x,b,t,c) -> lamdec_rec (add_rel_decl (x,Some b,t) l) n c
+ | Lambda (x,t,c) -> lamdec_rec (Context.Rel.add (x,None,t) l) (n-1) c
+ | LetIn (x,b,t,c) -> lamdec_rec (Context.Rel.add (x,Some b,t) l) n c
| Cast (c,_,_) -> lamdec_rec l n c
| c -> error "decompose_lam_n_assum: not enough abstractions"
in
- lamdec_rec empty_rel_context n
+ lamdec_rec Context.Rel.empty n
(* Same, counting let-in *)
let decompose_lam_n_decls n =
@@ -609,31 +640,12 @@ let decompose_lam_n_decls n =
let rec lamdec_rec l n c =
if Int.equal n 0 then l,c
else match kind_of_term c with
- | Lambda (x,t,c) -> lamdec_rec (add_rel_decl (x,None,t) l) (n-1) c
- | LetIn (x,b,t,c) -> lamdec_rec (add_rel_decl (x,Some b,t) l) (n-1) c
+ | Lambda (x,t,c) -> lamdec_rec (Context.Rel.add (x,None,t) l) (n-1) c
+ | LetIn (x,b,t,c) -> lamdec_rec (Context.Rel.add (x,Some b,t) l) (n-1) c
| Cast (c,_,_) -> lamdec_rec l n c
| c -> error "decompose_lam_n_decls: not enough abstractions"
in
- lamdec_rec empty_rel_context n
-
-(* (nb_lam [na1:T1]...[nan:Tan]c) where c is not an abstraction
- * gives n (casts are ignored) *)
-let nb_lam =
- let rec nbrec n c = match kind_of_term c with
- | Lambda (_,_,c) -> nbrec (n+1) c
- | Cast (c,_,_) -> nbrec n c
- | _ -> n
- in
- nbrec 0
-
-(* similar to nb_lam, but gives the number of products instead *)
-let nb_prod =
- let rec nbrec n c = match kind_of_term c with
- | Prod (_,_,c) -> nbrec (n+1) c
- | Cast (c,_,_) -> nbrec n c
- | _ -> n
- in
- nbrec 0
+ lamdec_rec Context.Rel.empty n
let prod_assum t = fst (decompose_prod_assum t)
let prod_n_assum n t = fst (decompose_prod_n_assum n t)
@@ -654,7 +666,7 @@ let strip_lam_n n t = snd (decompose_lam_n n t)
Such a term can canonically be seen as the pair of a context of types
and of a sort *)
-type arity = rel_context * sorts
+type arity = Context.Rel.t * sorts
let destArity =
let rec prodec_rec l c =