Skip to content

Commit

Permalink
Merge pull request #3753 from mtzguido/cmdline
Browse files Browse the repository at this point in the history
Simplifying command-line usage of F*
  • Loading branch information
mtzguido authored Feb 14, 2025
2 parents 69a61e9 + b95051a commit f46939e
Show file tree
Hide file tree
Showing 7 changed files with 60 additions and 22 deletions.
15 changes: 4 additions & 11 deletions mk/generic-1.mk
Original file line number Diff line number Diff line change
Expand Up @@ -69,24 +69,17 @@ FSTAR := $(WINWRAP) $(FSTAR_EXE) $(SIL) $(FSTAR_OPTIONS)
$(maybe_touch)

%.$(EEXT): FF=$(notdir $(subst $(EXTENSION),,$<))
%.$(EEXT): MM=$(basename $(FF))
%.$(EEXT): LBL=$(notdir $@)
# ^ HACK we use notdir to get the module name since we need to pass in
# the fst (not the checked file), but we don't know where it is, so this
# is relying on F* looking in its include path.
%.$(EEXT):
$(call msg, "EXTRACT", $(LBL))
$(FSTAR) $(if $(findstring FStarC.,$<),--MLish,) $(FF) --already_cached '*,' --codegen $(CODEGEN) --extract_module $(MM)
$(call msg, "EXTRACT", $(FF))
$(FSTAR) $< $(if $(findstring FStarC.,$<),--MLish,) $(FF) --already_cached '*,' --codegen $(CODEGEN)
@# HACK: finding FStarC modules and passing --MLish
@# for them and only them.
$(maybe_touch)

%.krml: FF=$(notdir $(subst $(EXTENSION),,$<))
%.krml: MM=$(basename $(FF))
%.krml: LBL=$(notdir $@)
%.krml:
$(call msg, "EXTRACT", $(LBL))
$(FSTAR) $(FF) --already_cached ',*' --codegen krml --extract_module $(MM)
$(call msg, "EXTRACT", $(FF))
$(FSTAR) $< --already_cached ',*' --codegen krml

DEPSTEM := $(CACHE_DIR)/.depend$(TAG)

Expand Down
8 changes: 4 additions & 4 deletions mk/test.mk
Original file line number Diff line number Diff line change
Expand Up @@ -100,13 +100,13 @@ $(OUTPUT_DIR)/%.fsti.json_output: %.fsti
@mkdir -p $(dir $@)
$(FSTAR) --message_format json --silent -f --print_expected_failures $< >$@ 2>&1

$(OUTPUT_DIR)/$(subst .,_,%).ml:
$(OUTPUT_DIR)/%.ml:
$(call msg, "EXTRACT", $(basename $(notdir $@)))
$(FSTAR) $(subst .checked,,$(notdir $<)) --codegen OCaml --extract_module $(subst .fst.checked,,$(notdir $<))
$(FSTAR) $< --codegen OCaml

$(OUTPUT_DIR)/$(subst .,_,%).fs:
$(OUTPUT_DIR)/%.fs:
$(call msg, "EXTRACT FS", $(basename $(notdir $@)))
$(FSTAR) $(subst .checked,,$(notdir $<)) --codegen FSharp --extract_module $(subst .fst.checked,,$(notdir $<))
$(FSTAR) $< --codegen FSharp

# No FSharp compilation in these makefiles, sorry.
$(OUTPUT_DIR)/%.exe: $(OUTPUT_DIR)/%.ml
Expand Down
9 changes: 8 additions & 1 deletion src/basic/FStarC.Options.fst
Original file line number Diff line number Diff line change
Expand Up @@ -1883,6 +1883,7 @@ let module_name_of_file_name f =
let f = String.substring f 0 (String.length f - String.length (Filepath.get_file_extension f) - 1) in
String.lowercase f
(* Basically returns true when the module is in the command line *)
let should_check m =
let l = get_verify_module () in
List.contains (String.lowercase m) l
Expand Down Expand Up @@ -2356,7 +2357,13 @@ let should_extract (m:string) (tgt:codegen_t) : bool =
in
not (no_extract m) &&
(match get_extract_namespace (), get_extract_module() with
| [], [] -> true //neither is set; extract everything
| [], [] ->
// Neither is set; extract only files given in the command line.
// Except for krml: there we retain the behavior of extracting everything
// into a single krml output file.
if tgt = Krml
then true
else should_check m
| _ -> should_extract_namespace m || should_extract_module m)
let should_be_already_cached m =
Expand Down
34 changes: 33 additions & 1 deletion src/fstar/FStarC.Main.fst
Original file line number Diff line number Diff line change
Expand Up @@ -136,7 +136,39 @@ let print_help_for (o : string) : unit =

(* Normal mode with some flags, files, etc *)
let go_normal () =
let res, filenames = process_args () in
let res, filenames0 = process_args () in

let chopsuf (suf s : string) : option string =
if ends_with s suf
then Some (String.substring s 0 (String.length s - String.length suf))
else None
in
let ( ||| ) x y =
match x with
| None -> y
| _ -> x
in
let checked_of (f:string) =
chopsuf ".checked" f ||| chopsuf ".checked.lax" f
in

let filenames =
filenames0 |>
List.map (fun f ->
if not (Filepath.file_exists f) then f else
(* ^ only rewrite if file exists *)
match checked_of f with
| Some f' ->
if Debug.any () then
print1 "Rewriting argument file %s to its source file\n" f;
(match Find.find_file (Filepath.basename f') with
| Some r -> r
| None -> failwith "Couldn't find source for file" ^ f' ^ "!\n")
| None -> f
)
in
if Debug.any () then
print2 "Rewrote\n%s\ninto\n%s\n\n" (show filenames0) (show filenames);

(* Compat: create the --odir and --cache_dir if they don't exist.
F* has done this for a long time, only sinc it simplified
Expand Down
7 changes: 5 additions & 2 deletions tests/dune_hello/Makefile
Original file line number Diff line number Diff line change
Expand Up @@ -8,8 +8,11 @@ ifneq ($(OS),Windows_NT) # See comment in ../simple_hello/Makefile
all: run.bc
endif

Hello.ml: Hello.fst
$(FSTAR_EXE) --codegen OCaml Hello.fst --extract Hello --z3version 4.13.3
%.checked: %
$(FSTAR_EXE) $< --cache_checked_modules --z3version 4.13.3

%.ml: %.fst.checked
$(FSTAR_EXE) --codegen OCaml $<

bin/hello.exe: Hello.ml
$(FSTAR_EXE) --ocamlenv dune build @install --profile=release
Expand Down
2 changes: 1 addition & 1 deletion tests/semiring/Makefile
Original file line number Diff line number Diff line change
Expand Up @@ -11,7 +11,7 @@ all: $(OUTPUT_DIR)/CanonCommSemiring.interpreted $(OUTPUT_DIR)/CanonCommSemiring
# If you add explicit dependencies here, it will fall back to the default rule.
$(OUTPUT_DIR)/%.ml:
$(call msg, "EXTRACT", $<)
$(FSTAR) --codegen Plugin --extract $* $*.fst
$(FSTAR) --codegen Plugin $<
cat $*.ml.fixup >> $@

$(OUTPUT_DIR)/%.cmxs: $(OUTPUT_DIR)/%.ml
Expand Down
7 changes: 5 additions & 2 deletions tests/simple_hello/Makefile
Original file line number Diff line number Diff line change
Expand Up @@ -20,8 +20,11 @@ endif
Hello.%.test: Hello.%
./$< | grep "Hello F\*!"

%.ml: %.fst
$(FSTAR_EXE) --codegen OCaml $< --extract $* --z3version 4.13.3
%.checked: %
$(FSTAR_EXE) $< --cache_checked_modules --z3version 4.13.3

%.ml: %.fst.checked
$(FSTAR_EXE) --codegen OCaml $<

%.exe: %.ml
$(FSTAR_EXE) --ocamlopt $< -o $@
Expand Down

0 comments on commit f46939e

Please sign in to comment.