Skip to content

Commit

Permalink
Extraction: replace unused variable names by _ in funs and matchs (fix
Browse files Browse the repository at this point in the history
…#2842)

 This is done via a unique pass which seems roughly linear in practice,
 even on big developments like CompCert. There's a List.nth in an env at
 each MLrel, that could be made logarithmic if necessary via Okasaki's
 skew list for instance.

 Another approach would be to keep names (as a form of documentation), but
 prefix them by _ to please OCaml's warnings. For now, let's be radical and
 use the _ pattern.
  • Loading branch information
letouzey committed Dec 15, 2015
1 parent 5cd9f58 commit 4b1429d
Show file tree
Hide file tree
Showing 3 changed files with 75 additions and 2 deletions.
69 changes: 69 additions & 0 deletions plugins/extraction/mlutil.ml
Original file line number Diff line number Diff line change
Expand Up @@ -459,6 +459,75 @@ let nb_occur_match =
| MLglob _ | MLexn _ | MLdummy | MLaxiom -> 0
in nb 1

(* Replace unused variables by _ *)

let dump_unused_vars a =
let dump_id = function
| Dummy -> Dummy
| Id _ -> Id dummy_name
| Tmp _ -> Tmp dummy_name
in
let rec ren env a = match a with
| MLrel i ->
let () = (List.nth env (i-1)) := true in a

| MLlam (id,b) ->
let occ_id = ref false in
let b' = ren (occ_id::env) b in
if !occ_id then if b' == b then a else MLlam(id,b')
else MLlam(dump_id id,b')

| MLletin (id,b,c) ->
let occ_id = ref false in
let b' = ren env b in
let c' = ren (occ_id::env) c in
if !occ_id then
if b' == b && c' == c then a else MLletin(id,b',c')
else
(* 'let' without occurrence: shouldn't happen after simpl *)
MLletin(dump_id id,b',c')

| MLcase (t,e,br) ->
let e' = ren env e in
let br' = array_smartmap (ren_branch env) br in
if e' == e && br' == br then a else MLcase (t,e',br')

| MLfix (i,ids,v) ->
let env' = list_tabulate (fun _ -> ref false) (Array.length ids) @ env in
let v' = array_smartmap (ren env') v in
if v' == v then a else MLfix (i,ids,v')

| MLapp (b,l) ->
let b' = ren env b and l' = list_smartmap (ren env) l in
if b' == b && l' == l then a else MLapp (b',l')

| MLcons(t,r,l) ->
let l' = list_smartmap (ren env) l in
if l' == l then a else MLcons (t,r,l')

| MLtuple l ->
let l' = list_smartmap (ren env) l in
if l' == l then a else MLtuple l'

| MLmagic b ->
let b' = ren env b in
if b' == b then a else MLmagic b'

| MLglob _ | MLexn _ | MLdummy | MLaxiom -> a

and ren_branch env ((ids,p,b) as tr) =
let occs = List.map (fun _ -> ref false) ids in
let b' = ren (List.rev_append occs env) b in
let ids' =
List.map2
(fun id occ -> if !occ then id else dump_id id)
ids occs
in
if b' == b && ids = ids' then tr
else (ids',p,b')
in
ren [] a

(*s Lifting on terms.
[ast_lift k t] lifts the binding depth of [t] across [k] bindings. *)

Expand Down
2 changes: 2 additions & 0 deletions plugins/extraction/mlutil.mli
Original file line number Diff line number Diff line change
Expand Up @@ -111,6 +111,8 @@ val ast_subst : ml_ast -> ml_ast -> ml_ast

val ast_glob_subst : ml_ast Refmap'.t -> ml_ast -> ml_ast

val dump_unused_vars : ml_ast -> ml_ast

val normalize : ml_ast -> ml_ast
val optimize_fix : ml_ast -> ml_ast
val inline : global_reference -> ml_ast -> bool
Expand Down
6 changes: 4 additions & 2 deletions plugins/extraction/modutil.ml
Original file line number Diff line number Diff line change
Expand Up @@ -260,10 +260,12 @@ let dfix_to_mlfix rv av i =
order to preserve the global interface, later [depcheck_se] will get
rid of them if possible *)

let optim_ast t = dump_unused_vars (normalize t)

let rec optim_se top to_appear s = function
| [] -> []
| (l,SEdecl (Dterm (r,a,t))) :: lse ->
let a = normalize (ast_glob_subst !s a) in
let a = optim_ast (ast_glob_subst !s a) in
let i = inline r a in
if i then s := Refmap'.add r a !s;
let d = match optimize_fix a with
Expand All @@ -273,7 +275,7 @@ let rec optim_se top to_appear s = function
in
(l,SEdecl d) :: (optim_se top to_appear s lse)
| (l,SEdecl (Dfix (rv,av,tv))) :: lse ->
let av = Array.map (fun a -> normalize (ast_glob_subst !s a)) av in
let av = Array.map (fun a -> optim_ast (ast_glob_subst !s a)) av in
(* This fake body ensures that no fixpoint will be auto-inlined. *)
let fake_body = MLfix (0,[||],[||]) in
for i = 0 to Array.length rv - 1 do
Expand Down

0 comments on commit 4b1429d

Please sign in to comment.