aboutsummaryrefslogtreecommitdiffhomepage
path: root/kernel/vars.ml
diff options
context:
space:
mode:
authorGravatar Hugo Herbelin <Hugo.Herbelin@inria.fr>2015-11-19 22:49:25 +0100
committerGravatar Hugo Herbelin <Hugo.Herbelin@inria.fr>2015-12-05 10:01:14 +0100
commitade2363e357db3ac3f258e645fe6bba988e7e7dd (patch)
treeade794510151d080d164be6d33d03aacbbe5064f /kernel/vars.ml
parentf66e604a9d714ee9dba09234d935ee208bc89d97 (diff)
About building of substitutions from instances.
Redefining adjust_subst_to_rel_context from instantiate_context who was hidden in inductiveops.ml, renamed the latter into subst_of_rel_context_instance and moving them to Vars. The new name highlights that the input is an instance (as for applist) and the output a substitution (as for substl). This is a clearer unified interface, centralizing the difficult de-Bruijn job in one place. It saves a couple of List.rev.
Diffstat (limited to 'kernel/vars.ml')
-rw-r--r--kernel/vars.ml15
1 files changed, 15 insertions, 0 deletions
diff --git a/kernel/vars.ml b/kernel/vars.ml
index f8c0a65e6..a00c7036f 100644
--- a/kernel/vars.ml
+++ b/kernel/vars.ml
@@ -164,6 +164,21 @@ let substnl_decl laml k r = map_rel_declaration (fun c -> substnl laml k c) r
let substl_decl laml r = map_rel_declaration (fun c -> substnl laml 0 c) r
let subst1_decl lam r = map_rel_declaration (fun c -> subst1 lam c) r
+(* Build a substitution from an instance, inserting missing let-ins *)
+
+let subst_of_rel_context_instance sign l =
+ let rec aux subst sign l =
+ match sign, l with
+ | (_,None,_)::sign', a::args' -> aux (a::subst) sign' args'
+ | (_,Some c,_)::sign', args' ->
+ aux (substl subst c :: subst) sign' args'
+ | [], [] -> subst
+ | _ -> Errors.anomaly (Pp.str "Instance and signature do not match")
+ in aux [] (List.rev sign) l
+
+let adjust_subst_to_rel_context sign l =
+ List.rev (subst_of_rel_context_instance sign l)
+
(* (thin_val sigma) removes identity substitutions from sigma *)
let rec thin_val = function