Skip to content

Commit

Permalink
Merge pull request #3420 from mtzguido/options
Browse files Browse the repository at this point in the history
Some rework of the options module
  • Loading branch information
mtzguido authored Aug 28, 2024
2 parents 8b8de3f + 0d9668a commit b2e21e3
Show file tree
Hide file tree
Showing 22 changed files with 871 additions and 742 deletions.
2 changes: 1 addition & 1 deletion ocaml/fstar-lib/generated/FStar_Compiler_Range_Ops.ml

Some generated files are not rendered by default. Learn more about how customized files appear on GitHub.

1,066 changes: 505 additions & 561 deletions ocaml/fstar-lib/generated/FStar_Options.ml

Large diffs are not rendered by default.

67 changes: 67 additions & 0 deletions ocaml/fstar-lib/generated/FStar_Options_Ext.ml

Some generated files are not rendered by default. Learn more about how customized files appear on GitHub.

12 changes: 6 additions & 6 deletions ocaml/fstar-lib/generated/FStar_SMTEncoding_Encode.ml

Some generated files are not rendered by default. Learn more about how customized files appear on GitHub.

14 changes: 3 additions & 11 deletions ocaml/fstar-lib/generated/FStar_Tactics_V1_Basic.ml

Some generated files are not rendered by default. Learn more about how customized files appear on GitHub.

20 changes: 6 additions & 14 deletions ocaml/fstar-lib/generated/FStar_Tactics_V2_Basic.ml

Some generated files are not rendered by default. Learn more about how customized files appear on GitHub.

12 changes: 6 additions & 6 deletions ocaml/fstar-lib/generated/FStar_ToSyntax_ToSyntax.ml

Some generated files are not rendered by default. Learn more about how customized files appear on GitHub.

2 changes: 1 addition & 1 deletion ocaml/fstar-lib/generated/FStar_TypeChecker_Cfg.ml

Some generated files are not rendered by default. Learn more about how customized files appear on GitHub.

2 changes: 1 addition & 1 deletion ocaml/fstar-lib/generated/FStar_TypeChecker_Rel.ml

Some generated files are not rendered by default. Learn more about how customized files appear on GitHub.

2 changes: 1 addition & 1 deletion src/Makefile.boot
Original file line number Diff line number Diff line change
Expand Up @@ -37,7 +37,7 @@ EXTRACT_NAMESPACES=FStar.Extraction FStar.Parser \
# TODO: Do we really need this anymore? Which (implementation) modules
# from src/basic are *not* extracted?
EXTRACT_MODULES=FStar.Pervasives FStar.Common FStar.Thunk \
FStar.VConfig FStar.Options FStar.Ident FStar.Errors FStar.Errors.Codes \
FStar.VConfig FStar.Options FStar.Options.Ext FStar.Ident FStar.Errors FStar.Errors.Codes \
FStar.Errors.Msg FStar.Errors.Raise FStar.Const \
FStar.Compiler.Order FStar.Order FStar.Dependencies \
FStar.Interactive.CompletionTable \
Expand Down
2 changes: 1 addition & 1 deletion src/basic/FStar.Compiler.Range.Ops.fst
Original file line number Diff line number Diff line change
Expand Up @@ -49,7 +49,7 @@ let string_of_pos pos =
format2 "%s,%s" (string_of_int pos.line) (string_of_int pos.col)
let string_of_file_name f =
if Options.ide () then
if Options.ext_getv "fstar:no_absolute_paths" = "1" then
if Options.Ext.get "fstar:no_absolute_paths" = "1" then
basename f
else begin
try
Expand Down
70 changes: 70 additions & 0 deletions src/basic/FStar.Options.Ext.fst
Original file line number Diff line number Diff line change
@@ -0,0 +1,70 @@
(*
Copyright 2008-2024 Microsoft Research
Licensed under the Apache License, Version 2.0 (the "License");
you may not use this file except in compliance with the License.
You may obtain a copy of the License at
http://www.apache.org/licenses/LICENSE-2.0
Unless required by applicable law or agreed to in writing, software
distributed under the License is distributed on an "AS IS" BASIS,
WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or implied.
See the License for the specific language governing permissions and
limitations under the License.
*)
module FStar.Options.Ext

open FStar.Compiler
open FStar.Compiler.Effect
open FStar.Class.Show
module BU = FStar.Compiler.Util

type ext_state =
| E : map : BU.psmap string -> ext_state

let cur_state = BU.mk_ref (E (BU.psmap_empty ()))

(* Set a key-value pair in the map *)
let set (k:key) (v:value) : unit =
cur_state := E (BU.psmap_add (!cur_state).map k v)

(* Get the value from the map, or return "" if not there *)
let get (k:key) : value =
let r =
match BU.psmap_try_find (!cur_state).map k with
| None -> ""
| Some v -> v
in
r

(* Find a home *)
let is_prefix (s1 s2 : string) : ML bool =
let open FStar.Compiler.String in
let l1 = length s1 in
let l2 = length s2 in
l2 >= l1 && substring s2 0 l1 = s1

(* Get a list of all KV pairs that "begin" with k, considered
as a namespace. *)
let getns (ns:string) : list (key & value) =
let f k v acc =
if (ns^":") `is_prefix` k
then (k, v) :: acc
else acc
in
BU.psmap_fold (!cur_state).map f []

let all () : list (key & value) =
let f k v acc = (k, v) :: acc in
BU.psmap_fold (!cur_state).map f []

let save () : ext_state =
!cur_state

let restore (s:ext_state) : unit =
cur_state := s;
()

let reset () : unit =
cur_state := E (BU.psmap_empty ())
42 changes: 42 additions & 0 deletions src/basic/FStar.Options.Ext.fsti
Original file line number Diff line number Diff line change
@@ -0,0 +1,42 @@
(*
Copyright 2008-2024 Microsoft Research
Licensed under the Apache License, Version 2.0 (the "License");
you may not use this file except in compliance with the License.
You may obtain a copy of the License at
http://www.apache.org/licenses/LICENSE-2.0
Unless required by applicable law or agreed to in writing, software
distributed under the License is distributed on an "AS IS" BASIS,
WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or implied.
See the License for the specific language governing permissions and
limitations under the License.
*)
module FStar.Options.Ext

open FStar.Compiler.Effect

type key = string
type value = string

new
val ext_state : Type0

(* Set a key-value pair in the map *)
val set (k:key) (v:value) : unit

(* Get the value from the map, or return "" if not there *)
val get (k:key) : value

(* Get a list of all KV pairs that "begin" with k, considered
as a namespace. *)
val getns (ns:string) : list (key & value)

(* List all pairs *)
val all () : list (key & value)

val save () : ext_state
val restore (s:ext_state) : unit

val reset () : unit
Loading

0 comments on commit b2e21e3

Please sign in to comment.