From 5f43c3df28d5d19a293b08ff1b6dc2eac1109c13 Mon Sep 17 00:00:00 2001 From: Lucas Franceschino Date: Wed, 11 Sep 2024 10:09:34 +0200 Subject: [PATCH 1/5] feat: add JSON serializers for errors --- src/basic/FStar.Compiler.Range.Type.fst | 18 ++++++++++++++++++ src/basic/FStar.Compiler.Range.Type.fsti | 4 ++++ src/basic/FStar.Errors.Msg.fst | 5 +++++ src/basic/FStar.Errors.Msg.fsti | 2 ++ src/basic/FStar.Errors.fst | 18 ++++++++++++++++++ src/basic/FStar.Errors.fsti | 5 +++++ 6 files changed, 52 insertions(+) diff --git a/src/basic/FStar.Compiler.Range.Type.fst b/src/basic/FStar.Compiler.Range.Type.fst index 32e5297af5f..3b7aba8c323 100644 --- a/src/basic/FStar.Compiler.Range.Type.fst +++ b/src/basic/FStar.Compiler.Range.Type.fst @@ -79,3 +79,21 @@ let mk_rng file_name start_pos end_pos = { } let mk_range f b e = let r = mk_rng f b e in range_of_rng r r + +open FStar.Json +let json_of_pos (r: pos): json + = JsonAssoc [ + "line", JsonInt r.line; + "col", JsonInt r.col; + ] +let json_of_rng (r: rng): json + = JsonAssoc [ + "file_name", JsonStr r.file_name; + "start_pos", json_of_pos r.start_pos; + "end_pos", json_of_pos r.end_pos; + ] +let json_of_range (r: range): json + = JsonAssoc [ + "def", json_of_rng r.def_range; + "use", json_of_rng r.use_range; + ] diff --git a/src/basic/FStar.Compiler.Range.Type.fsti b/src/basic/FStar.Compiler.Range.Type.fsti index e5a52f5a333..b55d17ea68f 100644 --- a/src/basic/FStar.Compiler.Range.Type.fsti +++ b/src/basic/FStar.Compiler.Range.Type.fsti @@ -37,3 +37,7 @@ val set_use_range: range -> rng -> range val set_def_range: range -> rng -> range val mk_pos: int -> int -> pos val mk_range: string -> pos -> pos -> range + +val json_of_pos: pos -> FStar.Json.json +val json_of_rng: rng -> FStar.Json.json +val json_of_range: range -> FStar.Json.json diff --git a/src/basic/FStar.Errors.Msg.fst b/src/basic/FStar.Errors.Msg.fst index b0252bd26cb..0799c2b3c22 100644 --- a/src/basic/FStar.Errors.Msg.fst +++ b/src/basic/FStar.Errors.Msg.fst @@ -56,3 +56,8 @@ let subdoc d = subdoc' true d let rendermsg (ds : list document) : string = renderdoc (concat (List.map (fun d -> subdoc (group d)) ds)) + +let json_of_error_message (err_msg: list document): FStar.Json.json + = FStar.Compiler.List.map + (fun doc -> FStar.Json.JsonStr (renderdoc doc)) err_msg + |> FStar.Json.JsonList diff --git a/src/basic/FStar.Errors.Msg.fsti b/src/basic/FStar.Errors.Msg.fsti index e30f81a38a3..90003ebcbf4 100644 --- a/src/basic/FStar.Errors.Msg.fsti +++ b/src/basic/FStar.Errors.Msg.fsti @@ -62,3 +62,5 @@ val backtrace_doc : unit -> document (* Render an error message as a string. *) val rendermsg : error_message -> string + +val json_of_error_message: list document -> FStar.Json.json diff --git a/src/basic/FStar.Errors.fst b/src/basic/FStar.Errors.fst index 8617b2b48ed..7810f89cafd 100644 --- a/src/basic/FStar.Errors.fst +++ b/src/basic/FStar.Errors.fst @@ -29,6 +29,7 @@ module PP = FStar.Pprint open FStar.Errors.Codes open FStar.Errors.Msg +open FStar.Json let fallback_range : ref (option range) = BU.mk_ref None @@ -121,6 +122,23 @@ exception Warning of error exception Stop exception Empty_frag +let json_of_issue_level level + = JsonStr ( match level with + | ENotImplemented -> "NotImplemented" + | EInfo -> "Info" + | EWarning -> "Warning" + | EError -> "Error") + +let json_of_issue issue = + JsonAssoc [ + "msg", json_of_error_message issue.issue_msg; + "level", json_of_issue_level issue.issue_level; + "range", dflt JsonNull (json_of_range <$> issue.issue_range); + "number", dflt JsonNull (JsonInt <$> issue.issue_number); + "ctx", JsonList (JsonStr <$> issue.issue_ctx); + ] + + let ctx_doc (ctx : list string) : PP.document = let open FStar.Pprint in if Options.error_contexts () diff --git a/src/basic/FStar.Errors.fsti b/src/basic/FStar.Errors.fsti index 0f9d6443e4f..a99b4ebd9d3 100644 --- a/src/basic/FStar.Errors.fsti +++ b/src/basic/FStar.Errors.fsti @@ -22,6 +22,7 @@ include FStar.Errors.Codes include FStar.Errors.Msg open FStar.Errors.Msg open FStar.Class.HasRange +open FStar.Json {json} (* This is a fallback to be used if an error is raised/logged with a dummy range. It is set by TypeChecker.Tc.process_one_decl to @@ -61,6 +62,8 @@ type issue_level = | EWarning | EError +val json_of_issue_level: issue_level -> json + type issue = { issue_msg: error_message; issue_level: issue_level; @@ -69,6 +72,8 @@ type issue = { issue_ctx: list string; } +val json_of_issue: issue -> json + type error_handler = { eh_name: string; (* just for debugging purposes *) eh_add_one: issue -> unit; From f17704ba342e0ccacc12fefa9bc042d0bbff9a0e Mon Sep 17 00:00:00 2001 From: Lucas Franceschino Date: Wed, 11 Sep 2024 10:10:42 +0200 Subject: [PATCH 2/5] feat: add `message_format` option: `human` or `json` --- src/basic/FStar.Errors.fst | 11 ++++++++++- src/basic/FStar.Options.fst | 12 ++++++++++++ src/basic/FStar.Options.fsti | 3 +++ 3 files changed, 25 insertions(+), 1 deletion(-) diff --git a/src/basic/FStar.Errors.fst b/src/basic/FStar.Errors.fst index 7810f89cafd..aacb51bb227 100644 --- a/src/basic/FStar.Errors.fst +++ b/src/basic/FStar.Errors.fst @@ -22,6 +22,7 @@ open FStar.Compiler.Effect open FStar.Compiler.List open FStar.Compiler.Util open FStar.Compiler.Range +open FStar.Class.Monad open FStar.Options module List = FStar.Compiler.List module BU = FStar.Compiler.Util @@ -226,7 +227,10 @@ let format_issue' (print_hdr:bool) (issue:issue) : string = let format_issue issue : string = format_issue' true issue -let print_issue issue = +let print_issue_json issue = + json_of_issue issue |> string_of_json |> BU.print1_error "%s\n" + +let print_issue_rendered issue = let printer = match issue.issue_level with | EInfo -> (fun s -> BU.print_string (colorize_cyan s)) @@ -235,6 +239,11 @@ let print_issue issue = | ENotImplemented -> BU.print_error in printer (format_issue issue ^ "\n") +let print_issue issue = + match FStar.Options.message_format () with + | Human -> print_issue_rendered issue + | Json -> print_issue_json issue + let compare_issues i1 i2 = match i1.issue_range, i2.issue_range with | None, None -> 0 diff --git a/src/basic/FStar.Options.fst b/src/basic/FStar.Options.fst index 32249b18f33..24d5ef26ea5 100644 --- a/src/basic/FStar.Options.fst +++ b/src/basic/FStar.Options.fst @@ -204,6 +204,7 @@ let defaults = ("eager_subtyping" , Bool false); ("error_contexts" , Bool false); ("expose_interfaces" , Bool false); + ("message_format" , String "human"); ("ext" , Unset); ("extract" , Unset); ("extract_all" , Bool false); @@ -470,6 +471,7 @@ let get_dump_module () = lookup_opt "dump_module" let get_eager_subtyping () = lookup_opt "eager_subtyping" as_bool let get_error_contexts () = lookup_opt "error_contexts" as_bool let get_expose_interfaces () = lookup_opt "expose_interfaces" as_bool +let get_message_format () = lookup_opt "message_format" as_string let get_extract () = lookup_opt "extract" (as_option (as_list as_string)) let get_extract_module () = lookup_opt "extract_module" (as_list as_string) let get_extract_namespace () = lookup_opt "extract_namespace" (as_list as_string) @@ -964,6 +966,11 @@ let rec specs_with_types warn_unsafe : list (char & string & opt_type & Pprint.d Const (Bool true), text "Explicitly break the abstraction imposed by the interface of any implementation file that appears on the command line (use with care!)"); + ( noshort, + "message_format", + EnumStr ["human"; "json"], + text "Format of the messages emitted by F* (default `human`)"); + ( noshort, "hide_uvar_nums", Const (Bool true), @@ -1998,6 +2005,11 @@ let dump_module s = get_dump_module() |> List.existsb (module_ let eager_subtyping () = get_eager_subtyping() let error_contexts () = get_error_contexts () let expose_interfaces () = get_expose_interfaces () +let message_format () = + match get_message_format () with + | "human" -> Human + | "json" -> Json + | illegal -> failwith ("print_issue: option `message_format` was expected to be `human` or `json`, not `" ^ illegal ^ "`. This should be impossible: `message_format` was supposed to be validated.") let force () = get_force () let full_context_dependency () = true let hide_uvar_nums () = get_hide_uvar_nums () diff --git a/src/basic/FStar.Options.fsti b/src/basic/FStar.Options.fsti index d33f3aa3c5d..b8df40644b7 100644 --- a/src/basic/FStar.Options.fsti +++ b/src/basic/FStar.Options.fsti @@ -28,6 +28,8 @@ type codegen_t = type split_queries_t = | No | OnFailure | Always +type message_format_t = | Json | Human + type option_val = | Bool of bool | String of string @@ -129,6 +131,7 @@ val dump_module : string -> bool val eager_subtyping : unit -> bool val error_contexts : unit -> bool val expose_interfaces : unit -> bool +val message_format : unit -> message_format_t val file_list : unit -> list string val find_file : (string -> option string) val force : unit -> bool From 9131a8b9b0f0f79e69b07d17e8c507c832416635 Mon Sep 17 00:00:00 2001 From: Lucas Franceschino Date: Tue, 24 Sep 2024 08:02:25 +0200 Subject: [PATCH 3/5] feat: `message_format`: JSON profiling messages --- src/basic/FStar.Profiling.fst | 50 +++++++++++++++++++++++++---------- 1 file changed, 36 insertions(+), 14 deletions(-) diff --git a/src/basic/FStar.Profiling.fst b/src/basic/FStar.Profiling.fst index 35c535551da..f7a6ed9cb51 100644 --- a/src/basic/FStar.Profiling.fst +++ b/src/basic/FStar.Profiling.fst @@ -21,6 +21,7 @@ open FStar.Compiler.Effect module List = FStar.Compiler.List open FStar.Options module BU = FStar.Compiler.Util +open FStar.Json (* A counter id is the name of a profiling phase; @@ -36,6 +37,14 @@ type counter = { undercount:ref bool; } +let json_of_counter (c: counter) = + JsonAssoc [ + "id", JsonStr c.cid; + "total_time", JsonInt !c.total_time; + "running", JsonBool !c.running; + "undercount", JsonBool !c.undercount; + ] + (* Creating a new counter *) let new_counter cid = { cid = cid; @@ -78,6 +87,32 @@ let profile (f: unit -> 'a) (module_name:option string) (cid:string) : 'a = end else f() +let report_json tag c = + let counter = json_of_counter c in + JsonAssoc [ + "tag", JsonStr tag; + "counter", counter; + ] |> string_of_json |> BU.print1_error "%s\n" + +let report_human tag c = + let warn = if !c.running + then " (Warning, this counter is still running)" + else if !c.undercount + then " (Warning, some operations raised exceptions and we not accounted for)" + else "" + in + //print each counter's profile + BU.print4 "%s, profiled %s:\t %s ms%s\n" + tag + c.cid + (BU.string_of_int (!c.total_time)) + warn + +let report tag c = + match FStar.Options.message_format () with + | Human -> report_human tag c + | Json -> report_json tag c + (* Report all profiles and clear all counters *) let report_and_clear tag = let ctrs = //all the counters as a list @@ -87,17 +122,4 @@ let report_and_clear tag = let ctrs = //sort counters in descending order by elapsed time BU.sort_with (fun c1 c2 -> !c2.total_time - !c1.total_time) ctrs in - ctrs |> List.iter - (fun c -> - let warn = if !c.running - then " (Warning, this counter is still running)" - else if !c.undercount - then " (Warning, some operations raised exceptions and we not accounted for)" - else "" - in - //print each counter's profile - BU.print4 "%s, profiled %s:\t %s ms%s\n" - tag - c.cid - (BU.string_of_int (!c.total_time)) - warn) + List.iter (report tag) ctrs From 4e7ce3073cbf9b387104d55a1a38877dbc07eec5 Mon Sep 17 00:00:00 2001 From: Lucas Franceschino Date: Sat, 21 Sep 2024 11:13:51 +0200 Subject: [PATCH 4/5] chore: refresh OCaml snapshot --- .../generated/FStar_Compiler_Range_Type.ml | 29 +- ocaml/fstar-lib/generated/FStar_Errors.ml | 119 +- ocaml/fstar-lib/generated/FStar_Errors_Msg.ml | 10 +- ocaml/fstar-lib/generated/FStar_Options.ml | 1024 +++++++++-------- ocaml/fstar-lib/generated/FStar_Profiling.ml | 78 +- 5 files changed, 728 insertions(+), 532 deletions(-) diff --git a/ocaml/fstar-lib/generated/FStar_Compiler_Range_Type.ml b/ocaml/fstar-lib/generated/FStar_Compiler_Range_Type.ml index 95a831e301d..0fb23b6e99a 100644 --- a/ocaml/fstar-lib/generated/FStar_Compiler_Range_Type.ml +++ b/ocaml/fstar-lib/generated/FStar_Compiler_Range_Type.ml @@ -66,4 +66,31 @@ let (mk_rng : Prims.string -> pos -> pos -> rng) = fun start_pos -> fun end_pos -> { file_name = file_name1; start_pos; end_pos } let (mk_range : Prims.string -> pos -> pos -> range) = - fun f -> fun b -> fun e -> let r = mk_rng f b e in range_of_rng r r \ No newline at end of file + fun f -> fun b -> fun e -> let r = mk_rng f b e in range_of_rng r r +let (json_of_pos : pos -> FStar_Json.json) = + fun r -> + FStar_Json.JsonAssoc + [("line", (FStar_Json.JsonInt (r.line))); + ("col", (FStar_Json.JsonInt (r.col)))] +let (json_of_rng : rng -> FStar_Json.json) = + fun r -> + let uu___ = + let uu___1 = + let uu___2 = + let uu___3 = json_of_pos r.start_pos in ("start_pos", uu___3) in + let uu___3 = + let uu___4 = + let uu___5 = json_of_pos r.end_pos in ("end_pos", uu___5) in + [uu___4] in + uu___2 :: uu___3 in + ("file_name", (FStar_Json.JsonStr (r.file_name))) :: uu___1 in + FStar_Json.JsonAssoc uu___ +let (json_of_range : range -> FStar_Json.json) = + fun r -> + let uu___ = + let uu___1 = let uu___2 = json_of_rng r.def_range in ("def", uu___2) in + let uu___2 = + let uu___3 = let uu___4 = json_of_rng r.use_range in ("use", uu___4) in + [uu___3] in + uu___1 :: uu___2 in + FStar_Json.JsonAssoc uu___ \ No newline at end of file diff --git a/ocaml/fstar-lib/generated/FStar_Errors.ml b/ocaml/fstar-lib/generated/FStar_Errors.ml index 7d229f3e368..353698f060a 100644 --- a/ocaml/fstar-lib/generated/FStar_Errors.ml +++ b/ocaml/fstar-lib/generated/FStar_Errors.ml @@ -176,6 +176,32 @@ let (uu___is_EWarning : issue_level -> Prims.bool) = fun projectee -> match projectee with | EWarning -> true | uu___ -> false let (uu___is_EError : issue_level -> Prims.bool) = fun projectee -> match projectee with | EError -> true | uu___ -> false +exception Error of error +let (uu___is_Error : Prims.exn -> Prims.bool) = + fun projectee -> + match projectee with | Error uu___ -> true | uu___ -> false +let (__proj__Error__item__uu___ : Prims.exn -> error) = + fun projectee -> match projectee with | Error uu___ -> uu___ +exception Warning of error +let (uu___is_Warning : Prims.exn -> Prims.bool) = + fun projectee -> + match projectee with | Warning uu___ -> true | uu___ -> false +let (__proj__Warning__item__uu___ : Prims.exn -> error) = + fun projectee -> match projectee with | Warning uu___ -> uu___ +exception Stop +let (uu___is_Stop : Prims.exn -> Prims.bool) = + fun projectee -> match projectee with | Stop -> true | uu___ -> false +exception Empty_frag +let (uu___is_Empty_frag : Prims.exn -> Prims.bool) = + fun projectee -> match projectee with | Empty_frag -> true | uu___ -> false +let (json_of_issue_level : issue_level -> FStar_Json.json) = + fun level -> + FStar_Json.JsonStr + (match level with + | ENotImplemented -> "NotImplemented" + | EInfo -> "Info" + | EWarning -> "Warning" + | EError -> "Error") type issue = { issue_msg: FStar_Errors_Msg.error_message ; @@ -211,6 +237,62 @@ let (__proj__Mkissue__item__issue_ctx : issue -> Prims.string Prims.list) = match projectee with | { issue_msg; issue_level = issue_level1; issue_range; issue_number; issue_ctx;_} -> issue_ctx +let (json_of_issue : issue -> FStar_Json.json) = + fun issue1 -> + let uu___ = + let uu___1 = + let uu___2 = FStar_Errors_Msg.json_of_error_message issue1.issue_msg in + ("msg", uu___2) in + let uu___2 = + let uu___3 = + let uu___4 = json_of_issue_level issue1.issue_level in + ("level", uu___4) in + let uu___4 = + let uu___5 = + let uu___6 = + let uu___7 = + Obj.magic + (FStar_Class_Monad.op_Less_Dollar_Greater + FStar_Class_Monad.monad_option () () + (fun uu___8 -> + (Obj.magic FStar_Compiler_Range_Type.json_of_range) + uu___8) (Obj.magic issue1.issue_range)) in + FStar_Compiler_Util.dflt FStar_Json.JsonNull uu___7 in + ("range", uu___6) in + let uu___6 = + let uu___7 = + let uu___8 = + let uu___9 = + Obj.magic + (FStar_Class_Monad.op_Less_Dollar_Greater + FStar_Class_Monad.monad_option () () + (fun uu___10 -> + (fun uu___10 -> + let uu___10 = Obj.magic uu___10 in + Obj.magic (FStar_Json.JsonInt uu___10)) uu___10) + (Obj.magic issue1.issue_number)) in + FStar_Compiler_Util.dflt FStar_Json.JsonNull uu___9 in + ("number", uu___8) in + let uu___8 = + let uu___9 = + let uu___10 = + let uu___11 = + Obj.magic + (FStar_Class_Monad.op_Less_Dollar_Greater + FStar_Class_Monad.monad_list () () + (fun uu___12 -> + (fun uu___12 -> + let uu___12 = Obj.magic uu___12 in + Obj.magic (FStar_Json.JsonStr uu___12)) + uu___12) (Obj.magic issue1.issue_ctx)) in + FStar_Json.JsonList uu___11 in + ("ctx", uu___10) in + [uu___9] in + uu___7 :: uu___8 in + uu___5 :: uu___6 in + uu___3 :: uu___4 in + uu___1 :: uu___2 in + FStar_Json.JsonAssoc uu___ type error_handler = { eh_name: Prims.string ; @@ -248,24 +330,6 @@ let (__proj__Mkerror_handler__item__eh_clear : error_handler -> unit -> unit) match projectee with | { eh_name; eh_add_one; eh_count_errors; eh_report; eh_clear;_} -> eh_clear -exception Error of error -let (uu___is_Error : Prims.exn -> Prims.bool) = - fun projectee -> - match projectee with | Error uu___ -> true | uu___ -> false -let (__proj__Error__item__uu___ : Prims.exn -> error) = - fun projectee -> match projectee with | Error uu___ -> uu___ -exception Warning of error -let (uu___is_Warning : Prims.exn -> Prims.bool) = - fun projectee -> - match projectee with | Warning uu___ -> true | uu___ -> false -let (__proj__Warning__item__uu___ : Prims.exn -> error) = - fun projectee -> match projectee with | Warning uu___ -> uu___ -exception Stop -let (uu___is_Stop : Prims.exn -> Prims.bool) = - fun projectee -> match projectee with | Stop -> true | uu___ -> false -exception Empty_frag -let (uu___is_Empty_frag : Prims.exn -> Prims.bool) = - fun projectee -> match projectee with | Empty_frag -> true | uu___ -> false let (ctx_doc : Prims.string Prims.list -> FStar_Pprint.document) = fun ctx -> let uu___ = FStar_Options.error_contexts () in @@ -420,7 +484,12 @@ let (format_issue' : Prims.bool -> issue -> Prims.string) = FStar_Errors_Msg.renderdoc doc let (format_issue : issue -> Prims.string) = fun issue1 -> format_issue' true issue1 -let (print_issue : issue -> unit) = +let (print_issue_json : issue -> unit) = + fun issue1 -> + let uu___ = + let uu___1 = json_of_issue issue1 in FStar_Json.string_of_json uu___1 in + FStar_Compiler_Util.print1_error "%s\n" uu___ +let (print_issue_rendered : issue -> unit) = fun issue1 -> let printer = match issue1.issue_level with @@ -433,6 +502,12 @@ let (print_issue : issue -> unit) = | ENotImplemented -> FStar_Compiler_Util.print_error in let uu___ = let uu___1 = format_issue issue1 in Prims.strcat uu___1 "\n" in printer uu___ +let (print_issue : issue -> unit) = + fun issue1 -> + let uu___ = FStar_Options.message_format () in + match uu___ with + | FStar_Options.Human -> print_issue_rendered issue1 + | FStar_Options.Json -> print_issue_json issue1 let (compare_issues : issue -> issue -> Prims.int) = fun i1 -> fun i2 -> @@ -703,7 +778,7 @@ let (set_option_warning_callback_range : FStar_Compiler_Range_Type.range FStar_Pervasives_Native.option -> unit) = fun ropt -> FStar_Options.set_option_warning_callback (warn_unsafe_options ropt) -let (uu___357 : +let (uu___369 : (((Prims.string -> FStar_Errors_Codes.error_setting Prims.list) -> unit) * (unit -> FStar_Errors_Codes.error_setting Prims.list))) = @@ -749,10 +824,10 @@ let (uu___357 : (set_callbacks, get_error_flags) let (t_set_parse_warn_error : (Prims.string -> FStar_Errors_Codes.error_setting Prims.list) -> unit) = - match uu___357 with + match uu___369 with | (t_set_parse_warn_error1, error_flags) -> t_set_parse_warn_error1 let (error_flags : unit -> FStar_Errors_Codes.error_setting Prims.list) = - match uu___357 with + match uu___369 with | (t_set_parse_warn_error1, error_flags1) -> error_flags1 let (set_parse_warn_error : (Prims.string -> FStar_Errors_Codes.error_setting Prims.list) -> unit) = diff --git a/ocaml/fstar-lib/generated/FStar_Errors_Msg.ml b/ocaml/fstar-lib/generated/FStar_Errors_Msg.ml index 9d3c87ba23a..a51739e963c 100644 --- a/ocaml/fstar-lib/generated/FStar_Errors_Msg.ml +++ b/ocaml/fstar-lib/generated/FStar_Errors_Msg.ml @@ -91,4 +91,12 @@ let (rendermsg : error_message -> Prims.string) = FStar_Compiler_List.map (fun d -> let uu___2 = FStar_Pprint.group d in subdoc uu___2) ds in FStar_Pprint.concat uu___1 in - renderdoc uu___ \ No newline at end of file + renderdoc uu___ +let (json_of_error_message : + FStar_Pprint.document Prims.list -> FStar_Json.json) = + fun err_msg -> + let uu___ = + FStar_Compiler_List.map + (fun doc -> let uu___1 = renderdoc doc in FStar_Json.JsonStr uu___1) + err_msg in + FStar_Json.JsonList uu___ \ No newline at end of file diff --git a/ocaml/fstar-lib/generated/FStar_Options.ml b/ocaml/fstar-lib/generated/FStar_Options.ml index 17e2cd563fe..80d441c7156 100644 --- a/ocaml/fstar-lib/generated/FStar_Options.ml +++ b/ocaml/fstar-lib/generated/FStar_Options.ml @@ -25,6 +25,13 @@ let (uu___is_OnFailure : split_queries_t -> Prims.bool) = fun projectee -> match projectee with | OnFailure -> true | uu___ -> false let (uu___is_Always : split_queries_t -> Prims.bool) = fun projectee -> match projectee with | Always -> true | uu___ -> false +type message_format_t = + | Json + | Human +let (uu___is_Json : message_format_t -> Prims.bool) = + fun projectee -> match projectee with | Json -> true | uu___ -> false +let (uu___is_Human : message_format_t -> Prims.bool) = + fun projectee -> match projectee with | Human -> true | uu___ -> false type option_val = | Bool of Prims.bool | String of Prims.string @@ -291,6 +298,7 @@ let (defaults : (Prims.string * option_val) Prims.list) = ("eager_subtyping", (Bool false)); ("error_contexts", (Bool false)); ("expose_interfaces", (Bool false)); + ("message_format", (String "human")); ("ext", Unset); ("extract", Unset); ("extract_all", (Bool false)); @@ -630,6 +638,8 @@ let (get_error_contexts : unit -> Prims.bool) = fun uu___ -> lookup_opt "error_contexts" as_bool let (get_expose_interfaces : unit -> Prims.bool) = fun uu___ -> lookup_opt "expose_interfaces" as_bool +let (get_message_format : unit -> Prims.string) = + fun uu___ -> lookup_opt "message_format" as_string let (get_extract : unit -> Prims.string Prims.list FStar_Pervasives_Native.option) = fun uu___ -> lookup_opt "extract" (as_option (as_list as_string)) @@ -1127,7 +1137,7 @@ let (interp_quake_arg : Prims.string -> (Prims.int * Prims.int * Prims.bool)) let uu___ = ios f1 in let uu___1 = ios f2 in (uu___, uu___1, true) else FStar_Compiler_Effect.failwith "unexpected value for --quake" | uu___ -> FStar_Compiler_Effect.failwith "unexpected value for --quake" -let (uu___528 : (((Prims.string -> unit) -> unit) * (Prims.string -> unit))) +let (uu___529 : (((Prims.string -> unit) -> unit) * (Prims.string -> unit))) = let cb = FStar_Compiler_Util.mk_ref FStar_Pervasives_Native.None in let set1 f = @@ -1139,11 +1149,11 @@ let (uu___528 : (((Prims.string -> unit) -> unit) * (Prims.string -> unit))) | FStar_Pervasives_Native.Some f -> f msg in (set1, call) let (set_option_warning_callback_aux : (Prims.string -> unit) -> unit) = - match uu___528 with + match uu___529 with | (set_option_warning_callback_aux1, option_warning_callback) -> set_option_warning_callback_aux1 let (option_warning_callback : Prims.string -> unit) = - match uu___528 with + match uu___529 with | (set_option_warning_callback_aux1, option_warning_callback1) -> option_warning_callback1 let (set_option_warning_callback : (Prims.string -> unit) -> unit) = @@ -1535,26 +1545,25 @@ let rec (specs_with_types : let uu___58 = let uu___59 = text - "Don't print unification variable numbers" in + "Format of the messages emitted by F* (default `human`)" in (FStar_Getopt.noshort, - "hide_uvar_nums", - (Const - (Bool - true)), + "message_format", + (EnumStr + ["human"; + "json"]), uu___59) in let uu___59 = let uu___60 = let uu___61 = text - "Read/write hints to dir/module_name.hints (instead of placing hint-file alongside source file)" in + "Don't print unification variable numbers" in (FStar_Getopt.noshort, - "hint_dir", + "hide_uvar_nums", ( - PostProcessed - (pp_validate_dir, - (PathStr - "dir"))), + Const + (Bool + true)), uu___61) in let uu___61 = let uu___62 @@ -1562,11 +1571,13 @@ let rec (specs_with_types : let uu___63 = text - "Read/write hints to path (instead of module-specific hints files; overrides hint_dir)" in + "Read/write hints to dir/module_name.hints (instead of placing hint-file alongside source file)" in (FStar_Getopt.noshort, - "hint_file", + "hint_dir", + (PostProcessed + (pp_validate_dir, (PathStr - "path"), + "dir"))), uu___63) in let uu___63 = @@ -1575,12 +1586,11 @@ let rec (specs_with_types : let uu___65 = text - "Print information regarding hints (deprecated; use --query_stats instead)" in + "Read/write hints to path (instead of module-specific hints files; overrides hint_dir)" in (FStar_Getopt.noshort, - "hint_info", - (Const - (Bool - true)), + "hint_file", + (PathStr + "path"), uu___65) in let uu___65 = @@ -1589,9 +1599,9 @@ let rec (specs_with_types : let uu___67 = text - "Legacy interactive mode; reads input from stdin" in + "Print information regarding hints (deprecated; use --query_stats instead)" in (FStar_Getopt.noshort, - "in", + "hint_info", (Const (Bool true)), @@ -1603,9 +1613,9 @@ let rec (specs_with_types : let uu___69 = text - "JSON-based interactive mode for IDEs" in + "Legacy interactive mode; reads input from stdin" in (FStar_Getopt.noshort, - "ide", + "in", (Const (Bool true)), @@ -1617,9 +1627,9 @@ let rec (specs_with_types : let uu___71 = text - "Disable identifier tables in IDE mode (temporary workaround useful in Steel)" in + "JSON-based interactive mode for IDEs" in (FStar_Getopt.noshort, - "ide_id_info_off", + "ide", (Const (Bool true)), @@ -1631,9 +1641,9 @@ let rec (specs_with_types : let uu___73 = text - "Language Server Protocol-based interactive mode for IDEs" in + "Disable identifier tables in IDE mode (temporary workaround useful in Steel)" in (FStar_Getopt.noshort, - "lsp", + "ide_id_info_off", (Const (Bool true)), @@ -1645,12 +1655,12 @@ let rec (specs_with_types : let uu___75 = text - "A directory in which to search for files included on the command line" in + "Language Server Protocol-based interactive mode for IDEs" in (FStar_Getopt.noshort, - "include", - (ReverseAccumulated - (PathStr - "path")), + "lsp", + (Const + (Bool + true)), uu___75) in let uu___75 = @@ -1659,12 +1669,12 @@ let rec (specs_with_types : let uu___77 = text - "Parses and prettyprints the files included on the command line" in + "A directory in which to search for files included on the command line" in (FStar_Getopt.noshort, - "print", - (Const - (Bool - true)), + "include", + (ReverseAccumulated + (PathStr + "path")), uu___77) in let uu___77 = @@ -1673,9 +1683,9 @@ let rec (specs_with_types : let uu___79 = text - "Parses and prettyprints in place the files included on the command line" in + "Parses and prettyprints the files included on the command line" in (FStar_Getopt.noshort, - "print_in_place", + "print", (Const (Bool true)), @@ -1687,9 +1697,9 @@ let rec (specs_with_types : let uu___81 = text - "Force checking the files given as arguments even if they have valid checked files" in - (102, - "force", + "Parses and prettyprints in place the files included on the command line" in + (FStar_Getopt.noshort, + "print_in_place", (Const (Bool true)), @@ -1701,26 +1711,40 @@ let rec (specs_with_types : let uu___83 = text + "Force checking the files given as arguments even if they have valid checked files" in + (102, + "force", + (Const + (Bool + true)), + uu___83) in + let uu___83 + = + let uu___84 + = + let uu___85 + = + text "Set initial_fuel and max_fuel at once" in (FStar_Getopt.noshort, "fuel", (PostProcessed ((fun - uu___84 + uu___86 -> - match uu___84 + match uu___86 with | String s -> let p f = - let uu___85 + let uu___87 = FStar_Compiler_Util.int_of_string f in Int - uu___85 in - let uu___85 + uu___87 in + let uu___87 = match FStar_Compiler_Util.split @@ -1734,40 +1758,40 @@ let rec (specs_with_types : -> (f1, f2) | - uu___86 + uu___88 -> FStar_Compiler_Effect.failwith "unexpected value for --fuel" in - (match uu___85 + (match uu___87 with | (min, max) -> (( - let uu___87 + let uu___89 = p min in set_option "initial_fuel" - uu___87); - (let uu___88 + uu___89); + (let uu___90 = p max in set_option "max_fuel" - uu___88); + uu___90); String s)) | - uu___85 + uu___87 -> FStar_Compiler_Effect.failwith "impos"), (SimpleStr "non-negative integer or pair of non-negative integers"))), - uu___83) in - let uu___83 + uu___85) in + let uu___85 = - let uu___84 + let uu___86 = - let uu___85 + let uu___87 = text "Set initial_ifuel and max_ifuel at once" in @@ -1775,21 +1799,21 @@ let rec (specs_with_types : "ifuel", (PostProcessed ((fun - uu___86 + uu___88 -> - match uu___86 + match uu___88 with | String s -> let p f = - let uu___87 + let uu___89 = FStar_Compiler_Util.int_of_string f in Int - uu___87 in - let uu___87 + uu___89 in + let uu___89 = match FStar_Compiler_Util.split @@ -1803,40 +1827,40 @@ let rec (specs_with_types : -> (f1, f2) | - uu___88 + uu___90 -> FStar_Compiler_Effect.failwith "unexpected value for --ifuel" in - (match uu___87 + (match uu___89 with | (min, max) -> (( - let uu___89 + let uu___91 = p min in set_option "initial_ifuel" - uu___89); - (let uu___90 + uu___91); + (let uu___92 = p max in set_option "max_ifuel" - uu___90); + uu___92); String s)) | - uu___87 + uu___89 -> FStar_Compiler_Effect.failwith "impos"), (SimpleStr "non-negative integer or pair of non-negative integers"))), - uu___85) in - let uu___85 + uu___87) in + let uu___87 = - let uu___86 + let uu___88 = - let uu___87 + let uu___89 = text "Number of unrolling of recursive functions to try initially (default 2)" in @@ -1844,12 +1868,12 @@ let rec (specs_with_types : "initial_fuel", (IntStr "non-negative integer"), - uu___87) in - let uu___87 + uu___89) in + let uu___89 = - let uu___88 + let uu___90 = - let uu___89 + let uu___91 = text "Number of unrolling of inductive datatypes to try at first (default 1)" in @@ -1857,24 +1881,24 @@ let rec (specs_with_types : "initial_ifuel", (IntStr "non-negative integer"), - uu___89) in - let uu___89 + uu___91) in + let uu___91 = - let uu___90 + let uu___92 = - let uu___91 + let uu___93 = text "Retain comments in the logged SMT queries (requires --log_queries or --log_failing_queries; default true)" in (FStar_Getopt.noshort, "keep_query_captions", BoolStr, - uu___91) in - let uu___91 + uu___93) in + let uu___93 = - let uu___92 + let uu___94 = - let uu___93 + let uu___95 = text "Run the lax-type checker only (admit all verification conditions)" in @@ -1882,7 +1906,7 @@ let rec (specs_with_types : "lax", (WithSideEffect ((fun - uu___94 + uu___96 -> if warn_unsafe @@ -1893,20 +1917,6 @@ let rec (specs_with_types : (Const (Bool true)))), - uu___93) in - let uu___93 - = - let uu___94 - = - let uu___95 - = - text - "Load OCaml module, compiling it if necessary" in - (FStar_Getopt.noshort, - "load", - (ReverseAccumulated - (PathStr - "module")), uu___95) in let uu___95 = @@ -1915,9 +1925,9 @@ let rec (specs_with_types : let uu___97 = text - "Load compiled module, fails hard if the module is not already compiled" in + "Load OCaml module, compiling it if necessary" in (FStar_Getopt.noshort, - "load_cmxs", + "load", (ReverseAccumulated (PathStr "module")), @@ -1929,12 +1939,12 @@ let rec (specs_with_types : let uu___99 = text - "Print types computed for data/val/let-bindings" in + "Load compiled module, fails hard if the module is not already compiled" in (FStar_Getopt.noshort, - "log_types", - (Const - (Bool - true)), + "load_cmxs", + (ReverseAccumulated + (PathStr + "module")), uu___99) in let uu___99 = @@ -1943,9 +1953,9 @@ let rec (specs_with_types : let uu___101 = text - "Log the Z3 queries in several queries-*.smt2 files, as we go" in + "Print types computed for data/val/let-bindings" in (FStar_Getopt.noshort, - "log_queries", + "log_types", (Const (Bool true)), @@ -1957,9 +1967,9 @@ let rec (specs_with_types : let uu___103 = text - "As --log_queries, but only save the failing queries. Each query is\n saved in its own file regardless of whether they were checked during the\n same invocation. The SMT2 file names begin with \"failedQueries\"" in + "Log the Z3 queries in several queries-*.smt2 files, as we go" in (FStar_Getopt.noshort, - "log_failing_queries", + "log_queries", (Const (Bool true)), @@ -1971,11 +1981,12 @@ let rec (specs_with_types : let uu___105 = text - "Number of unrolling of recursive functions to try at most (default 8)" in + "As --log_queries, but only save the failing queries. Each query is\n saved in its own file regardless of whether they were checked during the\n same invocation. The SMT2 file names begin with \"failedQueries\"" in (FStar_Getopt.noshort, - "max_fuel", - (IntStr - "non-negative integer"), + "log_failing_queries", + (Const + (Bool + true)), uu___105) in let uu___105 = @@ -1984,9 +1995,9 @@ let rec (specs_with_types : let uu___107 = text - "Number of unrolling of inductive datatypes to try at most (default 2)" in + "Number of unrolling of recursive functions to try at most (default 8)" in (FStar_Getopt.noshort, - "max_ifuel", + "max_fuel", (IntStr "non-negative integer"), uu___107) in @@ -1997,12 +2008,11 @@ let rec (specs_with_types : let uu___109 = text - "Trigger various specializations for compiling the F* compiler itself (not meant for user code)" in + "Number of unrolling of inductive datatypes to try at most (default 2)" in (FStar_Getopt.noshort, - "MLish", - (Const - (Bool - true)), + "max_ifuel", + (IntStr + "non-negative integer"), uu___109) in let uu___109 = @@ -2011,9 +2021,9 @@ let rec (specs_with_types : let uu___111 = text - "Ignore the default module search paths" in + "Trigger various specializations for compiling the F* compiler itself (not meant for user code)" in (FStar_Getopt.noshort, - "no_default_includes", + "MLish", (Const (Bool true)), @@ -2025,12 +2035,12 @@ let rec (specs_with_types : let uu___113 = text - "Deprecated: use --extract instead; Do not extract code from this module" in + "Ignore the default module search paths" in (FStar_Getopt.noshort, - "no_extract", - (Accumulated - (PathStr - "module name")), + "no_default_includes", + (Const + (Bool + true)), uu___113) in let uu___113 = @@ -2039,12 +2049,12 @@ let rec (specs_with_types : let uu___115 = text - "Suppress location information in the generated OCaml output (only relevant with --codegen OCaml)" in + "Deprecated: use --extract instead; Do not extract code from this module" in (FStar_Getopt.noshort, - "no_location_info", - (Const - (Bool - true)), + "no_extract", + (Accumulated + (PathStr + "module name")), uu___115) in let uu___115 = @@ -2053,9 +2063,9 @@ let rec (specs_with_types : let uu___117 = text - "Do not send any queries to the SMT solver, and fail on them instead" in + "Suppress location information in the generated OCaml output (only relevant with --codegen OCaml)" in (FStar_Getopt.noshort, - "no_smt", + "no_location_info", (Const (Bool true)), @@ -2067,9 +2077,9 @@ let rec (specs_with_types : let uu___119 = text - "Extract top-level pure terms after normalizing them. This can lead to very large code, but can result in more partial evaluation and compile-time specialization." in + "Do not send any queries to the SMT solver, and fail on them instead" in (FStar_Getopt.noshort, - "normalize_pure_terms_for_extraction", + "no_smt", (Const (Bool true)), @@ -2081,11 +2091,12 @@ let rec (specs_with_types : let uu___121 = text - "Place KaRaMeL extraction output in file . The path can be relative or absolute and does not dependon the --odir option." in + "Extract top-level pure terms after normalizing them. This can lead to very large code, but can result in more partial evaluation and compile-time specialization." in (FStar_Getopt.noshort, - "krmloutput", - (PathStr - "filename"), + "normalize_pure_terms_for_extraction", + (Const + (Bool + true)), uu___121) in let uu___121 = @@ -2094,13 +2105,11 @@ let rec (specs_with_types : let uu___123 = text - "Place output in directory dir" in + "Place KaRaMeL extraction output in file . The path can be relative or absolute and does not dependon the --odir option." in (FStar_Getopt.noshort, - "odir", - (PostProcessed - (pp_validate_dir, + "krmloutput", (PathStr - "dir"))), + "filename"), uu___123) in let uu___123 = @@ -2109,11 +2118,13 @@ let rec (specs_with_types : let uu___125 = text - "Output the result of --dep into this file instead of to standard output." in + "Place output in directory dir" in (FStar_Getopt.noshort, - "output_deps_to", + "odir", + (PostProcessed + (pp_validate_dir, (PathStr - "file"), + "dir"))), uu___125) in let uu___125 = @@ -2122,9 +2133,9 @@ let rec (specs_with_types : let uu___127 = text - "Use a custom Prims.fst file. Do not use if you do not know exactly what you're doing." in + "Output the result of --dep into this file instead of to standard output." in (FStar_Getopt.noshort, - "prims", + "output_deps_to", (PathStr "file"), uu___127) in @@ -2135,12 +2146,11 @@ let rec (specs_with_types : let uu___129 = text - "Print the types of bound variables" in + "Use a custom Prims.fst file. Do not use if you do not know exactly what you're doing." in (FStar_Getopt.noshort, - "print_bound_var_types", - (Const - (Bool - true)), + "prims", + (PathStr + "file"), uu___129) in let uu___129 = @@ -2149,9 +2159,9 @@ let rec (specs_with_types : let uu___131 = text - "Print inferred predicate transformers for all computation types" in + "Print the types of bound variables" in (FStar_Getopt.noshort, - "print_effect_args", + "print_bound_var_types", (Const (Bool true)), @@ -2163,9 +2173,9 @@ let rec (specs_with_types : let uu___133 = text - "Print the errors generated by declarations marked with expect_failure, useful for debugging error locations" in + "Print inferred predicate transformers for all computation types" in (FStar_Getopt.noshort, - "print_expected_failures", + "print_effect_args", (Const (Bool true)), @@ -2177,9 +2187,9 @@ let rec (specs_with_types : let uu___135 = text - "Print full names of variables" in + "Print the errors generated by declarations marked with expect_failure, useful for debugging error locations" in (FStar_Getopt.noshort, - "print_full_names", + "print_expected_failures", (Const (Bool true)), @@ -2191,9 +2201,9 @@ let rec (specs_with_types : let uu___137 = text - "Print implicit arguments" in + "Print full names of variables" in (FStar_Getopt.noshort, - "print_implicits", + "print_full_names", (Const (Bool true)), @@ -2205,9 +2215,9 @@ let rec (specs_with_types : let uu___139 = text - "Print universes" in + "Print implicit arguments" in (FStar_Getopt.noshort, - "print_universes", + "print_implicits", (Const (Bool true)), @@ -2219,9 +2229,9 @@ let rec (specs_with_types : let uu___141 = text - "Print Z3 statistics for each SMT query (details such as relevant modules, facts, etc. for each proof)" in + "Print universes" in (FStar_Getopt.noshort, - "print_z3_statistics", + "print_universes", (Const (Bool true)), @@ -2233,9 +2243,9 @@ let rec (specs_with_types : let uu___143 = text - "Print full names (deprecated; use --print_full_names instead)" in + "Print Z3 statistics for each SMT query (details such as relevant modules, facts, etc. for each proof)" in (FStar_Getopt.noshort, - "prn", + "print_z3_statistics", (Const (Bool true)), @@ -2247,9 +2257,9 @@ let rec (specs_with_types : let uu___145 = text - "Proof recovery mode: before failing an SMT query, retry 3 times, increasing rlimits. If the query goes through after retrying, verification will succeed, but a warning will be emitted. This feature is useful to restore a project after some change to its libraries or F* upgrade. Importantly, then, this option cannot be used in a pragma (#set-options, etc)." in + "Print full names (deprecated; use --print_full_names instead)" in (FStar_Getopt.noshort, - "proof_recovery", + "prn", (Const (Bool true)), @@ -2260,76 +2270,90 @@ let rec (specs_with_types : = let uu___147 = + text + "Proof recovery mode: before failing an SMT query, retry 3 times, increasing rlimits. If the query goes through after retrying, verification will succeed, but a warning will be emitted. This feature is useful to restore a project after some change to its libraries or F* upgrade. Importantly, then, this option cannot be used in a pragma (#set-options, etc)." in + (FStar_Getopt.noshort, + "proof_recovery", + (Const + (Bool + true)), + uu___147) in + let uu___147 + = let uu___148 = - text - "Repeats SMT queries to check for robustness" in let uu___149 = let uu___150 = + text + "Repeats SMT queries to check for robustness" in let uu___152 = let uu___153 = - text - "--quake N/M repeats each query checks that it succeeds at least N out of M times, aborting early if possible" in let uu___154 = let uu___155 = text - "--quake N/M/k works as above, except it will unconditionally run M times" in + "--quake N/M repeats each query checks that it succeeds at least N out of M times, aborting early if possible" in let uu___156 = let uu___157 = text - "--quake N is an alias for --quake N/N" in + "--quake N/M/k works as above, except it will unconditionally run M times" in let uu___158 = let uu___159 = text + "--quake N is an alias for --quake N/N" in + let uu___160 + = + let uu___161 + = + text "--quake N/k is an alias for --quake N/N/k" in - [uu___159] in + [uu___161] in + uu___159 + :: + uu___160 in uu___157 :: uu___158 in uu___155 :: uu___156 in - uu___153 - :: - uu___154 in FStar_Errors_Msg.bulleted - uu___152 in - let uu___152 + uu___154 in + let uu___154 = text "Using --quake disables --retry. When quake testing, queries are not splitted for error reporting unless '--split_queries always' is given. Queries from the smt_sync tactic are not quake-tested." in FStar_Pprint.op_Hat_Hat + uu___153 + uu___154 in + FStar_Pprint.op_Hat_Hat uu___150 uu___152 in - FStar_Pprint.op_Hat_Hat - uu___148 - uu___149 in (FStar_Getopt.noshort, "quake", (PostProcessed ((fun - uu___148 + uu___150 -> - match uu___148 + match uu___150 with | String s -> - let uu___149 + let uu___152 = interp_quake_arg s in - (match uu___149 + (match uu___152 with | (min, @@ -2350,18 +2374,18 @@ let rec (specs_with_types : false); String s)) | - uu___149 + uu___152 -> FStar_Compiler_Effect.failwith "impos"), (SimpleStr "positive integer or pair of positive integers"))), - uu___147) in - let uu___147 + uu___149) in + let uu___149 = - let uu___148 + let uu___150 = - let uu___149 + let uu___152 = text "Keep a running cache of SMT queries to make verification faster. Only available in the interactive mode. NOTE: This feature is experimental and potentially unsound! Hence why\n it is not allowed in batch mode (where it is also less useful). If you\n find a query that is mistakenly accepted with the cache, please\n report a bug to the F* issue tracker on GitHub." in @@ -2370,12 +2394,12 @@ let rec (specs_with_types : (Const (Bool true)), - uu___149) in - let uu___149 + uu___152) in + let uu___152 = - let uu___150 + let uu___153 = - let uu___152 + let uu___154 = text "Print SMT query statistics" in @@ -2384,12 +2408,12 @@ let rec (specs_with_types : (Const (Bool true)), - uu___152) in - let uu___152 + uu___154) in + let uu___154 = - let uu___153 + let uu___155 = - let uu___154 + let uu___156 = text "Read a checked file and dump it to standard output." in @@ -2397,12 +2421,12 @@ let rec (specs_with_types : "read_checked_file", (PathStr "path"), - uu___154) in - let uu___154 + uu___156) in + let uu___156 = - let uu___155 + let uu___157 = - let uu___156 + let uu___158 = text "Read a Karamel binary file and dump it to standard output." in @@ -2410,12 +2434,12 @@ let rec (specs_with_types : "read_krml_file", (PathStr "path"), - uu___156) in - let uu___156 + uu___158) in + let uu___158 = - let uu___157 + let uu___159 = - let uu___158 + let uu___160 = text "Record a database of hints for efficient proof replay" in @@ -2424,12 +2448,12 @@ let rec (specs_with_types : (Const (Bool true)), - uu___158) in - let uu___158 + uu___160) in + let uu___160 = - let uu___159 + let uu___161 = - let uu___160 + let uu___162 = text "Record the state of options used to check each sigelt, useful for the `check_with` attribute and metaprogramming. Note that this implies a performance hit and increases the size of checked files." in @@ -2438,12 +2462,12 @@ let rec (specs_with_types : (Const (Bool true)), - uu___160) in - let uu___160 + uu___162) in + let uu___162 = - let uu___161 + let uu___163 = - let uu___162 + let uu___164 = text "Retry each SMT query N times and succeed on the first try. Using --retry disables --quake." in @@ -2451,9 +2475,9 @@ let rec (specs_with_types : "retry", (PostProcessed ((fun - uu___163 + uu___165 -> - match uu___163 + match uu___165 with | Int i -> @@ -2474,18 +2498,18 @@ let rec (specs_with_types : true); Bool true) | - uu___164 + uu___166 -> FStar_Compiler_Effect.failwith "impos"), (IntStr "positive integer"))), - uu___162) in - let uu___162 + uu___164) in + let uu___164 = - let uu___163 + let uu___165 = - let uu___164 + let uu___166 = text "Optimistically, attempt using the recorded hint for toplevel_name (a top-level name in the current module) when trying to verify some other term 'g'" in @@ -2493,12 +2517,12 @@ let rec (specs_with_types : "reuse_hint_for", (SimpleStr "toplevel_name"), - uu___164) in - let uu___164 + uu___166) in + let uu___166 = - let uu___165 + let uu___167 = - let uu___166 + let uu___168 = text "Report every use of an escape hatch, include assume, admit, etc." in @@ -2507,12 +2531,12 @@ let rec (specs_with_types : (EnumStr ["warn"; "error"]), - uu___166) in - let uu___166 + uu___168) in + let uu___168 = - let uu___167 + let uu___169 = - let uu___168 + let uu___170 = text "Disable all non-critical output" in @@ -2521,12 +2545,12 @@ let rec (specs_with_types : (Const (Bool true)), - uu___168) in - let uu___168 + uu___170) in + let uu___170 = - let uu___169 + let uu___171 = - let uu___170 + let uu___172 = text "Path to the Z3 SMT solver (we could eventually support other solvers)" in @@ -2534,211 +2558,197 @@ let rec (specs_with_types : "smt", (PathStr "path"), - uu___170) in - let uu___170 + uu___172) in + let uu___172 = - let uu___171 + let uu___173 = - let uu___172 + let uu___174 = text "Toggle a peephole optimization that eliminates redundant uses of boxing/unboxing in the SMT encoding (default 'false')" in (FStar_Getopt.noshort, "smtencoding.elim_box", BoolStr, - uu___172) in - let uu___172 - = - let uu___173 - = + uu___174) in let uu___174 = let uu___175 = - text - "Control the representation of non-linear arithmetic functions in the SMT encoding:" in let uu___176 = let uu___177 = + text + "Control the representation of non-linear arithmetic functions in the SMT encoding:" in let uu___178 = let uu___179 = - text - "if 'boxwrap' use 'Prims.op_Multiply, Prims.op_Division, Prims.op_Modulus'" in let uu___180 = let uu___181 = text - "if 'native' use '*, div, mod'" in + "if 'boxwrap' use 'Prims.op_Multiply, Prims.op_Division, Prims.op_Modulus'" in let uu___182 = let uu___183 = text + "if 'native' use '*, div, mod'" in + let uu___184 + = + let uu___185 + = + text "if 'wrapped' use '_mul, _div, _mod : Int*Int -> Int'" in - [uu___183] in + [uu___185] in + uu___183 + :: + uu___184 in uu___181 :: uu___182 in - uu___179 - :: - uu___180 in FStar_Errors_Msg.bulleted - uu___178 in - let uu___178 + uu___180 in + let uu___180 = text "(default 'boxwrap')" in FStar_Pprint.op_Hat_Hat + uu___179 + uu___180 in + FStar_Pprint.op_Hat_Hat uu___177 uu___178 in - FStar_Pprint.op_Hat_Hat - uu___175 - uu___176 in (FStar_Getopt.noshort, "smtencoding.nl_arith_repr", (EnumStr ["native"; "wrapped"; "boxwrap"]), - uu___174) in - let uu___174 - = - let uu___175 - = + uu___176) in let uu___176 = let uu___177 = - text - "Toggle the representation of linear arithmetic functions in the SMT encoding:" in let uu___178 = let uu___179 = + text + "Toggle the representation of linear arithmetic functions in the SMT encoding:" in let uu___180 = let uu___181 = - text - "if 'boxwrap', use 'Prims.op_Addition, Prims.op_Subtraction, Prims.op_Minus'" in let uu___182 = let uu___183 = text + "if 'boxwrap', use 'Prims.op_Addition, Prims.op_Subtraction, Prims.op_Minus'" in + let uu___184 + = + let uu___185 + = + text "if 'native', use '+, -, -'" in - [uu___183] in - uu___181 + [uu___185] in + uu___183 :: - uu___182 in + uu___184 in FStar_Errors_Msg.bulleted - uu___180 in - let uu___180 + uu___182 in + let uu___182 = text "(default 'boxwrap')" in FStar_Pprint.op_Hat_Hat + uu___181 + uu___182 in + FStar_Pprint.op_Hat_Hat uu___179 uu___180 in - FStar_Pprint.op_Hat_Hat - uu___177 - uu___178 in (FStar_Getopt.noshort, "smtencoding.l_arith_repr", (EnumStr ["native"; "boxwrap"]), - uu___176) in - let uu___176 + uu___178) in + let uu___178 = - let uu___177 + let uu___179 = - let uu___178 + let uu___180 = text "Include an axiom in the SMT encoding to introduce proof-irrelevance from a constructive proof" in (FStar_Getopt.noshort, "smtencoding.valid_intro", BoolStr, - uu___178) in - let uu___178 + uu___180) in + let uu___180 = - let uu___179 + let uu___181 = - let uu___180 + let uu___182 = text "Include an axiom in the SMT encoding to eliminate proof-irrelevance into the existence of a proof witness" in (FStar_Getopt.noshort, "smtencoding.valid_elim", BoolStr, - uu___180) in - let uu___180 - = - let uu___181 - = + uu___182) in let uu___182 = let uu___183 = - text - "Split SMT verification conditions into several separate queries, one per goal. Helps with localizing errors." in let uu___184 = let uu___185 = + text + "Split SMT verification conditions into several separate queries, one per goal. Helps with localizing errors." in let uu___186 = - text - "Use 'no' to disable (this may reduce the quality of error messages)." in let uu___187 = let uu___188 = text - "Use 'on_failure' to split queries and retry when discharging fails (the default)" in + "Use 'no' to disable (this may reduce the quality of error messages)." in let uu___189 = let uu___190 = text + "Use 'on_failure' to split queries and retry when discharging fails (the default)" in + let uu___191 + = + let uu___192 + = + text "Use 'yes' to always split." in - [uu___190] in + [uu___192] in + uu___190 + :: + uu___191 in uu___188 :: uu___189 in - uu___186 - :: - uu___187 in FStar_Errors_Msg.bulleted - uu___185 in + uu___187 in FStar_Pprint.op_Hat_Hat - uu___183 - uu___184 in + uu___185 + uu___186 in (FStar_Getopt.noshort, "split_queries", (EnumStr ["no"; "on_failure"; "always"]), - uu___182) in - let uu___182 - = - let uu___183 - = - let uu___184 - = - text - "Do not use the lexical scope of tactics to improve binder names" in - (FStar_Getopt.noshort, - "tactic_raw_binders", - (Const - (Bool - true)), uu___184) in let uu___184 = @@ -2747,9 +2757,9 @@ let rec (specs_with_types : let uu___186 = text - "Do not recover from metaprogramming errors, and abort if one occurs" in + "Do not use the lexical scope of tactics to improve binder names" in (FStar_Getopt.noshort, - "tactics_failhard", + "tactic_raw_binders", (Const (Bool true)), @@ -2761,9 +2771,9 @@ let rec (specs_with_types : let uu___188 = text - "Print some rough information on tactics, such as the time they take to run" in + "Do not recover from metaprogramming errors, and abort if one occurs" in (FStar_Getopt.noshort, - "tactics_info", + "tactics_failhard", (Const (Bool true)), @@ -2775,9 +2785,9 @@ let rec (specs_with_types : let uu___190 = text - "Print a depth-indexed trace of tactic execution (Warning: very verbose)" in + "Print some rough information on tactics, such as the time they take to run" in (FStar_Getopt.noshort, - "tactic_trace", + "tactics_info", (Const (Bool true)), @@ -2789,11 +2799,12 @@ let rec (specs_with_types : let uu___192 = text - "Trace tactics up to a certain binding depth" in + "Print a depth-indexed trace of tactic execution (Warning: very verbose)" in (FStar_Getopt.noshort, - "tactic_trace_d", - (IntStr - "positive_integer"), + "tactic_trace", + (Const + (Bool + true)), uu___192) in let uu___192 = @@ -2802,12 +2813,11 @@ let rec (specs_with_types : let uu___194 = text - "Use NBE to evaluate metaprograms (experimental)" in + "Trace tactics up to a certain binding depth" in (FStar_Getopt.noshort, - "__tactics_nbe", - (Const - (Bool - true)), + "tactic_trace_d", + (IntStr + "positive_integer"), uu___194) in let uu___194 = @@ -2816,10 +2826,12 @@ let rec (specs_with_types : let uu___196 = text - "Attempt to normalize definitions marked as tcnorm (default 'true')" in + "Use NBE to evaluate metaprograms (experimental)" in (FStar_Getopt.noshort, - "tcnorm", - BoolStr, + "__tactics_nbe", + (Const + (Bool + true)), uu___196) in let uu___196 = @@ -2828,12 +2840,10 @@ let rec (specs_with_types : let uu___198 = text - "Print the time it takes to verify each top-level definition. This is just an alias for an invocation of the profiler, so it may not work well if combined with --profile. In particular, it implies --profile_group_by_decl." in + "Attempt to normalize definitions marked as tcnorm (default 'true')" in (FStar_Getopt.noshort, - "timing", - (Const - (Bool - true)), + "tcnorm", + BoolStr, uu___198) in let uu___198 = @@ -2842,9 +2852,9 @@ let rec (specs_with_types : let uu___200 = text - "Attach stack traces on errors" in + "Print the time it takes to verify each top-level definition. This is just an alias for an invocation of the profiler, so it may not work well if combined with --profile. In particular, it implies --profile_group_by_decl." in (FStar_Getopt.noshort, - "trace_error", + "timing", (Const (Bool true)), @@ -2856,9 +2866,9 @@ let rec (specs_with_types : let uu___202 = text - "Emit output formatted for debugging" in + "Attach stack traces on errors" in (FStar_Getopt.noshort, - "ugly", + "trace_error", (Const (Bool true)), @@ -2870,9 +2880,9 @@ let rec (specs_with_types : let uu___204 = text - "Let the SMT solver unfold inductive types to arbitrary depths (may affect verifier performance)" in + "Emit output formatted for debugging" in (FStar_Getopt.noshort, - "unthrottle_inductives", + "ugly", (Const (Bool true)), @@ -2884,9 +2894,9 @@ let rec (specs_with_types : let uu___206 = text - "Allow tactics to run external processes. WARNING: checking an untrusted F* file while using this option can have disastrous effects." in + "Let the SMT solver unfold inductive types to arbitrary depths (may affect verifier performance)" in (FStar_Getopt.noshort, - "unsafe_tactic_exec", + "unthrottle_inductives", (Const (Bool true)), @@ -2898,9 +2908,9 @@ let rec (specs_with_types : let uu___208 = text - "Use equality constraints when comparing higher-order types (Temporary)" in + "Allow tactics to run external processes. WARNING: checking an untrusted F* file while using this option can have disastrous effects." in (FStar_Getopt.noshort, - "use_eq_at_higher_order", + "unsafe_tactic_exec", (Const (Bool true)), @@ -2912,18 +2922,32 @@ let rec (specs_with_types : let uu___210 = text - "Use a previously recorded hints database for proof replay" in + "Use equality constraints when comparing higher-order types (Temporary)" in (FStar_Getopt.noshort, - "use_hints", + "use_eq_at_higher_order", (Const (Bool true)), uu___210) in let uu___210 = - let uu___211 + let uu___211 + = + let uu___212 + = + text + "Use a previously recorded hints database for proof replay" in + (FStar_Getopt.noshort, + "use_hints", + (Const + (Bool + true)), + uu___212) in + let uu___212 + = + let uu___213 = - let uu___212 + let uu___214 = text "Admit queries if their hash matches the hash recorded in the hints database" in @@ -2932,12 +2956,12 @@ let rec (specs_with_types : (Const (Bool true)), - uu___212) in - let uu___212 + uu___214) in + let uu___214 = - let uu___213 + let uu___215 = - let uu___214 + let uu___216 = text "Use compiled tactics from path" in @@ -2945,12 +2969,12 @@ let rec (specs_with_types : "use_native_tactics", (PathStr "path"), - uu___214) in - let uu___214 + uu___216) in + let uu___216 = - let uu___215 + let uu___217 = - let uu___216 + let uu___218 = text "Do not run plugins natively and interpret them as usual instead" in @@ -2959,12 +2983,12 @@ let rec (specs_with_types : (Const (Bool true)), - uu___216) in - let uu___216 + uu___218) in + let uu___218 = - let uu___217 + let uu___219 = - let uu___218 + let uu___220 = text "Do not run the tactic engine before discharging a VC" in @@ -2973,12 +2997,12 @@ let rec (specs_with_types : (Const (Bool true)), - uu___218) in - let uu___218 + uu___220) in + let uu___220 = - let uu___219 + let uu___221 = - let uu___220 + let uu___222 = text "Prunes the context to include only the facts from the given namespace or fact id. Facts can be include or excluded using the [+|-] qualifier. For example --using_facts_from '* -FStar.Reflection +FStar.Compiler.List -FStar.Compiler.List.Tot' will remove all facts from FStar.Compiler.List.Tot.*, retain all remaining facts from FStar.Compiler.List.*, remove all facts from FStar.Reflection.*, and retain all the rest. Note, the '+' is optional: --using_facts_from 'FStar.Compiler.List' is equivalent to --using_facts_from '+FStar.Compiler.List'. Multiple uses of this option accumulate, e.g., --using_facts_from A --using_facts_from B is interpreted as --using_facts_from A^B." in @@ -2987,12 +3011,12 @@ let rec (specs_with_types : (ReverseAccumulated (SimpleStr "One or more space-separated occurrences of '[+|-]( * | namespace | fact id)'")), - uu___220) in - let uu___220 + uu___222) in + let uu___222 = - let uu___221 + let uu___223 = - let uu___222 + let uu___224 = text "This does nothing and will be removed" in @@ -3001,12 +3025,12 @@ let rec (specs_with_types : (Const (Bool true)), - uu___222) in - let uu___222 + uu___224) in + let uu___224 = - let uu___223 + let uu___225 = - let uu___224 + let uu___226 = text "Display version number" in @@ -3014,7 +3038,7 @@ let rec (specs_with_types : "version", (WithSideEffect ((fun - uu___225 + uu___227 -> display_version (); @@ -3023,12 +3047,12 @@ let rec (specs_with_types : (Const (Bool true)))), - uu___224) in - let uu___224 + uu___226) in + let uu___226 = - let uu___225 + let uu___227 = - let uu___226 + let uu___228 = text "Warn when (a -> b) is desugared to (a -> Tot b)" in @@ -3037,12 +3061,12 @@ let rec (specs_with_types : (Const (Bool true)), - uu___226) in - let uu___226 + uu___228) in + let uu___228 = - let uu___227 + let uu___229 = - let uu___228 + let uu___230 = text "Z3 command line options" in @@ -3051,12 +3075,12 @@ let rec (specs_with_types : (ReverseAccumulated (SimpleStr "option")), - uu___228) in - let uu___228 + uu___230) in + let uu___230 = - let uu___229 + let uu___231 = - let uu___230 + let uu___232 = text "Z3 options in smt2 format" in @@ -3065,12 +3089,12 @@ let rec (specs_with_types : (ReverseAccumulated (SimpleStr "option")), - uu___230) in - let uu___230 + uu___232) in + let uu___232 = - let uu___231 + let uu___233 = - let uu___232 + let uu___234 = text "Restart Z3 after each query; useful for ensuring proof robustness" in @@ -3079,12 +3103,12 @@ let rec (specs_with_types : (Const (Bool true)), - uu___232) in - let uu___232 + uu___234) in + let uu___234 = - let uu___233 + let uu___235 = - let uu___234 + let uu___236 = text "Set the Z3 per-query resource limit (default 5 units, taking roughtly 5s)" in @@ -3092,12 +3116,12 @@ let rec (specs_with_types : "z3rlimit", (IntStr "positive_integer"), - uu___234) in - let uu___234 + uu___236) in + let uu___236 = - let uu___235 + let uu___237 = - let uu___236 + let uu___238 = text "Set the Z3 per-query resource limit multiplier. This is useful when, say, regenerating hints and you want to be more lax. (default 1)" in @@ -3105,12 +3129,12 @@ let rec (specs_with_types : "z3rlimit_factor", (IntStr "positive_integer"), - uu___236) in - let uu___236 + uu___238) in + let uu___238 = - let uu___237 + let uu___239 = - let uu___238 + let uu___240 = text "Set the Z3 random seed (default 0)" in @@ -3118,12 +3142,12 @@ let rec (specs_with_types : "z3seed", (IntStr "positive_integer"), - uu___238) in - let uu___238 + uu___240) in + let uu___240 = - let uu___239 + let uu___241 = - let uu___240 + let uu___242 = text "Set the version of Z3 that is to be used. Default: 4.8.5" in @@ -3131,12 +3155,12 @@ let rec (specs_with_types : "z3version", (SimpleStr "version"), - uu___240) in - let uu___240 + uu___242) in + let uu___242 = - let uu___241 + let uu___243 = - let uu___242 + let uu___244 = text "Don't check positivity of inductive types" in @@ -3144,7 +3168,7 @@ let rec (specs_with_types : "__no_positivity", (WithSideEffect ((fun - uu___243 + uu___245 -> if warn_unsafe @@ -3155,75 +3179,63 @@ let rec (specs_with_types : (Const (Bool true)))), - uu___242) in - let uu___242 - = - let uu___243 - = + uu___244) in let uu___244 = let uu___245 = - text - "The [-warn_error] option follows the OCaml syntax, namely:" in let uu___246 = let uu___247 = + text + "The [-warn_error] option follows the OCaml syntax, namely:" in let uu___248 = - text - "[r] is a range of warnings (either a number [n], or a range [n..n])" in let uu___249 = let uu___250 = text - "[-r] silences range [r]" in + "[r] is a range of warnings (either a number [n], or a range [n..n])" in let uu___251 = let uu___252 = text - "[+r] enables range [r] as warnings (NOTE: \"enabling\" an error will downgrade it to a warning)" in + "[-r] silences range [r]" in let uu___253 = let uu___254 = text + "[+r] enables range [r] as warnings (NOTE: \"enabling\" an error will downgrade it to a warning)" in + let uu___255 + = + let uu___256 + = + text "[@r] makes range [r] fatal." in - [uu___254] in + [uu___256] in + uu___254 + :: + uu___255 in uu___252 :: uu___253 in uu___250 :: uu___251 in - uu___248 - :: - uu___249 in FStar_Errors_Msg.bulleted - uu___247 in + uu___249 in FStar_Pprint.op_Hat_Hat - uu___245 - uu___246 in + uu___247 + uu___248 in (FStar_Getopt.noshort, "warn_error", (ReverseAccumulated (SimpleStr "")), - uu___244) in - let uu___244 - = - let uu___245 - = - let uu___246 - = - text - "Use normalization by evaluation as the default normalization strategy (default 'false')" in - (FStar_Getopt.noshort, - "use_nbe", - BoolStr, uu___246) in let uu___246 = @@ -3232,9 +3244,9 @@ let rec (specs_with_types : let uu___248 = text - "Use normalization by evaluation for normalizing terms before extraction (default 'false')" in + "Use normalization by evaluation as the default normalization strategy (default 'false')" in (FStar_Getopt.noshort, - "use_nbe_for_extraction", + "use_nbe", BoolStr, uu___248) in let uu___248 @@ -3244,9 +3256,9 @@ let rec (specs_with_types : let uu___250 = text - "Enforce trivial preconditions for unannotated effectful functions (default 'true')" in + "Use normalization by evaluation for normalizing terms before extraction (default 'false')" in (FStar_Getopt.noshort, - "trivial_pre_for_unannotated_effectful_fns", + "use_nbe_for_extraction", BoolStr, uu___250) in let uu___250 @@ -3256,12 +3268,24 @@ let rec (specs_with_types : let uu___252 = text + "Enforce trivial preconditions for unannotated effectful functions (default 'true')" in + (FStar_Getopt.noshort, + "trivial_pre_for_unannotated_effectful_fns", + BoolStr, + uu___252) in + let uu___252 + = + let uu___253 + = + let uu___254 + = + text "Debug messages for embeddings/unembeddings of natively compiled terms" in (FStar_Getopt.noshort, "__debug_embedding", (WithSideEffect ((fun - uu___253 + uu___255 -> FStar_Compiler_Effect.op_Colon_Equals debug_embedding @@ -3269,12 +3293,12 @@ let rec (specs_with_types : (Const (Bool true)))), - uu___252) in - let uu___252 + uu___254) in + let uu___254 = - let uu___253 + let uu___255 = - let uu___254 + let uu___256 = text "Eagerly embed and unembed terms to primitive operations and plugins: not recommended except for benchmarking" in @@ -3282,7 +3306,7 @@ let rec (specs_with_types : "eager_embedding", (WithSideEffect ((fun - uu___255 + uu___257 -> FStar_Compiler_Effect.op_Colon_Equals eager_embedding @@ -3290,12 +3314,12 @@ let rec (specs_with_types : (Const (Bool true)))), - uu___254) in - let uu___254 + uu___256) in + let uu___256 = - let uu___255 + let uu___257 = - let uu___256 + let uu___258 = text "Emit profiles grouped by declaration rather than by module" in @@ -3304,12 +3328,12 @@ let rec (specs_with_types : (Const (Bool true)), - uu___256) in - let uu___256 + uu___258) in + let uu___258 = - let uu___257 + let uu___259 = - let uu___258 + let uu___260 = text "Specific source locations in the compiler are instrumented with profiling counters. Pass `--profile_component FStar.TypeChecker` to enable all counters in the FStar.TypeChecker namespace. This option is a module or namespace selector, like many other options (e.g., `--extract`)" in @@ -3318,12 +3342,12 @@ let rec (specs_with_types : (Accumulated (SimpleStr "One or more space-separated occurrences of '[+|-]( * | namespace | module | identifier)'")), - uu___258) in - let uu___258 + uu___260) in + let uu___260 = - let uu___259 + let uu___261 = - let uu___260 + let uu___262 = text "Profiling can be enabled when the compiler is processing a given set of source modules. Pass `--profile FStar.Pervasives` to enable profiling when the compiler is processing any module in FStar.Pervasives. This option is a module or namespace selector, like many other options (e.g., `--extract`)" in @@ -3332,12 +3356,12 @@ let rec (specs_with_types : (Accumulated (SimpleStr "One or more space-separated occurrences of '[+|-]( * | namespace | module)'")), - uu___260) in - let uu___260 + uu___262) in + let uu___262 = - let uu___261 + let uu___263 = - let uu___262 + let uu___264 = text "Display this information" in @@ -3345,26 +3369,26 @@ let rec (specs_with_types : "help", (WithSideEffect ((fun - uu___263 + uu___265 -> ( - let uu___265 + let uu___267 = specs warn_unsafe in display_usage_aux - uu___265); + uu___267); FStar_Compiler_Effect.exit Prims.int_zero), (Const (Bool true)))), - uu___262) in - let uu___262 + uu___264) in + let uu___264 = - let uu___263 + let uu___265 = - let uu___264 + let uu___266 = text "List all debug keys and exit" in @@ -3372,7 +3396,7 @@ let rec (specs_with_types : "list_debug_keys", (WithSideEffect ((fun - uu___265 + uu___267 -> display_debug_keys (); @@ -3381,8 +3405,11 @@ let rec (specs_with_types : (Const (Bool true)))), - uu___264) in - [uu___263] in + uu___266) in + [uu___265] in + uu___263 + :: + uu___264 in uu___261 :: uu___262 in @@ -3838,7 +3865,7 @@ let (settable_specs : (fun uu___ -> match uu___ with | ((uu___1, x, uu___2), uu___3) -> settable x) all_specs -let (uu___755 : +let (uu___756 : (((unit -> FStar_Getopt.parse_cmdline_res) -> unit) * (unit -> FStar_Getopt.parse_cmdline_res))) = @@ -3855,11 +3882,11 @@ let (uu___755 : (set1, call) let (set_error_flags_callback_aux : (unit -> FStar_Getopt.parse_cmdline_res) -> unit) = - match uu___755 with + match uu___756 with | (set_error_flags_callback_aux1, set_error_flags) -> set_error_flags_callback_aux1 let (set_error_flags : unit -> FStar_Getopt.parse_cmdline_res) = - match uu___755 with + match uu___756 with | (set_error_flags_callback_aux1, set_error_flags1) -> set_error_flags1 let (set_error_flags_callback : (unit -> FStar_Getopt.parse_cmdline_res) -> unit) = @@ -4273,6 +4300,21 @@ let (error_contexts : unit -> Prims.bool) = fun uu___ -> get_error_contexts () let (expose_interfaces : unit -> Prims.bool) = fun uu___ -> get_expose_interfaces () +let (message_format : unit -> message_format_t) = + fun uu___ -> + let uu___1 = get_message_format () in + match uu___1 with + | "human" -> Human + | "json" -> Json + | illegal -> + let uu___2 = + let uu___3 = + FStar_Compiler_String.op_Hat illegal + "`. This should be impossible: `message_format` was supposed to be validated." in + FStar_Compiler_String.op_Hat + "print_issue: option `message_format` was expected to be `human` or `json`, not `" + uu___3 in + FStar_Compiler_Effect.failwith uu___2 let (force : unit -> Prims.bool) = fun uu___ -> get_force () let (full_context_dependency : unit -> Prims.bool) = fun uu___ -> true let (hide_uvar_nums : unit -> Prims.bool) = diff --git a/ocaml/fstar-lib/generated/FStar_Profiling.ml b/ocaml/fstar-lib/generated/FStar_Profiling.ml index 57f5869e33d..102afbffe44 100644 --- a/ocaml/fstar-lib/generated/FStar_Profiling.ml +++ b/ocaml/fstar-lib/generated/FStar_Profiling.ml @@ -23,6 +23,32 @@ let (__proj__Mkcounter__item__undercount : fun projectee -> match projectee with | { cid; total_time; running; undercount;_} -> undercount +let (json_of_counter : counter -> FStar_Json.json) = + fun c -> + let uu___ = + let uu___1 = + let uu___2 = + let uu___3 = + let uu___4 = FStar_Compiler_Effect.op_Bang c.total_time in + FStar_Json.JsonInt uu___4 in + ("total_time", uu___3) in + let uu___3 = + let uu___4 = + let uu___5 = + let uu___6 = FStar_Compiler_Effect.op_Bang c.running in + FStar_Json.JsonBool uu___6 in + ("running", uu___5) in + let uu___5 = + let uu___6 = + let uu___7 = + let uu___8 = FStar_Compiler_Effect.op_Bang c.undercount in + FStar_Json.JsonBool uu___8 in + ("undercount", uu___7) in + [uu___6] in + uu___4 :: uu___5 in + uu___2 :: uu___3 in + ("id", (FStar_Json.JsonStr (c.cid))) :: uu___1 in + FStar_Json.JsonAssoc uu___ let (new_counter : Prims.string -> counter) = fun cid -> let uu___ = FStar_Compiler_Util.mk_ref Prims.int_zero in @@ -78,6 +104,40 @@ let profile : FStar_Compiler_Effect.op_Colon_Equals c.undercount true; FStar_Compiler_Effect.raise uu___3))) else f () +let (report_json : Prims.string -> counter -> unit) = + fun tag -> + fun c -> + let counter1 = json_of_counter c in + let uu___ = + FStar_Json.string_of_json + (FStar_Json.JsonAssoc + [("tag", (FStar_Json.JsonStr tag)); ("counter", counter1)]) in + FStar_Compiler_Util.print1_error "%s\n" uu___ +let (report_human : Prims.string -> counter -> unit) = + fun tag -> + fun c -> + let warn = + let uu___ = FStar_Compiler_Effect.op_Bang c.running in + if uu___ + then " (Warning, this counter is still running)" + else + (let uu___2 = FStar_Compiler_Effect.op_Bang c.undercount in + if uu___2 + then + " (Warning, some operations raised exceptions and we not accounted for)" + else "") in + let uu___ = + let uu___1 = FStar_Compiler_Effect.op_Bang c.total_time in + FStar_Compiler_Util.string_of_int uu___1 in + FStar_Compiler_Util.print4 "%s, profiled %s:\t %s ms%s\n" tag c.cid + uu___ warn +let (report : Prims.string -> counter -> unit) = + fun tag -> + fun c -> + let uu___ = FStar_Options.message_format () in + match uu___ with + | FStar_Options.Human -> report_human tag c + | FStar_Options.Json -> report_json tag c let (report_and_clear : Prims.string -> unit) = fun tag -> let ctrs = @@ -91,20 +151,4 @@ let (report_and_clear : Prims.string -> unit) = let uu___1 = FStar_Compiler_Effect.op_Bang c2.total_time in let uu___2 = FStar_Compiler_Effect.op_Bang c1.total_time in uu___1 - uu___2) ctrs in - FStar_Compiler_List.iter - (fun c -> - let warn = - let uu___1 = FStar_Compiler_Effect.op_Bang c.running in - if uu___1 - then " (Warning, this counter is still running)" - else - (let uu___3 = FStar_Compiler_Effect.op_Bang c.undercount in - if uu___3 - then - " (Warning, some operations raised exceptions and we not accounted for)" - else "") in - let uu___1 = - let uu___2 = FStar_Compiler_Effect.op_Bang c.total_time in - FStar_Compiler_Util.string_of_int uu___2 in - FStar_Compiler_Util.print4 "%s, profiled %s:\t %s ms%s\n" tag - c.cid uu___1 warn) ctrs1) \ No newline at end of file + FStar_Compiler_List.iter (report tag) ctrs1) \ No newline at end of file From 043421206adecc4d347dda94ca3bd29f02129839 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Guido=20Mart=C3=ADnez?= Date: Thu, 26 Sep 2024 07:01:03 -0700 Subject: [PATCH 5/5] Adding tests for --message_format json --- examples/Makefile.common | 8 + tests/error-messages/.gitignore | 3 + .../AQualMismatch.fst.json_expected | 7 + .../error-messages/AnotType.fst.json_expected | 8 + .../ArgsAndQuals.fst.json_expected | 7 + .../ArrowRanges.fst.json_expected | 8 + .../AssertNorm.fst.json_expected | 5 + .../error-messages/Asserts.fst.json_expected | 11 + .../BadEmptyRecord.fst.json_expected | 5 + tests/error-messages/Basic.fst.json_expected | 38 +++ .../error-messages/Bug1918.fst.json_expected | 5 + .../error-messages/Bug1988.fst.json_expected | 11 + .../error-messages/Bug1997.fst.json_expected | 295 ++++++++++++++++++ .../error-messages/Bug2010.fst.json_expected | 5 + .../error-messages/Bug2021.fst.json_expected | 20 ++ .../error-messages/Bug2245.fst.json_expected | 5 + .../error-messages/Bug2820.fst.json_expected | 21 ++ .../error-messages/Bug2899.fst.json_expected | 14 + .../error-messages/Bug2953.fst.json_expected | 5 + .../error-messages/Bug2980.fst.json_expected | 5 + .../error-messages/Bug2981.fst.json_expected | 5 + .../error-messages/Bug3099.fst.json_expected | 7 + .../error-messages/Bug3102.fst.json_expected | 20 ++ .../error-messages/Bug3145.fst.json_expected | 29 ++ .../error-messages/Bug3227.fst.json_expected | 76 +++++ .../error-messages/Bug3292.fst.json_expected | 83 +++++ tests/error-messages/Calc.fst.json_expected | 35 +++ .../error-messages/CalcImpl.fst.json_expected | 161 ++++++++++ .../Coercions.fst.json_expected | 23 ++ .../error-messages/DTuples.fst.json_expected | 143 +++++++++ .../DecreasesTypeWarning.fst.json_expected | 4 + .../DumpUvarsOf.fst.json_expected | 15 + .../error-messages/Erasable.fst.json_expected | 23 ++ .../ExpectFailure.fst.json_expected | 23 ++ .../GhostImplicits.fst.json_expected | 5 + tests/error-messages/IfCond.fst.json_expected | 5 + tests/error-messages/IfThen.fst.json_expected | 5 + .../Inference.fst.json_expected | 6 + tests/error-messages/Makefile | 55 +++- .../NegativeTests.BST.fst.json_expected | 11 + .../NegativeTests.Bug260.fst.json_expected | 7 + .../NegativeTests.False.fst.json_expected | 12 + .../NegativeTests.Heap.fst.json_expected | 5 + .../NegativeTests.Neg.fst.json_expected | 41 +++ ...NegativeTests.Positivity.fst.json_expected | 20 ++ .../NegativeTests.Set.fst.json_expected | 15 + ...iveTests.ShortCircuiting.fst.json_expected | 11 + ...egativeTests.Termination.fst.json_expected | 43 +++ ...iveTests.ZZImplicitFalse.fst.json_expected | 7 + .../OccursCheckOnArrows.fst.json_expected | 5 + .../OptionStack.fst.json_expected | 14 + .../error-messages/PatAnnot.fst.json_expected | 20 ++ .../PatCoerce.fst.json_expected | 7 + .../PatternMatch.fst.json_expected | 44 +++ .../QuickTest.fst.json_expected | 5 + .../QuickTestNBE.fst.json_expected | 5 + .../RecordFields.fst.json_expected | 11 + ...ResolveImplicitsErrorPos.fst.json_expected | 10 + .../SMTPatSymbols.fst.json_expected | 3 + tests/error-messages/SeqLit.fst.json_expected | 78 +++++ .../StableErr.fst.json_expected | 5 + .../StrictUnfolding.fst.json_expected | 25 ++ .../StringNormalization.fst.json_expected | 5 + ...FunctionalExtensionality.fst.json_expected | 14 + .../TestErrorLocations.fst.json_expected | 51 +++ .../TestHasEq.fst.json_expected | 11 + .../UnboundOp.fst.json_expected | 5 + tests/error-messages/Unit2.fst.json_expected | 5 + .../UnresolvedFields.fst.json_expected | 11 + .../WPExtensionality.fst.json_expected | 5 + 70 files changed, 1703 insertions(+), 17 deletions(-) create mode 100644 tests/error-messages/AQualMismatch.fst.json_expected create mode 100644 tests/error-messages/AnotType.fst.json_expected create mode 100644 tests/error-messages/ArgsAndQuals.fst.json_expected create mode 100644 tests/error-messages/ArrowRanges.fst.json_expected create mode 100644 tests/error-messages/AssertNorm.fst.json_expected create mode 100644 tests/error-messages/Asserts.fst.json_expected create mode 100644 tests/error-messages/BadEmptyRecord.fst.json_expected create mode 100644 tests/error-messages/Basic.fst.json_expected create mode 100644 tests/error-messages/Bug1918.fst.json_expected create mode 100644 tests/error-messages/Bug1988.fst.json_expected create mode 100644 tests/error-messages/Bug1997.fst.json_expected create mode 100644 tests/error-messages/Bug2010.fst.json_expected create mode 100644 tests/error-messages/Bug2021.fst.json_expected create mode 100644 tests/error-messages/Bug2245.fst.json_expected create mode 100644 tests/error-messages/Bug2820.fst.json_expected create mode 100644 tests/error-messages/Bug2899.fst.json_expected create mode 100644 tests/error-messages/Bug2953.fst.json_expected create mode 100644 tests/error-messages/Bug2980.fst.json_expected create mode 100644 tests/error-messages/Bug2981.fst.json_expected create mode 100644 tests/error-messages/Bug3099.fst.json_expected create mode 100644 tests/error-messages/Bug3102.fst.json_expected create mode 100644 tests/error-messages/Bug3145.fst.json_expected create mode 100644 tests/error-messages/Bug3227.fst.json_expected create mode 100644 tests/error-messages/Bug3292.fst.json_expected create mode 100644 tests/error-messages/Calc.fst.json_expected create mode 100644 tests/error-messages/CalcImpl.fst.json_expected create mode 100644 tests/error-messages/Coercions.fst.json_expected create mode 100644 tests/error-messages/DTuples.fst.json_expected create mode 100644 tests/error-messages/DecreasesTypeWarning.fst.json_expected create mode 100644 tests/error-messages/DumpUvarsOf.fst.json_expected create mode 100644 tests/error-messages/Erasable.fst.json_expected create mode 100644 tests/error-messages/ExpectFailure.fst.json_expected create mode 100644 tests/error-messages/GhostImplicits.fst.json_expected create mode 100644 tests/error-messages/IfCond.fst.json_expected create mode 100644 tests/error-messages/IfThen.fst.json_expected create mode 100644 tests/error-messages/Inference.fst.json_expected create mode 100644 tests/error-messages/NegativeTests.BST.fst.json_expected create mode 100644 tests/error-messages/NegativeTests.Bug260.fst.json_expected create mode 100644 tests/error-messages/NegativeTests.False.fst.json_expected create mode 100644 tests/error-messages/NegativeTests.Heap.fst.json_expected create mode 100644 tests/error-messages/NegativeTests.Neg.fst.json_expected create mode 100644 tests/error-messages/NegativeTests.Positivity.fst.json_expected create mode 100644 tests/error-messages/NegativeTests.Set.fst.json_expected create mode 100644 tests/error-messages/NegativeTests.ShortCircuiting.fst.json_expected create mode 100644 tests/error-messages/NegativeTests.Termination.fst.json_expected create mode 100644 tests/error-messages/NegativeTests.ZZImplicitFalse.fst.json_expected create mode 100644 tests/error-messages/OccursCheckOnArrows.fst.json_expected create mode 100644 tests/error-messages/OptionStack.fst.json_expected create mode 100644 tests/error-messages/PatAnnot.fst.json_expected create mode 100644 tests/error-messages/PatCoerce.fst.json_expected create mode 100644 tests/error-messages/PatternMatch.fst.json_expected create mode 100644 tests/error-messages/QuickTest.fst.json_expected create mode 100644 tests/error-messages/QuickTestNBE.fst.json_expected create mode 100644 tests/error-messages/RecordFields.fst.json_expected create mode 100644 tests/error-messages/ResolveImplicitsErrorPos.fst.json_expected create mode 100644 tests/error-messages/SMTPatSymbols.fst.json_expected create mode 100644 tests/error-messages/SeqLit.fst.json_expected create mode 100644 tests/error-messages/StableErr.fst.json_expected create mode 100644 tests/error-messages/StrictUnfolding.fst.json_expected create mode 100644 tests/error-messages/StringNormalization.fst.json_expected create mode 100644 tests/error-messages/Test.FunctionalExtensionality.fst.json_expected create mode 100644 tests/error-messages/TestErrorLocations.fst.json_expected create mode 100644 tests/error-messages/TestHasEq.fst.json_expected create mode 100644 tests/error-messages/UnboundOp.fst.json_expected create mode 100644 tests/error-messages/Unit2.fst.json_expected create mode 100644 tests/error-messages/UnresolvedFields.fst.json_expected create mode 100644 tests/error-messages/WPExtensionality.fst.json_expected diff --git a/examples/Makefile.common b/examples/Makefile.common index 505092b1445..53814a59a23 100644 --- a/examples/Makefile.common +++ b/examples/Makefile.common @@ -102,6 +102,14 @@ include .depend $(call msg, "OUTPUT", $(basename $(notdir $@))) $(Q)$(VERBOSE_FSTAR) -f --print_expected_failures $< >$@ 2>&1 +%.fst.json_output: %.fst + $(call msg, "JSONOUT", $(basename $(notdir $@))) + $(Q)$(VERBOSE_FSTAR) --message_format json -f --print_expected_failures $< >$@ 2>&1 + +%.fsti.json_output: %.fsti + $(call msg, "JSONOUT", $(basename $(notdir $@))) + $(Q)$(VERBOSE_FSTAR) --message_format json -f --print_expected_failures $< >$@ 2>&1 + clean: rm -rf $(CACHE_DIR) diff --git a/tests/error-messages/.gitignore b/tests/error-messages/.gitignore index 15878c51119..99e3e101bb3 100644 --- a/tests/error-messages/.gitignore +++ b/tests/error-messages/.gitignore @@ -1,2 +1,5 @@ *.output +*.json_output *.check +*.human_check +*.json_check diff --git a/tests/error-messages/AQualMismatch.fst.json_expected b/tests/error-messages/AQualMismatch.fst.json_expected new file mode 100644 index 00000000000..0c709512895 --- /dev/null +++ b/tests/error-messages/AQualMismatch.fst.json_expected @@ -0,0 +1,7 @@ +>> Got issues: [ +{"msg":["Inconsistent implicit argument annotation on argument x","Got: '#'","Expected: ''"],"level":"Error","range":{"def":{"file_name":"AQualMismatch.fst","start_pos":{"line":6,"col":7},"end_pos":{"line":6,"col":8}},"use":{"file_name":"AQualMismatch.fst","start_pos":{"line":6,"col":7},"end_pos":{"line":6,"col":8}}},"number":91,"ctx":["While typechecking the top-level declaration `let f`","While typechecking the top-level declaration `[@@expect_failure] let f`"]} +>>] +{"msg":["AQualMismatch.f\nis declared but no definition was found","Add an 'assume' if this is intentional"],"level":"Warning","range":{"def":{"file_name":"AQualMismatch.fst","start_pos":{"line":3,"col":4},"end_pos":{"line":3,"col":5}},"use":{"file_name":"AQualMismatch.fst","start_pos":{"line":3,"col":4},"end_pos":{"line":3,"col":5}}},"number":240,"ctx":["While desugaring module AQualMismatch"]} +{"msg":["Missing definitions in module AQualMismatch: f"],"level":"Warning","range":{"def":{"file_name":"dummy","start_pos":{"line":0,"col":0},"end_pos":{"line":0,"col":0}},"use":{"file_name":"AQualMismatch.fst","start_pos":{"line":6,"col":0},"end_pos":{"line":6,"col":12}}},"number":240,"ctx":[]} +Verified module: AQualMismatch +All verification conditions discharged successfully diff --git a/tests/error-messages/AnotType.fst.json_expected b/tests/error-messages/AnotType.fst.json_expected new file mode 100644 index 00000000000..b3bff7f6336 --- /dev/null +++ b/tests/error-messages/AnotType.fst.json_expected @@ -0,0 +1,8 @@ +>> Got issues: [ +{"msg":["Type annotation Prims.int for inductive AnotType.tc is not Type or eqtype, or it\nis eqtype but contains noeq/unopteq qualifiers"],"level":"Error","range":{"def":{"file_name":"AnotType.fst","start_pos":{"line":19,"col":5},"end_pos":{"line":19,"col":7}},"use":{"file_name":"AnotType.fst","start_pos":{"line":19,"col":5},"end_pos":{"line":19,"col":7}}},"number":309,"ctx":["While typechecking the top-level declaration `type AnotType.tc`","While typechecking the top-level declaration `[@@expect_failure] type AnotType.tc`"]} +>>] +>> Got issues: [ +{"msg":["Expected expression of type Type0\ngot expression tb\nof type Type"],"level":"Error","range":{"def":{"file_name":"AnotType.fst","start_pos":{"line":27,"col":14},"end_pos":{"line":27,"col":16}},"use":{"file_name":"AnotType.fst","start_pos":{"line":27,"col":14},"end_pos":{"line":27,"col":16}}},"number":189,"ctx":["While typechecking the top-level declaration `let uu___0`","While typechecking the top-level declaration `[@@expect_failure] let uu___0`"]} +>>] +Verified module: AnotType +All verification conditions discharged successfully diff --git a/tests/error-messages/ArgsAndQuals.fst.json_expected b/tests/error-messages/ArgsAndQuals.fst.json_expected new file mode 100644 index 00000000000..f1545bc8d52 --- /dev/null +++ b/tests/error-messages/ArgsAndQuals.fst.json_expected @@ -0,0 +1,7 @@ +>> Got issues: [ +{"msg":["Inconsistent implicit argument annotation on argument uu___","Got: '#'","Expected: '$'"],"level":"Error","range":{"def":{"file_name":"ArgsAndQuals.fst","start_pos":{"line":25,"col":16},"end_pos":{"line":25,"col":18}},"use":{"file_name":"ArgsAndQuals.fst","start_pos":{"line":25,"col":16},"end_pos":{"line":25,"col":18}}},"number":91,"ctx":["While typechecking the top-level declaration `let test1`","While typechecking the top-level declaration `[@@expect_failure] let test1`"]} +>>] +{"msg":["ArgsAndQuals.test1\nis declared but no definition was found","Add an 'assume' if this is intentional"],"level":"Warning","range":{"def":{"file_name":"ArgsAndQuals.fst","start_pos":{"line":23,"col":4},"end_pos":{"line":23,"col":9}},"use":{"file_name":"ArgsAndQuals.fst","start_pos":{"line":23,"col":4},"end_pos":{"line":23,"col":9}}},"number":240,"ctx":["While desugaring module ArgsAndQuals"]} +{"msg":["Missing definitions in module ArgsAndQuals: test1"],"level":"Warning","range":{"def":{"file_name":"dummy","start_pos":{"line":0,"col":0},"end_pos":{"line":0,"col":0}},"use":{"file_name":"ArgsAndQuals.fst","start_pos":{"line":25,"col":0},"end_pos":{"line":25,"col":29}}},"number":240,"ctx":[]} +Verified module: ArgsAndQuals +All verification conditions discharged successfully diff --git a/tests/error-messages/ArrowRanges.fst.json_expected b/tests/error-messages/ArrowRanges.fst.json_expected new file mode 100644 index 00000000000..d019cc1a11d --- /dev/null +++ b/tests/error-messages/ArrowRanges.fst.json_expected @@ -0,0 +1,8 @@ +>> Got issues: [ +{"msg":["Subtyping check failed","Expected type Prims.eqtype\ngot type Type0","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"Prims.fst","start_pos":{"line":73,"col":23},"end_pos":{"line":73,"col":30}},"use":{"file_name":"ArrowRanges.fst","start_pos":{"line":4,"col":30},"end_pos":{"line":4,"col":39}}},"number":19,"ctx":["While typechecking the top-level declaration `let ppof`","While typechecking the top-level declaration `[@@expect_failure] let ppof`"]} +>>] +>> Got issues: [ +{"msg":["Failed to prove that the type\n'ArrowRanges.ppof'\nsupports decidable equality because of this argument.","Add either the 'noeq' or 'unopteq' qualifier","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"ArrowRanges.fst","start_pos":{"line":7,"col":0},"end_pos":{"line":11,"col":1}},"use":{"file_name":"ArrowRanges.fst","start_pos":{"line":8,"col":10},"end_pos":{"line":8,"col":28}}},"number":19,"ctx":["While typechecking the top-level declaration `type ArrowRanges.ppof`","While typechecking the top-level declaration `[@@expect_failure] type ArrowRanges.ppof`"]} +>>] +Verified module: ArrowRanges +All verification conditions discharged successfully diff --git a/tests/error-messages/AssertNorm.fst.json_expected b/tests/error-messages/AssertNorm.fst.json_expected new file mode 100644 index 00000000000..9f6b602a8a1 --- /dev/null +++ b/tests/error-messages/AssertNorm.fst.json_expected @@ -0,0 +1,5 @@ +>> Got issues: [ +{"msg":["Assertion failed","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"AssertNorm.fst","start_pos":{"line":7,"col":14},"end_pos":{"line":7,"col":30}},"use":{"file_name":"AssertNorm.fst","start_pos":{"line":7,"col":2},"end_pos":{"line":7,"col":13}}},"number":19,"ctx":["While typechecking the top-level declaration `let f`","While typechecking the top-level declaration `[@@expect_failure] let f`"]} +>>] +Verified module: AssertNorm +All verification conditions discharged successfully diff --git a/tests/error-messages/Asserts.fst.json_expected b/tests/error-messages/Asserts.fst.json_expected new file mode 100644 index 00000000000..b51bae7c3fa --- /dev/null +++ b/tests/error-messages/Asserts.fst.json_expected @@ -0,0 +1,11 @@ +>> Got issues: [ +{"msg":["Assertion failed","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"Asserts.fst","start_pos":{"line":6,"col":9},"end_pos":{"line":6,"col":17}},"use":{"file_name":"Asserts.fst","start_pos":{"line":6,"col":9},"end_pos":{"line":6,"col":17}}},"number":19,"ctx":["While typechecking the top-level declaration `let test`","While typechecking the top-level declaration `[@@expect_failure] let test`"]} +>>] +>> Got issues: [ +{"msg":["Assertion failed","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"Asserts.fst","start_pos":{"line":11,"col":9},"end_pos":{"line":11,"col":17}},"use":{"file_name":"Asserts.fst","start_pos":{"line":11,"col":2},"end_pos":{"line":11,"col":8}}},"number":19,"ctx":["While typechecking the top-level declaration `let test2`","While typechecking the top-level declaration `[@@expect_failure] let test2`"]} +>>] +>> Got issues: [ +{"msg":["Assertion failed","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"Asserts.fst","start_pos":{"line":16,"col":9},"end_pos":{"line":16,"col":14}},"use":{"file_name":"Asserts.fst","start_pos":{"line":16,"col":2},"end_pos":{"line":16,"col":8}}},"number":19,"ctx":["While typechecking the top-level declaration `let test3`","While typechecking the top-level declaration `[@@expect_failure] let test3`"]} +>>] +Verified module: Asserts +All verification conditions discharged successfully diff --git a/tests/error-messages/BadEmptyRecord.fst.json_expected b/tests/error-messages/BadEmptyRecord.fst.json_expected new file mode 100644 index 00000000000..2d14bacf5d7 --- /dev/null +++ b/tests/error-messages/BadEmptyRecord.fst.json_expected @@ -0,0 +1,5 @@ +>> Got issues: [ +{"msg":["Could not resolve the type for this record."],"level":"Error","range":{"def":{"file_name":"BadEmptyRecord.fst","start_pos":{"line":9,"col":6},"end_pos":{"line":9,"col":8}},"use":{"file_name":"BadEmptyRecord.fst","start_pos":{"line":9,"col":6},"end_pos":{"line":9,"col":8}}},"number":360,"ctx":["While typechecking the top-level declaration `let test`","While typechecking the top-level declaration `[@@expect_failure] let test`"]} +>>] +Verified module: BadEmptyRecord +All verification conditions discharged successfully diff --git a/tests/error-messages/Basic.fst.json_expected b/tests/error-messages/Basic.fst.json_expected new file mode 100644 index 00000000000..7a5d46ffd34 --- /dev/null +++ b/tests/error-messages/Basic.fst.json_expected @@ -0,0 +1,38 @@ +>> Got issues: [ +{"msg":["Expected expression of type Prims.int\ngot expression true\nof type Prims.bool"],"level":"Error","range":{"def":{"file_name":"Basic.fst","start_pos":{"line":4,"col":13},"end_pos":{"line":4,"col":17}},"use":{"file_name":"Basic.fst","start_pos":{"line":4,"col":13},"end_pos":{"line":4,"col":17}}},"number":189,"ctx":["While typechecking the top-level declaration `let uu___0`","While typechecking the top-level declaration `[@@expect_failure] let uu___0`"]} +>>] +>> Got issues: [ +{"msg":["Assertion failed","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"Basic.fst","start_pos":{"line":6,"col":45},"end_pos":{"line":6,"col":50}},"use":{"file_name":"Basic.fst","start_pos":{"line":6,"col":38},"end_pos":{"line":6,"col":44}}},"number":19,"ctx":["While checking for top-level effects","While typechecking the top-level declaration `let uu___1`","While typechecking the top-level declaration `[@@expect_failure] let uu___1`"]} +>>] +>> Got issues: [ +{"msg":["Assertion failed","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"Basic.fst","start_pos":{"line":7,"col":45},"end_pos":{"line":7,"col":50}},"use":{"file_name":"Basic.fst","start_pos":{"line":7,"col":38},"end_pos":{"line":7,"col":44}}},"number":19,"ctx":["While checking for top-level effects","While typechecking the top-level declaration `let uu___2`","While typechecking the top-level declaration `[@@expect_failure] let uu___2`"]} +>>] +>> Got issues: [ +{"msg":["Assertion failed","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"Basic.fst","start_pos":{"line":8,"col":45},"end_pos":{"line":8,"col":50}},"use":{"file_name":"Basic.fst","start_pos":{"line":8,"col":38},"end_pos":{"line":8,"col":44}}},"number":19,"ctx":["While checking for top-level effects","While typechecking the top-level declaration `let uu___4`","While typechecking the top-level declaration `[@@expect_failure] let uu___4`"]} +>>] +>> Got issues: [ +{"msg":["Assertion failed","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"Basic.fst","start_pos":{"line":9,"col":45},"end_pos":{"line":9,"col":50}},"use":{"file_name":"Basic.fst","start_pos":{"line":9,"col":38},"end_pos":{"line":9,"col":44}}},"number":19,"ctx":["While checking for top-level effects","While typechecking the top-level declaration `let uu___6`","While typechecking the top-level declaration `[@@expect_failure] let uu___6`"]} +>>] +>> Got issues: [ +{"msg":["Assertion failed","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"Basic.fst","start_pos":{"line":11,"col":50},"end_pos":{"line":11,"col":55}},"use":{"file_name":"Basic.fst","start_pos":{"line":11,"col":38},"end_pos":{"line":11,"col":49}}},"number":19,"ctx":["While checking for top-level effects","While typechecking the top-level declaration `let uu___9`","While typechecking the top-level declaration `[@@expect_failure] let uu___9`"]} +>>] +>> Got issues: [ +{"msg":["Assertion failed","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"Basic.fst","start_pos":{"line":12,"col":50},"end_pos":{"line":12,"col":55}},"use":{"file_name":"Basic.fst","start_pos":{"line":12,"col":38},"end_pos":{"line":12,"col":49}}},"number":19,"ctx":["While checking for top-level effects","While typechecking the top-level declaration `let uu___10`","While typechecking the top-level declaration `[@@expect_failure] let uu___10`"]} +>>] +>> Got issues: [ +{"msg":["Assertion failed","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"Basic.fst","start_pos":{"line":13,"col":50},"end_pos":{"line":13,"col":55}},"use":{"file_name":"Basic.fst","start_pos":{"line":13,"col":38},"end_pos":{"line":13,"col":49}}},"number":19,"ctx":["While checking for top-level effects","While typechecking the top-level declaration `let uu___12`","While typechecking the top-level declaration `[@@expect_failure] let uu___12`"]} +>>] +>> Got issues: [ +{"msg":["Assertion failed","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"Basic.fst","start_pos":{"line":14,"col":50},"end_pos":{"line":14,"col":55}},"use":{"file_name":"Basic.fst","start_pos":{"line":14,"col":38},"end_pos":{"line":14,"col":49}}},"number":19,"ctx":["While checking for top-level effects","While typechecking the top-level declaration `let uu___14`","While typechecking the top-level declaration `[@@expect_failure] let uu___14`"]} +>>] +>> Got issues: [ +{"msg":["Assertion failed","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"Prims.fst","start_pos":{"line":451,"col":90},"end_pos":{"line":451,"col":102}},"use":{"file_name":"Basic.fst","start_pos":{"line":17,"col":29},"end_pos":{"line":17,"col":31}}},"number":19,"ctx":["While typechecking the top-level declaration `let test2`","While typechecking the top-level declaration `[@@expect_failure] let test2`"]} +>>] +>> Got issues: [ +{"msg":["Assertion failed","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"Prims.fst","start_pos":{"line":451,"col":90},"end_pos":{"line":451,"col":102}},"use":{"file_name":"Basic.fst","start_pos":{"line":20,"col":29},"end_pos":{"line":20,"col":31}}},"number":19,"ctx":["While typechecking the top-level declaration `let test3`","While typechecking the top-level declaration `[@@expect_failure] let test3`"]} +>>] +>> Got issues: [ +{"msg":["Assertion failed","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"Prims.fst","start_pos":{"line":451,"col":90},"end_pos":{"line":451,"col":102}},"use":{"file_name":"Basic.fst","start_pos":{"line":23,"col":46},"end_pos":{"line":23,"col":48}}},"number":19,"ctx":["While typechecking the top-level declaration `let test6`","While typechecking the top-level declaration `[@@expect_failure] let test6`"]} +>>] +Verified module: Basic +All verification conditions discharged successfully diff --git a/tests/error-messages/Bug1918.fst.json_expected b/tests/error-messages/Bug1918.fst.json_expected new file mode 100644 index 00000000000..5c42dd627ac --- /dev/null +++ b/tests/error-messages/Bug1918.fst.json_expected @@ -0,0 +1,5 @@ +>> Got issues: [ +{"msg":["Could not solve typeclass constraint `Bug1918.mon`"],"level":"Error","range":{"def":{"file_name":"FStar.Tactics.Typeclasses.fst","start_pos":{"line":297,"col":6},"end_pos":{"line":300,"col":7}},"use":{"file_name":"Bug1918.fst","start_pos":{"line":11,"col":13},"end_pos":{"line":11,"col":14}}},"number":228,"ctx":["While synthesizing term with a tactic","Running tactic for meta-arg","While typechecking the top-level declaration `let comp2`","While typechecking the top-level declaration `[@@expect_failure] let comp2`"]} +>>] +Verified module: Bug1918 +All verification conditions discharged successfully diff --git a/tests/error-messages/Bug1988.fst.json_expected b/tests/error-messages/Bug1988.fst.json_expected new file mode 100644 index 00000000000..3c8aa50053d --- /dev/null +++ b/tests/error-messages/Bug1988.fst.json_expected @@ -0,0 +1,11 @@ +>> Got issues: [ +{"msg":["Expected expression of type Prims.int\ngot expression \"string literal\"\nof type Prims.string"],"level":"Error","range":{"def":{"file_name":"Bug1988.fst","start_pos":{"line":4,"col":14},"end_pos":{"line":4,"col":30}},"use":{"file_name":"Bug1988.fst","start_pos":{"line":4,"col":14},"end_pos":{"line":4,"col":30}}},"number":189,"ctx":["While typechecking the top-level declaration `let x`","While typechecking the top-level declaration `[@@expect_failure] let x`"]} +>>] +>> Got issues: [ +{"msg":["Prims.string is not equal to the expected type _: ident -> Prims.string"],"level":"Error","range":{"def":{"file_name":"dummy","start_pos":{"line":0,"col":0},"end_pos":{"line":0,"col":0}},"use":{"file_name":"Bug1988.fst","start_pos":{"line":15,"col":32},"end_pos":{"line":15,"col":38}}},"number":54,"ctx":["While solving deferred constraints","solve_non_tactic_deferred_constraints","While typechecking the top-level declaration `let f1`","While typechecking the top-level declaration `[@@expect_failure] let f1`"]} +>>] +>> Got issues: [ +{"msg":["Prims.string is not a subtype of the expected type _: (*?u21*) _ -> (*?u22*) _"],"level":"Error","range":{"def":{"file_name":"Bug1988.fst","start_pos":{"line":21,"col":32},"end_pos":{"line":21,"col":38}},"use":{"file_name":"Bug1988.fst","start_pos":{"line":21,"col":32},"end_pos":{"line":21,"col":38}}},"number":54,"ctx":["While solving deferred constraints","solve_non_tactic_deferred_constraints","While typechecking the top-level declaration `let f2`","While typechecking the top-level declaration `[@@expect_failure] let f2`"]} +>>] +Verified module: Bug1988 +All verification conditions discharged successfully diff --git a/tests/error-messages/Bug1997.fst.json_expected b/tests/error-messages/Bug1997.fst.json_expected new file mode 100644 index 00000000000..786b8ea9a7e --- /dev/null +++ b/tests/error-messages/Bug1997.fst.json_expected @@ -0,0 +1,295 @@ +Module after desugaring: +module Bug1997 +Declarations: [ +let op_Equals_Hat a b = a + b +let test1 a b c = a =^ b =^ c +let test2 a b c = a =^ b =^ c +let test3 a b c = a =^ (b =^ c) +private +let _ = FStar.Pervasives.assert_norm (Bug1997.test1 == Bug1997.test2) + +let imp_assoc_0 p q r = p ==> q ==> r +let imp_assoc_1 p q r = p ==> q ==> r +let imp_assoc_2 p q r = (p ==> q) ==> r +private +let _ = FStar.Pervasives.assert_norm (Bug1997.imp_assoc_0 == Bug1997.imp_assoc_1) + +let imp_iff_0 p q r = p ==> q <==> r +let imp_iff_1 p q r = p ==> q <==> r +let imp_iff_2 p q r = p ==> (q <==> r) +private +let _ = FStar.Pervasives.assert_norm (Bug1997.imp_iff_0 == Bug1997.imp_iff_1) + +let imp_conj_1 p q r s = p ==> q /\ r ==> s +let imp_conj_2 p q r s = (p ==> q) /\ (r ==> s) +let imp_conj_3 p q r s = p ==> q /\ (r ==> s) +let imp_conj_4 p q r s = p ==> q /\ r ==> s +private +let _ = FStar.Pervasives.assert_norm (Bug1997.imp_conj_1 == Bug1997.imp_conj_4) + + + + + +let impneg1 p q r = p ==> q /\ ~p ==> r +let impneg2 p q r = (p ==> q) /\ (~p ==> r) +let impneg3 p q r = p ==> q /\ (~p ==> r) +let impneg4 p q r = p ==> q /\ ~p ==> r +private +let _ = FStar.Pervasives.assert_norm (Bug1997.impneg1 == Bug1997.impneg4) + + + + + +let cd1 p q r = p /\ q \/ r +let cd2 p q r = p /\ q \/ r +let cd3 p q r = p /\ (q \/ r) +private +let _ = FStar.Pervasives.assert_norm (Bug1997.cd1 == Bug1997.cd2) + + +let m0 p q r s = p /\ (q \/ r) /\ s +let m1 p q r s = p /\ q \/ r /\ s +let m2 p q r s = p /\ q \/ r /\ s +let m3 p q r s = p /\ q \/ r /\ s +let m4 p q r s = p /\ q \/ r /\ s +private +let _ = FStar.Pervasives.assert_norm (Bug1997.m1 == Bug1997.m2) +private +let _ = FStar.Pervasives.assert_norm (Bug1997.m1 == Bug1997.m3) +private +let _ = FStar.Pervasives.assert_norm (Bug1997.m1 == Bug1997.m4) + + + + +let n0 p q r s = p \/ q /\ r \/ s +let n1 p q r s = p \/ q /\ r \/ s +let n2 p q r s = (p \/ q) /\ r \/ s +let n3 p q r s = (p \/ q) /\ (r \/ s) +let n4 p q r s = p \/ q /\ (r \/ s) +private +let _ = FStar.Pervasives.assert_norm (Bug1997.n0 == Bug1997.n1) + + + + + + +] + +Module before type checking: +module Bug1997 +Declarations: [ +let op_Equals_Hat a b = a + b +let test1 a b c = a =^ b =^ c +let test2 a b c = a =^ b =^ c +let test3 a b c = a =^ (b =^ c) +private +let _ = FStar.Pervasives.assert_norm (Bug1997.test1 == Bug1997.test2) + +let imp_assoc_0 p q r = p ==> q ==> r +let imp_assoc_1 p q r = p ==> q ==> r +let imp_assoc_2 p q r = (p ==> q) ==> r +private +let _ = FStar.Pervasives.assert_norm (Bug1997.imp_assoc_0 == Bug1997.imp_assoc_1) + +let imp_iff_0 p q r = p ==> q <==> r +let imp_iff_1 p q r = p ==> q <==> r +let imp_iff_2 p q r = p ==> (q <==> r) +private +let _ = FStar.Pervasives.assert_norm (Bug1997.imp_iff_0 == Bug1997.imp_iff_1) + +let imp_conj_1 p q r s = p ==> q /\ r ==> s +let imp_conj_2 p q r s = (p ==> q) /\ (r ==> s) +let imp_conj_3 p q r s = p ==> q /\ (r ==> s) +let imp_conj_4 p q r s = p ==> q /\ r ==> s +private +let _ = FStar.Pervasives.assert_norm (Bug1997.imp_conj_1 == Bug1997.imp_conj_4) + + + + + +let impneg1 p q r = p ==> q /\ ~p ==> r +let impneg2 p q r = (p ==> q) /\ (~p ==> r) +let impneg3 p q r = p ==> q /\ (~p ==> r) +let impneg4 p q r = p ==> q /\ ~p ==> r +private +let _ = FStar.Pervasives.assert_norm (Bug1997.impneg1 == Bug1997.impneg4) + + + + + +let cd1 p q r = p /\ q \/ r +let cd2 p q r = p /\ q \/ r +let cd3 p q r = p /\ (q \/ r) +private +let _ = FStar.Pervasives.assert_norm (Bug1997.cd1 == Bug1997.cd2) + + +let m0 p q r s = p /\ (q \/ r) /\ s +let m1 p q r s = p /\ q \/ r /\ s +let m2 p q r s = p /\ q \/ r /\ s +let m3 p q r s = p /\ q \/ r /\ s +let m4 p q r s = p /\ q \/ r /\ s +private +let _ = FStar.Pervasives.assert_norm (Bug1997.m1 == Bug1997.m2) +private +let _ = FStar.Pervasives.assert_norm (Bug1997.m1 == Bug1997.m3) +private +let _ = FStar.Pervasives.assert_norm (Bug1997.m1 == Bug1997.m4) + + + + +let n0 p q r s = p \/ q /\ r \/ s +let n1 p q r s = p \/ q /\ r \/ s +let n2 p q r s = (p \/ q) /\ r \/ s +let n3 p q r s = (p \/ q) /\ (r \/ s) +let n4 p q r s = p \/ q /\ (r \/ s) +private +let _ = FStar.Pervasives.assert_norm (Bug1997.n0 == Bug1997.n1) + + + + + + +] + +>> Got issues: [ +{"msg":["Assertion failed","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"Bug1997.fst","start_pos":{"line":13,"col":39},"end_pos":{"line":13,"col":55}},"use":{"file_name":"Bug1997.fst","start_pos":{"line":13,"col":27},"end_pos":{"line":13,"col":38}}},"number":19,"ctx":["While checking for top-level effects","While typechecking the top-level declaration `let uu___12`","While typechecking the top-level declaration `[@@expect_failure] let uu___12`"]} +>>] +>> Got issues: [ +{"msg":["Assertion failed","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"Bug1997.fst","start_pos":{"line":21,"col":39},"end_pos":{"line":21,"col":67}},"use":{"file_name":"Bug1997.fst","start_pos":{"line":21,"col":27},"end_pos":{"line":21,"col":38}}},"number":19,"ctx":["While checking for top-level effects","While typechecking the top-level declaration `let uu___23`","While typechecking the top-level declaration `[@@expect_failure] let uu___23`"]} +>>] +>> Got issues: [ +{"msg":["Assertion failed","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"Bug1997.fst","start_pos":{"line":29,"col":39},"end_pos":{"line":29,"col":63}},"use":{"file_name":"Bug1997.fst","start_pos":{"line":29,"col":27},"end_pos":{"line":29,"col":38}}},"number":19,"ctx":["While checking for top-level effects","While typechecking the top-level declaration `let uu___34`","While typechecking the top-level declaration `[@@expect_failure] let uu___34`"]} +>>] +>> Got issues: [ +{"msg":["Assertion failed","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"Bug1997.fst","start_pos":{"line":38,"col":39},"end_pos":{"line":38,"col":65}},"use":{"file_name":"Bug1997.fst","start_pos":{"line":38,"col":27},"end_pos":{"line":38,"col":38}}},"number":19,"ctx":["While checking for top-level effects","While typechecking the top-level declaration `let uu___52`","While typechecking the top-level declaration `[@@expect_failure] let uu___52`"]} +>>] +>> Got issues: [ +{"msg":["Assertion failed","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"Bug1997.fst","start_pos":{"line":39,"col":39},"end_pos":{"line":39,"col":65}},"use":{"file_name":"Bug1997.fst","start_pos":{"line":39,"col":27},"end_pos":{"line":39,"col":38}}},"number":19,"ctx":["While checking for top-level effects","While typechecking the top-level declaration `let uu___53`","While typechecking the top-level declaration `[@@expect_failure] let uu___53`"]} +>>] +>> Got issues: [ +{"msg":["Assertion failed","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"Bug1997.fst","start_pos":{"line":40,"col":39},"end_pos":{"line":40,"col":65}},"use":{"file_name":"Bug1997.fst","start_pos":{"line":40,"col":27},"end_pos":{"line":40,"col":38}}},"number":19,"ctx":["While checking for top-level effects","While typechecking the top-level declaration `let uu___54`","While typechecking the top-level declaration `[@@expect_failure] let uu___54`"]} +>>] +>> Got issues: [ +{"msg":["Assertion failed","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"Bug1997.fst","start_pos":{"line":41,"col":39},"end_pos":{"line":41,"col":65}},"use":{"file_name":"Bug1997.fst","start_pos":{"line":41,"col":27},"end_pos":{"line":41,"col":38}}},"number":19,"ctx":["While checking for top-level effects","While typechecking the top-level declaration `let uu___55`","While typechecking the top-level declaration `[@@expect_failure] let uu___55`"]} +>>] +>> Got issues: [ +{"msg":["Assertion failed","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"Bug1997.fst","start_pos":{"line":42,"col":39},"end_pos":{"line":42,"col":65}},"use":{"file_name":"Bug1997.fst","start_pos":{"line":42,"col":27},"end_pos":{"line":42,"col":38}}},"number":19,"ctx":["While checking for top-level effects","While typechecking the top-level declaration `let uu___56`","While typechecking the top-level declaration `[@@expect_failure] let uu___56`"]} +>>] +>> Got issues: [ +{"msg":["Assertion failed","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"Bug1997.fst","start_pos":{"line":51,"col":39},"end_pos":{"line":51,"col":59}},"use":{"file_name":"Bug1997.fst","start_pos":{"line":51,"col":27},"end_pos":{"line":51,"col":38}}},"number":19,"ctx":["While checking for top-level effects","While typechecking the top-level declaration `let uu___70`","While typechecking the top-level declaration `[@@expect_failure] let uu___70`"]} +>>] +>> Got issues: [ +{"msg":["Assertion failed","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"Bug1997.fst","start_pos":{"line":52,"col":39},"end_pos":{"line":52,"col":59}},"use":{"file_name":"Bug1997.fst","start_pos":{"line":52,"col":27},"end_pos":{"line":52,"col":38}}},"number":19,"ctx":["While checking for top-level effects","While typechecking the top-level declaration `let uu___71`","While typechecking the top-level declaration `[@@expect_failure] let uu___71`"]} +>>] +>> Got issues: [ +{"msg":["Assertion failed","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"Bug1997.fst","start_pos":{"line":53,"col":39},"end_pos":{"line":53,"col":59}},"use":{"file_name":"Bug1997.fst","start_pos":{"line":53,"col":27},"end_pos":{"line":53,"col":38}}},"number":19,"ctx":["While checking for top-level effects","While typechecking the top-level declaration `let uu___72`","While typechecking the top-level declaration `[@@expect_failure] let uu___72`"]} +>>] +>> Got issues: [ +{"msg":["Assertion failed","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"Bug1997.fst","start_pos":{"line":54,"col":39},"end_pos":{"line":54,"col":59}},"use":{"file_name":"Bug1997.fst","start_pos":{"line":54,"col":27},"end_pos":{"line":54,"col":38}}},"number":19,"ctx":["While checking for top-level effects","While typechecking the top-level declaration `let uu___73`","While typechecking the top-level declaration `[@@expect_failure] let uu___73`"]} +>>] +>> Got issues: [ +{"msg":["Assertion failed","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"Bug1997.fst","start_pos":{"line":55,"col":39},"end_pos":{"line":55,"col":59}},"use":{"file_name":"Bug1997.fst","start_pos":{"line":55,"col":27},"end_pos":{"line":55,"col":38}}},"number":19,"ctx":["While checking for top-level effects","While typechecking the top-level declaration `let uu___74`","While typechecking the top-level declaration `[@@expect_failure] let uu___74`"]} +>>] +>> Got issues: [ +{"msg":["Assertion failed","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"Bug1997.fst","start_pos":{"line":63,"col":39},"end_pos":{"line":63,"col":51}},"use":{"file_name":"Bug1997.fst","start_pos":{"line":63,"col":27},"end_pos":{"line":63,"col":38}}},"number":19,"ctx":["While checking for top-level effects","While typechecking the top-level declaration `let uu___85`","While typechecking the top-level declaration `[@@expect_failure] let uu___85`"]} +>>] +>> Got issues: [ +{"msg":["Assertion failed","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"Bug1997.fst","start_pos":{"line":64,"col":39},"end_pos":{"line":64,"col":51}},"use":{"file_name":"Bug1997.fst","start_pos":{"line":64,"col":27},"end_pos":{"line":64,"col":38}}},"number":19,"ctx":["While checking for top-level effects","While typechecking the top-level declaration `let uu___86`","While typechecking the top-level declaration `[@@expect_failure] let uu___86`"]} +>>] +>> Got issues: [ +{"msg":["Assertion failed","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"Bug1997.fst","start_pos":{"line":75,"col":39},"end_pos":{"line":75,"col":49}},"use":{"file_name":"Bug1997.fst","start_pos":{"line":75,"col":27},"end_pos":{"line":75,"col":38}}},"number":19,"ctx":["While checking for top-level effects","While typechecking the top-level declaration `let uu___110`","While typechecking the top-level declaration `[@@expect_failure] let uu___110`"]} +>>] +>> Got issues: [ +{"msg":["Assertion failed","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"Bug1997.fst","start_pos":{"line":76,"col":39},"end_pos":{"line":76,"col":49}},"use":{"file_name":"Bug1997.fst","start_pos":{"line":76,"col":27},"end_pos":{"line":76,"col":38}}},"number":19,"ctx":["While checking for top-level effects","While typechecking the top-level declaration `let uu___111`","While typechecking the top-level declaration `[@@expect_failure] let uu___111`"]} +>>] +>> Got issues: [ +{"msg":["Assertion failed","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"Bug1997.fst","start_pos":{"line":77,"col":39},"end_pos":{"line":77,"col":49}},"use":{"file_name":"Bug1997.fst","start_pos":{"line":77,"col":27},"end_pos":{"line":77,"col":38}}},"number":19,"ctx":["While checking for top-level effects","While typechecking the top-level declaration `let uu___112`","While typechecking the top-level declaration `[@@expect_failure] let uu___112`"]} +>>] +>> Got issues: [ +{"msg":["Assertion failed","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"Bug1997.fst","start_pos":{"line":78,"col":39},"end_pos":{"line":78,"col":49}},"use":{"file_name":"Bug1997.fst","start_pos":{"line":78,"col":27},"end_pos":{"line":78,"col":38}}},"number":19,"ctx":["While checking for top-level effects","While typechecking the top-level declaration `let uu___113`","While typechecking the top-level declaration `[@@expect_failure] let uu___113`"]} +>>] +>> Got issues: [ +{"msg":["Assertion failed","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"Bug1997.fst","start_pos":{"line":87,"col":39},"end_pos":{"line":87,"col":49}},"use":{"file_name":"Bug1997.fst","start_pos":{"line":87,"col":27},"end_pos":{"line":87,"col":38}}},"number":19,"ctx":["While checking for top-level effects","While typechecking the top-level declaration `let uu___135`","While typechecking the top-level declaration `[@@expect_failure] let uu___135`"]} +>>] +>> Got issues: [ +{"msg":["Assertion failed","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"Bug1997.fst","start_pos":{"line":88,"col":39},"end_pos":{"line":88,"col":49}},"use":{"file_name":"Bug1997.fst","start_pos":{"line":88,"col":27},"end_pos":{"line":88,"col":38}}},"number":19,"ctx":["While checking for top-level effects","While typechecking the top-level declaration `let uu___136`","While typechecking the top-level declaration `[@@expect_failure] let uu___136`"]} +>>] +>> Got issues: [ +{"msg":["Assertion failed","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"Bug1997.fst","start_pos":{"line":89,"col":39},"end_pos":{"line":89,"col":49}},"use":{"file_name":"Bug1997.fst","start_pos":{"line":89,"col":27},"end_pos":{"line":89,"col":38}}},"number":19,"ctx":["While checking for top-level effects","While typechecking the top-level declaration `let uu___137`","While typechecking the top-level declaration `[@@expect_failure] let uu___137`"]} +>>] +>> Got issues: [ +{"msg":["Assertion failed","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"Bug1997.fst","start_pos":{"line":90,"col":39},"end_pos":{"line":90,"col":49}},"use":{"file_name":"Bug1997.fst","start_pos":{"line":90,"col":27},"end_pos":{"line":90,"col":38}}},"number":19,"ctx":["While checking for top-level effects","While typechecking the top-level declaration `let uu___138`","While typechecking the top-level declaration `[@@expect_failure] let uu___138`"]} +>>] +>> Got issues: [ +{"msg":["Assertion failed","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"Bug1997.fst","start_pos":{"line":91,"col":39},"end_pos":{"line":91,"col":49}},"use":{"file_name":"Bug1997.fst","start_pos":{"line":91,"col":27},"end_pos":{"line":91,"col":38}}},"number":19,"ctx":["While checking for top-level effects","While typechecking the top-level declaration `let uu___139`","While typechecking the top-level declaration `[@@expect_failure] let uu___139`"]} +>>] +>> Got issues: [ +{"msg":["Assertion failed","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"Bug1997.fst","start_pos":{"line":92,"col":39},"end_pos":{"line":92,"col":49}},"use":{"file_name":"Bug1997.fst","start_pos":{"line":92,"col":27},"end_pos":{"line":92,"col":38}}},"number":19,"ctx":["While checking for top-level effects","While typechecking the top-level declaration `let uu___140`","While typechecking the top-level declaration `[@@expect_failure] let uu___140`"]} +>>] +Module after type checking: +module Bug1997 +Declarations: [ +let op_Equals_Hat a b = a + b +let test1 a b c = a =^ b =^ c +let test2 a b c = a =^ b =^ c +let test3 a b c = a =^ (b =^ c) +private +let _ = FStar.Pervasives.assert_norm (Bug1997.test1 == Bug1997.test2) +let imp_assoc_0 p q r = p ==> q ==> r +let imp_assoc_1 p q r = p ==> q ==> r +let imp_assoc_2 p q r = (p ==> q) ==> r +private +let _ = FStar.Pervasives.assert_norm (Bug1997.imp_assoc_0 == Bug1997.imp_assoc_1) +let imp_iff_0 p q r = p ==> q <==> r +let imp_iff_1 p q r = p ==> q <==> r +let imp_iff_2 p q r = p ==> (q <==> r) +private +let _ = FStar.Pervasives.assert_norm (Bug1997.imp_iff_0 == Bug1997.imp_iff_1) +let imp_conj_1 p q r s = p ==> q /\ r ==> s +let imp_conj_2 p q r s = (p ==> q) /\ (r ==> s) +let imp_conj_3 p q r s = p ==> q /\ (r ==> s) +let imp_conj_4 p q r s = p ==> q /\ r ==> s +private +let _ = FStar.Pervasives.assert_norm (Bug1997.imp_conj_1 == Bug1997.imp_conj_4) +let impneg1 p q r = p ==> q /\ ~p ==> r +let impneg2 p q r = (p ==> q) /\ (~p ==> r) +let impneg3 p q r = p ==> q /\ (~p ==> r) +let impneg4 p q r = p ==> q /\ ~p ==> r +private +let _ = FStar.Pervasives.assert_norm (Bug1997.impneg1 == Bug1997.impneg4) +let cd1 p q r = p /\ q \/ r +let cd2 p q r = p /\ q \/ r +let cd3 p q r = p /\ (q \/ r) +private +let _ = FStar.Pervasives.assert_norm (Bug1997.cd1 == Bug1997.cd2) +let m0 p q r s = p /\ (q \/ r) /\ s +let m1 p q r s = p /\ q \/ r /\ s +let m2 p q r s = p /\ q \/ r /\ s +let m3 p q r s = p /\ q \/ r /\ s +let m4 p q r s = p /\ q \/ r /\ s +private +let _ = FStar.Pervasives.assert_norm (Bug1997.m1 == Bug1997.m2) +private +let _ = FStar.Pervasives.assert_norm (Bug1997.m1 == Bug1997.m3) +private +let _ = FStar.Pervasives.assert_norm (Bug1997.m1 == Bug1997.m4) +let n0 p q r s = p \/ q /\ r \/ s +let n1 p q r s = p \/ q /\ r \/ s +let n2 p q r s = (p \/ q) /\ r \/ s +let n3 p q r s = (p \/ q) /\ (r \/ s) +let n4 p q r s = p \/ q /\ (r \/ s) +private +let _ = FStar.Pervasives.assert_norm (Bug1997.n0 == Bug1997.n1) +] + +Verified module: Bug1997 +All verification conditions discharged successfully diff --git a/tests/error-messages/Bug2010.fst.json_expected b/tests/error-messages/Bug2010.fst.json_expected new file mode 100644 index 00000000000..2c0b1ccafd8 --- /dev/null +++ b/tests/error-messages/Bug2010.fst.json_expected @@ -0,0 +1,5 @@ +>> Got issues: [ +{"msg":["Failed to resolve implicit argument ?14\nof type Prims.nat\nintroduced for user-provided implicit term at Bug2010.fst(6,66-6,67)"],"level":"Error","range":{"def":{"file_name":"Bug2010.fst","start_pos":{"line":6,"col":66},"end_pos":{"line":6,"col":67}},"use":{"file_name":"Bug2010.fst","start_pos":{"line":6,"col":66},"end_pos":{"line":6,"col":67}}},"number":66,"ctx":["While checking for top-level effects","While typechecking the top-level declaration `let concat'`","While typechecking the top-level declaration `[@@expect_failure] let concat'`"]} +>>] +Verified module: Bug2010 +All verification conditions discharged successfully diff --git a/tests/error-messages/Bug2021.fst.json_expected b/tests/error-messages/Bug2021.fst.json_expected new file mode 100644 index 00000000000..66aa2effca8 --- /dev/null +++ b/tests/error-messages/Bug2021.fst.json_expected @@ -0,0 +1,20 @@ +>> Got issues: [ +{"msg":["Failed to resolve implicit argument ?2\nof type Type\nintroduced for user-provided implicit term at Bug2021.fst(5,15-5,16)"],"level":"Error","range":{"def":{"file_name":"Bug2021.fst","start_pos":{"line":5,"col":15},"end_pos":{"line":5,"col":16}},"use":{"file_name":"Bug2021.fst","start_pos":{"line":5,"col":15},"end_pos":{"line":5,"col":16}}},"number":66,"ctx":["While checking for top-level effects","While typechecking the top-level declaration `let test0`","While typechecking the top-level declaration `[@@expect_failure] let test0`"]} +>>] +>> Got issues: [ +{"msg":["Failed to resolve implicit argument ?2\nof type Type\nintroduced for user-provided implicit term at Bug2021.fst(10,8-10,9)"],"level":"Error","range":{"def":{"file_name":"Bug2021.fst","start_pos":{"line":10,"col":8},"end_pos":{"line":10,"col":9}},"use":{"file_name":"Bug2021.fst","start_pos":{"line":10,"col":8},"end_pos":{"line":10,"col":9}}},"number":66,"ctx":["While checking for top-level effects","While typechecking the top-level declaration `let test1`","While typechecking the top-level declaration `[@@expect_failure] let test1`"]} +>>] +>> Got issues: [ +{"msg":["Failed to resolve implicit argument ?5\nof type Type\nintroduced for user-provided implicit term at Bug2021.fst(16,12-16,13)"],"level":"Error","range":{"def":{"file_name":"Bug2021.fst","start_pos":{"line":16,"col":12},"end_pos":{"line":16,"col":13}},"use":{"file_name":"Bug2021.fst","start_pos":{"line":16,"col":12},"end_pos":{"line":16,"col":13}}},"number":66,"ctx":["While checking for top-level effects","While typechecking the top-level declaration `let test2`","While typechecking the top-level declaration `[@@expect_failure] let test2`"]} +>>] +>> Got issues: [ +{"msg":["Failed to resolve implicit argument ?9\nof type Prims.int\nintroduced for Instantiating implicit argument"],"level":"Error","range":{"def":{"file_name":"Bug2021.fst","start_pos":{"line":23,"col":11},"end_pos":{"line":23,"col":12}},"use":{"file_name":"Bug2021.fst","start_pos":{"line":24,"col":13},"end_pos":{"line":24,"col":14}}},"number":66,"ctx":["While checking for top-level effects","While typechecking the top-level declaration `let test3`","While typechecking the top-level declaration `[@@expect_failure] let test3`"]} +>>] +>> Got issues: [ +{"msg":["Failed to resolve implicit argument ?10\nof type Prims.int\nintroduced for Instantiating implicit argument"],"level":"Error","range":{"def":{"file_name":"Bug2021.fst","start_pos":{"line":29,"col":11},"end_pos":{"line":29,"col":12}},"use":{"file_name":"Bug2021.fst","start_pos":{"line":30,"col":13},"end_pos":{"line":30,"col":17}}},"number":66,"ctx":["While checking for top-level effects","While typechecking the top-level declaration `let test4`","While typechecking the top-level declaration `[@@expect_failure] let test4`"]} +>>] +>> Got issues: [ +{"msg":["Failed to resolve implicit argument ?13\nof type Prims.int\nintroduced for Instantiating implicit argument"],"level":"Error","range":{"def":{"file_name":"Bug2021.fst","start_pos":{"line":36,"col":11},"end_pos":{"line":36,"col":12}},"use":{"file_name":"Bug2021.fst","start_pos":{"line":37,"col":13},"end_pos":{"line":37,"col":17}}},"number":66,"ctx":["While checking for top-level effects","While typechecking the top-level declaration `let test5`","While typechecking the top-level declaration `[@@expect_failure] let test5`"]} +>>] +Verified module: Bug2021 +All verification conditions discharged successfully diff --git a/tests/error-messages/Bug2245.fst.json_expected b/tests/error-messages/Bug2245.fst.json_expected new file mode 100644 index 00000000000..1d868b63567 --- /dev/null +++ b/tests/error-messages/Bug2245.fst.json_expected @@ -0,0 +1,5 @@ +>> Got issues: [ +{"msg":["Expected expression of type (Prims.int & Prims.int) & Prims.int\ngot expression a\nof type Prims.int & Prims.int & Prims.int"],"level":"Error","range":{"def":{"file_name":"Bug2245.fst","start_pos":{"line":6,"col":23},"end_pos":{"line":6,"col":24}},"use":{"file_name":"Bug2245.fst","start_pos":{"line":6,"col":23},"end_pos":{"line":6,"col":24}}},"number":189,"ctx":["While typechecking the top-level declaration `let b`","While typechecking the top-level declaration `[@@expect_failure] let b`"]} +>>] +Verified module: Bug2245 +All verification conditions discharged successfully diff --git a/tests/error-messages/Bug2820.fst.json_expected b/tests/error-messages/Bug2820.fst.json_expected new file mode 100644 index 00000000000..6a7d8d3ec05 --- /dev/null +++ b/tests/error-messages/Bug2820.fst.json_expected @@ -0,0 +1,21 @@ +Module after desugaring: +module Bug2820 +Declarations: [ +let basic _ = () <: FStar.Pervasives.Lemma (ensures forall (x: x: _{x + 1 == 1 + x}). Prims.l_True) +] + +Module before type checking: +module Bug2820 +Declarations: [ +let basic _ = () <: FStar.Pervasives.Lemma (ensures forall (x: x: _{x + 1 == 1 + x}). Prims.l_True) +] + +Module after type checking: +module Bug2820 +Declarations: [ +let basic _ = + () <: FStar.Pervasives.Lemma (ensures forall (x: Prims.int{x + 1 == 1 + x}). Prims.l_True) +] + +Verified module: Bug2820 +All verification conditions discharged successfully diff --git a/tests/error-messages/Bug2899.fst.json_expected b/tests/error-messages/Bug2899.fst.json_expected new file mode 100644 index 00000000000..740eddb1cc0 --- /dev/null +++ b/tests/error-messages/Bug2899.fst.json_expected @@ -0,0 +1,14 @@ +>> Got issues: [ +{"msg":["Tactic failed","Tactic got stuck!","Reduction stopped at: FStar.Stubs.Tactics.Result.Success (Prims.admit ())\n \"(((proofstate)))\"","The term contains an `admit`, which will not reduce. Did you mean `tadmit()`?"],"level":"Error","range":{"def":{"file_name":"FStar.Tactics.Effect.fsti","start_pos":{"line":212,"col":48},"end_pos":{"line":212,"col":58}},"use":{"file_name":"Bug2899.fst","start_pos":{"line":7,"col":12},"end_pos":{"line":7,"col":18}}},"number":170,"ctx":["While preprocessing VC with a tactic","While checking for top-level effects","While typechecking the top-level declaration `let test0`","While typechecking the top-level declaration `[@@expect_failure] let test0`"]} +>>] +>> Got issues: [ +{"msg":["Tactic failed","Tactic got stuck!","Reduction stopped at: Prims.admit ()","The term contains an `admit`, which will not reduce. Did you mean `tadmit()`?"],"level":"Error","range":{"def":{"file_name":"FStar.Tactics.Effect.fsti","start_pos":{"line":212,"col":48},"end_pos":{"line":212,"col":58}},"use":{"file_name":"Bug2899.fst","start_pos":{"line":10,"col":12},"end_pos":{"line":10,"col":18}}},"number":170,"ctx":["While preprocessing VC with a tactic","While checking for top-level effects","While typechecking the top-level declaration `let test1`","While typechecking the top-level declaration `[@@expect_failure] let test1`"]} +>>] +>> Got issues: [ +{"msg":["Tactic failed","Tactic got stuck!","Reduction stopped at: FStar.Stubs.Tactics.V2.Builtins.dump (Prims.admit ())\n \"(((proofstate)))\"","The term contains an `admit`, which will not reduce. Did you mean `tadmit()`?"],"level":"Error","range":{"def":{"file_name":"FStar.Tactics.Effect.fsti","start_pos":{"line":212,"col":48},"end_pos":{"line":212,"col":58}},"use":{"file_name":"Bug2899.fst","start_pos":{"line":13,"col":12},"end_pos":{"line":13,"col":18}}},"number":170,"ctx":["While preprocessing VC with a tactic","While checking for top-level effects","While typechecking the top-level declaration `let test2`","While typechecking the top-level declaration `[@@expect_failure] let test2`"]} +>>] +>> Got issues: [ +{"msg":["Tactic failed","Tactic got stuck!","Reduction stopped at: reify (tac ()) \"(((proofstate)))\"",""],"level":"Error","range":{"def":{"file_name":"FStar.Tactics.Effect.fsti","start_pos":{"line":212,"col":48},"end_pos":{"line":212,"col":58}},"use":{"file_name":"Bug2899.fst","start_pos":{"line":17,"col":50},"end_pos":{"line":17,"col":66}}},"number":170,"ctx":["While preprocessing VC with a tactic","While typechecking the top-level declaration `let eval_tactic`","While typechecking the top-level declaration `[@@expect_failure] let eval_tactic`"]} +>>] +Verified module: Bug2899 +All verification conditions discharged successfully diff --git a/tests/error-messages/Bug2953.fst.json_expected b/tests/error-messages/Bug2953.fst.json_expected new file mode 100644 index 00000000000..1d33a583424 --- /dev/null +++ b/tests/error-messages/Bug2953.fst.json_expected @@ -0,0 +1,5 @@ +>> Got issues: [ +{"msg":["When clauses are not yet supported in --verify mode; they will be some day"],"level":"Error","range":{"def":{"file_name":"Bug2953.fst","start_pos":{"line":14,"col":21},"end_pos":{"line":14,"col":31}},"use":{"file_name":"Bug2953.fst","start_pos":{"line":14,"col":21},"end_pos":{"line":14,"col":31}}},"number":236,"ctx":["While typechecking the top-level declaration `let open_view`","While typechecking the top-level declaration `[@@expect_failure] let open_view`"]} +>>] +Verified module: Bug2953 +All verification conditions discharged successfully diff --git a/tests/error-messages/Bug2980.fst.json_expected b/tests/error-messages/Bug2980.fst.json_expected new file mode 100644 index 00000000000..0f65468c6bf --- /dev/null +++ b/tests/error-messages/Bug2980.fst.json_expected @@ -0,0 +1,5 @@ +>> Got issues: [ +{"msg":["Expected expression of type Type0\ngot expression Bug2980.t\nof type _: Prims.int -> Type0"],"level":"Error","range":{"def":{"file_name":"Bug2980.fst","start_pos":{"line":5,"col":4},"end_pos":{"line":5,"col":5}},"use":{"file_name":"Bug2980.fst","start_pos":{"line":5,"col":4},"end_pos":{"line":5,"col":5}}},"number":189,"ctx":["While typechecking the top-level declaration `type Bug2980.t`","While typechecking the top-level declaration `[@@expect_failure] type Bug2980.t`"]} +>>] +Verified module: Bug2980 +All verification conditions discharged successfully diff --git a/tests/error-messages/Bug2981.fst.json_expected b/tests/error-messages/Bug2981.fst.json_expected new file mode 100644 index 00000000000..b372344e166 --- /dev/null +++ b/tests/error-messages/Bug2981.fst.json_expected @@ -0,0 +1,5 @@ +>> Got issues: [ +{"msg":["Bound variable\n'g'\nescapes because of impure applications in the type of\n'check'","Add explicit let-bindings to avoid this"],"level":"Error","range":{"def":{"file_name":"Bug2981.fst","start_pos":{"line":10,"col":10},"end_pos":{"line":10,"col":27}},"use":{"file_name":"Bug2981.fst","start_pos":{"line":10,"col":10},"end_pos":{"line":10,"col":27}}},"number":56,"ctx":["While checking for escaped variables","While typechecking the top-level declaration `let test1`","While typechecking the top-level declaration `[@@expect_failure] let test1`"]} +>>] +Verified module: Bug2981 +All verification conditions discharged successfully diff --git a/tests/error-messages/Bug3099.fst.json_expected b/tests/error-messages/Bug3099.fst.json_expected new file mode 100644 index 00000000000..d079194cd8a --- /dev/null +++ b/tests/error-messages/Bug3099.fst.json_expected @@ -0,0 +1,7 @@ +proof-state: State dump @ depth 0 (X): +Location: Bug3099.fst(11,50-11,64) +Goal 1/1: +(_: Prims.unit), (return_val: Prims.eqtype), (_: return_val == Prims.int & (Prims.int & (Prims.int & Prims.int))), (any_result: Prims.bool), (_: true == any_result), (any_result'0: Prims.logical), (_: Prims.l_True == any_result'0) |- _ : Prims.squash ((1, (2, (3, 4))) = (1, (2, (3, 4)))) + +Verified module: Bug3099 +All verification conditions discharged successfully diff --git a/tests/error-messages/Bug3102.fst.json_expected b/tests/error-messages/Bug3102.fst.json_expected new file mode 100644 index 00000000000..8e4f47922ec --- /dev/null +++ b/tests/error-messages/Bug3102.fst.json_expected @@ -0,0 +1,20 @@ +>> Got issues: [ +{"msg":["Bound variable\n'e1'\nwould escape in the type of this letbinding","Add a type annotation that does not mention it"],"level":"Error","range":{"def":{"file_name":"Bug3102.fst","start_pos":{"line":8,"col":17},"end_pos":{"line":10,"col":9}},"use":{"file_name":"Bug3102.fst","start_pos":{"line":8,"col":17},"end_pos":{"line":10,"col":9}}},"number":56,"ctx":["While checking for escaped variables","While typechecking the top-level declaration `let min`","While typechecking the top-level declaration `[@@expect_failure] let min`"]} +>>] +>> Got issues: [ +{"msg":["(*?u23*) _ g t1 t2 is not equal to the expected type e2"],"level":"Error","range":{"def":{"file_name":"Bug3102.fst","start_pos":{"line":23,"col":4},"end_pos":{"line":23,"col":27}},"use":{"file_name":"Bug3102.fst","start_pos":{"line":23,"col":4},"end_pos":{"line":23,"col":27}}},"number":54,"ctx":["While solving deferred constraints","solve_non_tactic_deferred_constraints","While typechecking the top-level declaration `let test1`","While typechecking the top-level declaration `[@@expect_failure] let test1`"]} +>>] +>> Got issues: [ +{"msg":["(*?u20*) _ is not equal to the expected type e2"],"level":"Error","range":{"def":{"file_name":"Bug3102.fst","start_pos":{"line":29,"col":4},"end_pos":{"line":29,"col":27}},"use":{"file_name":"Bug3102.fst","start_pos":{"line":29,"col":4},"end_pos":{"line":29,"col":27}}},"number":54,"ctx":["While solving deferred constraints","solve_non_tactic_deferred_constraints","While typechecking the top-level declaration `let test2`","While typechecking the top-level declaration `[@@expect_failure] let test2`"]} +>>] +>> Got issues: [ +{"msg":["Bound variable\n'e2'\nwould escape in the type of this letbinding","Add a type annotation that does not mention it"],"level":"Error","range":{"def":{"file_name":"Bug3102.fst","start_pos":{"line":33,"col":29},"end_pos":{"line":35,"col":27}},"use":{"file_name":"Bug3102.fst","start_pos":{"line":33,"col":29},"end_pos":{"line":35,"col":27}}},"number":56,"ctx":["While checking for escaped variables","While typechecking the top-level declaration `let test3`","While typechecking the top-level declaration `[@@expect_failure] let test3`"]} +>>] +>> Got issues: [ +{"msg":["Bound variable\n'z'\nwould escape in the type of this letbinding","Add a type annotation that does not mention it"],"level":"Error","range":{"def":{"file_name":"Bug3102.fst","start_pos":{"line":41,"col":16},"end_pos":{"line":43,"col":8}},"use":{"file_name":"Bug3102.fst","start_pos":{"line":41,"col":16},"end_pos":{"line":43,"col":8}}},"number":56,"ctx":["While checking for escaped variables","While typechecking the top-level declaration `let gg`","While typechecking the top-level declaration `[@@expect_failure] let gg`"]} +>>] +>> Got issues: [ +{"msg":["Bound variable\n'z'\nwould escape in the type of this letbinding","Add a type annotation that does not mention it"],"level":"Error","range":{"def":{"file_name":"Bug3102.fst","start_pos":{"line":49,"col":16},"end_pos":{"line":51,"col":7}},"use":{"file_name":"Bug3102.fst","start_pos":{"line":49,"col":16},"end_pos":{"line":51,"col":7}}},"number":56,"ctx":["While checking for escaped variables","While typechecking the top-level declaration `let g`","While typechecking the top-level declaration `[@@expect_failure] let g`"]} +>>] +Verified module: Bug3102 +All verification conditions discharged successfully diff --git a/tests/error-messages/Bug3145.fst.json_expected b/tests/error-messages/Bug3145.fst.json_expected new file mode 100644 index 00000000000..f2c2f6f2486 --- /dev/null +++ b/tests/error-messages/Bug3145.fst.json_expected @@ -0,0 +1,29 @@ +Module after desugaring: +module Bug3145 +Declarations: [ +let t1 = Prims.int & Prims.int & Prims.int +let t2 = (Prims.int & Prims.int) & Prims.int +let t3 = Prims.int & (Prims.int & Prims.int) +let t4 = Prims.int & Prims.int & Prims.int +] + +Module before type checking: +module Bug3145 +Declarations: [ +let t1 = Prims.int & Prims.int & Prims.int +let t2 = (Prims.int & Prims.int) & Prims.int +let t3 = Prims.int & (Prims.int & Prims.int) +let t4 = Prims.int & Prims.int & Prims.int +] + +Module after type checking: +module Bug3145 +Declarations: [ +let t1 = Prims.int & Prims.int & Prims.int +let t2 = (Prims.int & Prims.int) & Prims.int +let t3 = Prims.int & (Prims.int & Prims.int) +let t4 = Prims.int & Prims.int & Prims.int +] + +Verified module: Bug3145 +All verification conditions discharged successfully diff --git a/tests/error-messages/Bug3227.fst.json_expected b/tests/error-messages/Bug3227.fst.json_expected new file mode 100644 index 00000000000..2da662f37fd --- /dev/null +++ b/tests/error-messages/Bug3227.fst.json_expected @@ -0,0 +1,76 @@ +Module after desugaring: +module Bug3227 +Declarations: [ +type box (a: Type) = { } + +let proj b = x (x (x b)) <: Prims.int +type box2 (a: Type) = | Box2 : x: a -> Bug3227.box2 a + + + +let test b = Box2? b && Box2? b.x +noeq +type boxf (a: Type) = { } + +let test2 r = ff r 5 +noeq +type boxfi (a: Type) = { } + +let test3 r = ff r +let test4 r x = ff r x <: a +] + +Module before type checking: +module Bug3227 +Declarations: [ +type box (a: Type) = { } + +let proj b = x (x (x b)) <: Prims.int +type box2 (a: Type) = | Box2 : x: a -> Bug3227.box2 a + + + +let test b = Box2? b && Box2? b.x +noeq +type boxf (a: Type) = { } + +let test2 r = ff r 5 +noeq +type boxfi (a: Type) = { } + +let test3 r = ff r +let test4 r x = ff r x <: a +] + +Module after type checking: +module Bug3227 +Declarations: [ +type box (a: Type) = { x:a } +val box__uu___haseq: forall (a: Type). {:pattern Prims.hasEq (Bug3227.box a)} + Prims.l_True /\ Prims.hasEq a ==> Prims.hasEq (Bug3227.box a) + + +let proj b = b.x.x.x <: Prims.int +type box2 (a: Type) = | Box2 : x: a -> Bug3227.box2 a +val box2__uu___haseq: forall (a: Type). {:pattern Prims.hasEq (Bug3227.box2 a)} + Prims.l_True /\ Prims.hasEq a ==> Prims.hasEq (Bug3227.box2 a) + + + + +let test b = Box2? b && Box2? b.x +noeq +type boxf (a: Type) = { ff:_: a -> a } + + +let test2 r = r.ff 5 +noeq +type boxfi (a: Type) = { ff:a } + + +let test3 r = r.ff +let test4 r x = r.ff x <: a +] + +Verified module: Bug3227 +All verification conditions discharged successfully diff --git a/tests/error-messages/Bug3292.fst.json_expected b/tests/error-messages/Bug3292.fst.json_expected new file mode 100644 index 00000000000..c8b1ebd3fc7 --- /dev/null +++ b/tests/error-messages/Bug3292.fst.json_expected @@ -0,0 +1,83 @@ +Module after desugaring: +module Bug3292 +Declarations: [ +#set-options "--print_implicits" +let op_Plus #a x y = x, y +let op_Minus #a x y = x, y +let op_Slash #a x y = x, y +let op_Greater #a x y = x, y +let op_Less #a x y = x, y +let op_GreaterEquals #a x y = x, y +let op_LessEquals #a x y = x, y +private +let _ = 1 + 1 +private +let _ = 1 - 1 +private +let _ = 1 / 1 +private +let _ = 1 > 1 +private +let _ = 1 < 1 +private +let _ = 1 >= 1 +private +let _ = 1 <= 1 +] + +Module before type checking: +module Bug3292 +Declarations: [ +#set-options "--print_implicits" +let op_Plus x y = x, y +let op_Minus x y = x, y +let op_Slash x y = x, y +let op_Greater x y = x, y +let op_Less x y = x, y +let op_GreaterEquals x y = x, y +let op_LessEquals x y = x, y +private +let _ = 1 + 1 +private +let _ = 1 - 1 +private +let _ = 1 / 1 +private +let _ = 1 > 1 +private +let _ = 1 < 1 +private +let _ = 1 >= 1 +private +let _ = 1 <= 1 +] + +Module after type checking: +module Bug3292 +Declarations: [ +#set-options "--print_implicits" +let op_Plus #a x y = FStar.Pervasives.Native.Mktuple2 #a #a x y +let op_Minus #a x y = FStar.Pervasives.Native.Mktuple2 #a #a x y +let op_Slash #a x y = FStar.Pervasives.Native.Mktuple2 #a #a x y +let op_Greater #a x y = FStar.Pervasives.Native.Mktuple2 #a #a x y +let op_Less #a x y = FStar.Pervasives.Native.Mktuple2 #a #a x y +let op_GreaterEquals #a x y = FStar.Pervasives.Native.Mktuple2 #a #a x y +let op_LessEquals #a x y = FStar.Pervasives.Native.Mktuple2 #a #a x y +private +let _ = Bug3292.op_Plus #Prims.int 1 1 +private +let _ = 1 - 1 +private +let _ = Bug3292.op_Slash #Prims.int 1 1 +private +let _ = Bug3292.op_Greater #Prims.int 1 1 +private +let _ = Bug3292.op_Less #Prims.int 1 1 +private +let _ = 1 >= 1 +private +let _ = 1 <= 1 +] + +Verified module: Bug3292 +All verification conditions discharged successfully diff --git a/tests/error-messages/Calc.fst.json_expected b/tests/error-messages/Calc.fst.json_expected new file mode 100644 index 00000000000..a67e9e7dba8 --- /dev/null +++ b/tests/error-messages/Calc.fst.json_expected @@ -0,0 +1,35 @@ +>> Got issues: [ +{"msg":["Could not prove that this calc-chain is compatible","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"FStar.Calc.fsti","start_pos":{"line":47,"col":50},"end_pos":{"line":47,"col":55}},"use":{"file_name":"Calc.fst","start_pos":{"line":12,"col":2},"end_pos":{"line":12,"col":13}}},"number":19,"ctx":["While typechecking the top-level declaration `let test_gt_lt_elab`","While typechecking the top-level declaration `[@@expect_failure] let test_gt_lt_elab`"]} +>>] +>> Got issues: [ +{"msg":["Could not prove that this calc-chain is compatible","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"Calc.fst","start_pos":{"line":22,"col":7},"end_pos":{"line":22,"col":11}},"use":{"file_name":"Calc.fst","start_pos":{"line":22,"col":2},"end_pos":{"line":28,"col":3}}},"number":19,"ctx":["While typechecking the top-level declaration `let test_gt_lt`","While typechecking the top-level declaration `[@@expect_failure] let test_gt_lt`"]} +>>] +>> Got issues: [ +{"msg":["Could not prove that this calc-chain is compatible","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"Calc.fst","start_pos":{"line":33,"col":7},"end_pos":{"line":33,"col":10}},"use":{"file_name":"Calc.fst","start_pos":{"line":33,"col":2},"end_pos":{"line":35,"col":3}}},"number":19,"ctx":["While typechecking the top-level declaration `let test_singl`","While typechecking the top-level declaration `[@@expect_failure] let test_singl`"]} +>>] +>> Got issues: [ +{"msg":["Could not prove that this calc-chain is compatible","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"Calc.fst","start_pos":{"line":40,"col":7},"end_pos":{"line":40,"col":10}},"use":{"file_name":"Calc.fst","start_pos":{"line":40,"col":2},"end_pos":{"line":44,"col":3}}},"number":19,"ctx":["While typechecking the top-level declaration `let test`","While typechecking the top-level declaration `[@@expect_failure] let test`"]} +>>] +>> Got issues: [ +{"msg":["Subtyping check failed","Expected type Prims.squash (1 == 2)\ngot type Prims.unit","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"Calc.fst","start_pos":{"line":51,"col":3},"end_pos":{"line":51,"col":5}},"use":{"file_name":"Calc.fst","start_pos":{"line":51,"col":6},"end_pos":{"line":51,"col":8}}},"number":19,"ctx":["While typechecking the top-level declaration `let fail1`","While typechecking the top-level declaration `[@@expect_failure] let fail1`"]} +>>] +>> Got issues: [ +{"msg":["Subtyping check failed","Expected type Prims.squash (2 == 3)\ngot type Prims.unit","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"Calc.fst","start_pos":{"line":65,"col":3},"end_pos":{"line":65,"col":5}},"use":{"file_name":"Calc.fst","start_pos":{"line":65,"col":6},"end_pos":{"line":65,"col":8}}},"number":19,"ctx":["While typechecking the top-level declaration `let fail2`","While typechecking the top-level declaration `[@@expect_failure] let fail2`"]} +>>] +>> Got issues: [ +{"msg":["Subtyping check failed","Expected type Prims.squash (3 == 4)\ngot type Prims.unit","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"Calc.fst","start_pos":{"line":79,"col":3},"end_pos":{"line":79,"col":5}},"use":{"file_name":"Calc.fst","start_pos":{"line":79,"col":6},"end_pos":{"line":79,"col":8}}},"number":19,"ctx":["While typechecking the top-level declaration `let fail3`","While typechecking the top-level declaration `[@@expect_failure] let fail3`"]} +>>] +>> Got issues: [ +{"msg":["Subtyping check failed","Expected type Prims.squash q\ngot type Prims.unit","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"Calc.fst","start_pos":{"line":91,"col":20},"end_pos":{"line":91,"col":21}},"use":{"file_name":"Calc.fst","start_pos":{"line":93,"col":42},"end_pos":{"line":93,"col":44}}},"number":19,"ctx":["While typechecking the top-level declaration `let test_impl_elab`","While typechecking the top-level declaration `[@@expect_failure] let test_impl_elab`"]} +>>] +>> Got issues: [ +{"msg":["Subtyping check failed","Expected type Prims.squash q\ngot type Prims.unit","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"Calc.fst","start_pos":{"line":101,"col":4},"end_pos":{"line":101,"col":5}},"use":{"file_name":"Calc.fst","start_pos":{"line":100,"col":10},"end_pos":{"line":100,"col":12}}},"number":19,"ctx":["While typechecking the top-level declaration `let test_impl`","While typechecking the top-level declaration `[@@expect_failure] let test_impl`"]} +>>] +>> Got issues: [ +{"msg":["Assertion failed","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"Calc.fst","start_pos":{"line":105,"col":12},"end_pos":{"line":105,"col":17}},"use":{"file_name":"Calc.fst","start_pos":{"line":114,"col":17},"end_pos":{"line":114,"col":25}}},"number":19,"ctx":["While typechecking the top-level declaration `let test_1763_elab`","While typechecking the top-level declaration `[@@expect_failure] let test_1763_elab`"]} +>>] +>> Got issues: [ +{"msg":["Assertion failed","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"Calc.fst","start_pos":{"line":105,"col":12},"end_pos":{"line":105,"col":17}},"use":{"file_name":"Calc.fst","start_pos":{"line":121,"col":9},"end_pos":{"line":121,"col":17}}},"number":19,"ctx":["While typechecking the top-level declaration `let test_1763`","While typechecking the top-level declaration `[@@expect_failure] let test_1763`"]} +>>] +Verified module: Calc +All verification conditions discharged successfully diff --git a/tests/error-messages/CalcImpl.fst.json_expected b/tests/error-messages/CalcImpl.fst.json_expected new file mode 100644 index 00000000000..a966d0fb0a4 --- /dev/null +++ b/tests/error-messages/CalcImpl.fst.json_expected @@ -0,0 +1,161 @@ +Module after desugaring: +module CalcImpl +Declarations: [ +assume +val p:Prims.prop +assume +val q:Prims.prop +assume +val lem (_: Prims.unit) : FStar.Pervasives.Lemma (requires CalcImpl.p) (ensures CalcImpl.q) +let test _ = + FStar.Calc.calc_finish (fun x3 y4 -> x3 ==> y4 <: Type0) + (fun _ -> + FStar.Calc.calc_step (fun x5 y6 -> x5 ==> y6 <: Type0) + CalcImpl.q + (fun _ -> FStar.Calc.calc_init CalcImpl.p) + (fun _ -> FStar.Calc.calc_push_impl (fun _ -> CalcImpl.lem ()))) +let any p q = Prims.l_True +let test3 _ = + FStar.Calc.calc_finish (fun x18 y19 -> CalcImpl.any x18 y19 <: Type0) + (fun _ -> + FStar.Calc.calc_step (fun x24 y25 -> x24 <==> y25 <: Type0) + CalcImpl.q + (fun _ -> + FStar.Calc.calc_step (fun x22 y23 -> x22 ==> y23 <: Type0) + (CalcImpl.q /\ CalcImpl.q) + (fun _ -> + FStar.Calc.calc_step (fun x20 y21 -> x20 <==> y21 <: Type0) + CalcImpl.p + (fun _ -> FStar.Calc.calc_init (CalcImpl.p /\ CalcImpl.p)) + (fun _ -> ())) + (fun _ -> FStar.Calc.calc_push_impl (fun _ -> CalcImpl.lem ()))) + (fun _ -> ())) +let test4 _ = + FStar.Calc.calc_finish (fun x43 y44 -> CalcImpl.any x43 y44 <: Type0) + (fun _ -> + FStar.Calc.calc_step (fun x49 y50 -> x49 <==> y50 <: Type0) + CalcImpl.q + (fun _ -> + FStar.Calc.calc_step (fun x47 y48 -> x47 ==> y48 <: Type0) + (CalcImpl.q /\ CalcImpl.q) + (fun _ -> + FStar.Calc.calc_step (fun x45 y46 -> x45 <==> y46 <: Type0) + CalcImpl.p + (fun _ -> FStar.Calc.calc_init (CalcImpl.p /\ CalcImpl.p)) + (fun _ -> ())) + (fun _ -> FStar.Calc.calc_push_impl (fun _ -> CalcImpl.lem ()))) + (fun _ -> ())) +let op_Equals_Equals_Greater = Prims.op_LessThan +let test5 _ = + FStar.Calc.calc_finish (fun x68 y69 -> x68 ==> y69 <: Type0) + (fun _ -> + FStar.Calc.calc_step (fun x70 y71 -> x70 ==> y71 <: Type0) + 2 + (fun _ -> FStar.Calc.calc_init 1) + (fun _ -> ())) +] + +Module before type checking: +module CalcImpl +Declarations: [ +assume +val p:Prims.prop +assume +val q:Prims.prop +assume +val lem (_: Prims.unit) : FStar.Pervasives.Lemma (requires CalcImpl.p) (ensures CalcImpl.q) +let test _ = + FStar.Calc.calc_finish (fun x3 y4 -> x3 ==> y4 <: Type0) + (fun _ -> + FStar.Calc.calc_step (fun x5 y6 -> x5 ==> y6 <: Type0) + CalcImpl.q + (fun _ -> FStar.Calc.calc_init CalcImpl.p) + (fun _ -> FStar.Calc.calc_push_impl (fun _ -> CalcImpl.lem ()))) +let any p q = Prims.l_True +let test3 _ = + FStar.Calc.calc_finish (fun x18 y19 -> CalcImpl.any x18 y19 <: Type0) + (fun _ -> + FStar.Calc.calc_step (fun x24 y25 -> x24 <==> y25 <: Type0) + CalcImpl.q + (fun _ -> + FStar.Calc.calc_step (fun x22 y23 -> x22 ==> y23 <: Type0) + (CalcImpl.q /\ CalcImpl.q) + (fun _ -> + FStar.Calc.calc_step (fun x20 y21 -> x20 <==> y21 <: Type0) + CalcImpl.p + (fun _ -> FStar.Calc.calc_init (CalcImpl.p /\ CalcImpl.p)) + (fun _ -> ())) + (fun _ -> FStar.Calc.calc_push_impl (fun _ -> CalcImpl.lem ()))) + (fun _ -> ())) +let test4 _ = + FStar.Calc.calc_finish (fun x43 y44 -> CalcImpl.any x43 y44 <: Type0) + (fun _ -> + FStar.Calc.calc_step (fun x49 y50 -> x49 <==> y50 <: Type0) + CalcImpl.q + (fun _ -> + FStar.Calc.calc_step (fun x47 y48 -> x47 ==> y48 <: Type0) + (CalcImpl.q /\ CalcImpl.q) + (fun _ -> + FStar.Calc.calc_step (fun x45 y46 -> x45 <==> y46 <: Type0) + CalcImpl.p + (fun _ -> FStar.Calc.calc_init (CalcImpl.p /\ CalcImpl.p)) + (fun _ -> ())) + (fun _ -> FStar.Calc.calc_push_impl (fun _ -> CalcImpl.lem ()))) + (fun _ -> ())) +let op_Equals_Equals_Greater = Prims.op_LessThan +let test5 _ = + FStar.Calc.calc_finish (fun x68 y69 -> x68 ==> y69 <: Type0) + (fun _ -> + FStar.Calc.calc_step (fun x70 y71 -> x70 ==> y71 <: Type0) + 2 + (fun _ -> FStar.Calc.calc_init 1) + (fun _ -> ())) +] + +Module after type checking: +module CalcImpl +Declarations: [ +assume +val p:Prims.prop +assume +val q:Prims.prop +assume +val lem (_: Prims.unit) : FStar.Pervasives.Lemma (requires CalcImpl.p) (ensures CalcImpl.q) +let test _ = + calc ( ==> ) { + CalcImpl.p; + ( ==> ) { FStar.Calc.calc_push_impl (fun _ -> CalcImpl.lem ()) } + CalcImpl.q; + } +let any p q = Prims.l_True +let test3 _ = + calc (CalcImpl.any) { + CalcImpl.p /\ CalcImpl.p; + ( <==> ) { () } + CalcImpl.p; + ( ==> ) { FStar.Calc.calc_push_impl (fun _ -> CalcImpl.lem ()) } + CalcImpl.q /\ CalcImpl.q; + ( <==> ) { () } + CalcImpl.q; + } +let test4 _ = + calc (CalcImpl.any) { + CalcImpl.p /\ CalcImpl.p; + ( <==> ) { () } + CalcImpl.p; + ( ==> ) { FStar.Calc.calc_push_impl (fun _ -> CalcImpl.lem ()) } + CalcImpl.q /\ CalcImpl.q; + ( <==> ) { () } + CalcImpl.q; + } +let op_Equals_Equals_Greater = Prims.op_LessThan +let test5 _ = + calc ( ==> ) { + 1; + ( ==> ) { () } + 2; + } +] + +Verified module: CalcImpl +All verification conditions discharged successfully diff --git a/tests/error-messages/Coercions.fst.json_expected b/tests/error-messages/Coercions.fst.json_expected new file mode 100644 index 00000000000..5a3cfd212bc --- /dev/null +++ b/tests/error-messages/Coercions.fst.json_expected @@ -0,0 +1,23 @@ +>> Got issues: [ +{"msg":["Computed type Prims.int\nand effect GTot\nis not compatible with the annotated type Prims.int\nand effect Tot"],"level":"Error","range":{"def":{"file_name":"Coercions.fst","start_pos":{"line":6,"col":38},"end_pos":{"line":6,"col":39}},"use":{"file_name":"Coercions.fst","start_pos":{"line":6,"col":38},"end_pos":{"line":6,"col":39}}},"number":34,"ctx":["While typechecking the top-level declaration `let test0`","While typechecking the top-level declaration `[@@expect_failure] let test0`"]} +>>] +>> Got issues: [ +{"msg":["Computed type 'a\nand effect GTot\nis not compatible with the annotated type 'a\nand effect Tot"],"level":"Error","range":{"def":{"file_name":"Coercions.fst","start_pos":{"line":19,"col":37},"end_pos":{"line":19,"col":38}},"use":{"file_name":"Coercions.fst","start_pos":{"line":19,"col":37},"end_pos":{"line":19,"col":38}}},"number":34,"ctx":["While typechecking the top-level declaration `let test0'`","While typechecking the top-level declaration `[@@expect_failure] let test0'`"]} +>>] +>> Got issues: [ +{"msg":["Subtyping check failed","Expected type Prims.nat\ngot type Prims.int","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"Prims.fst","start_pos":{"line":680,"col":18},"end_pos":{"line":680,"col":24}},"use":{"file_name":"Coercions.fst","start_pos":{"line":71,"col":4},"end_pos":{"line":71,"col":8}}},"number":19,"ctx":["While typechecking the top-level declaration `let test_literal_bad`","While typechecking the top-level declaration `[@@expect_failure] let test_literal_bad`"]} +>>] +>> Got issues: [ +{"msg":["Subtyping check failed","Expected type Prims.nat\ngot type Prims.int","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"Prims.fst","start_pos":{"line":680,"col":18},"end_pos":{"line":680,"col":24}},"use":{"file_name":"Coercions.fst","start_pos":{"line":74,"col":49},"end_pos":{"line":74,"col":57}}},"number":19,"ctx":["While typechecking the top-level declaration `let test_int_nat_1`","While typechecking the top-level declaration `[@@expect_failure] let test_int_nat_1`"]} +>>] +>> Got issues: [ +{"msg":["Subtyping check failed","Expected type Prims.nat\ngot type Prims.int","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"Prims.fst","start_pos":{"line":680,"col":18},"end_pos":{"line":680,"col":24}},"use":{"file_name":"Coercions.fst","start_pos":{"line":76,"col":55},"end_pos":{"line":76,"col":56}}},"number":19,"ctx":["While typechecking the top-level declaration `let test_int_nat_2`","While typechecking the top-level declaration `[@@expect_failure] let test_int_nat_2`"]} +>>] +>> Got issues: [ +{"msg":["Subtyping check failed","Expected type Prims.nat\ngot type Prims.int","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"Prims.fst","start_pos":{"line":680,"col":18},"end_pos":{"line":680,"col":24}},"use":{"file_name":"Coercions.fst","start_pos":{"line":78,"col":50},"end_pos":{"line":78,"col":51}}},"number":19,"ctx":["While typechecking the top-level declaration `let test_int_nat_1'`","While typechecking the top-level declaration `[@@expect_failure] let test_int_nat_1'`"]} +>>] +>> Got issues: [ +{"msg":["Subtyping check failed","Expected type Prims.nat\ngot type Prims.int","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"Prims.fst","start_pos":{"line":680,"col":18},"end_pos":{"line":680,"col":24}},"use":{"file_name":"Coercions.fst","start_pos":{"line":80,"col":51},"end_pos":{"line":80,"col":52}}},"number":19,"ctx":["While typechecking the top-level declaration `let test_int_nat_2'`","While typechecking the top-level declaration `[@@expect_failure] let test_int_nat_2'`"]} +>>] +Verified module: Coercions +All verification conditions discharged successfully diff --git a/tests/error-messages/DTuples.fst.json_expected b/tests/error-messages/DTuples.fst.json_expected new file mode 100644 index 00000000000..060c892d55e --- /dev/null +++ b/tests/error-messages/DTuples.fst.json_expected @@ -0,0 +1,143 @@ +Module after desugaring: +module DTuples +Declarations: [ +private +let _ = a: Prims.int & b: Prims.int{a > b} & c: Prims.int{c == a + b} +private +let _ = a: Prims.int & b: Prims.int{a > b} & c: Prims.int{c == a + b} +private +let _ = a: Prims.int & z: b: Prims.int{a > b} & c: Prims.int{c == a + z} +private +let _ = a: Prims.int & b: Prims.int{a > b} & Prims.unit +private +let _ = a: Prims.int & b: Prims.int & Prims.int +private +let _ = a: Prims.int & b: Prims.int & c: Prims.int & Prims.int +private +let _ = a: Prims.int & b: Prims.int & c: Prims.int & d: Prims.int & Prims.int +private +let _ = Prims.int & Prims.int +private +let _ = Prims.int & Prims.int & Prims.int +private +let _ = Prims.int & Prims.int & Prims.int & Prims.int +private +let _ = Prims.int & Prims.int & Prims.int & Prims.int & Prims.int +private +let _ = Prims.int & (Prims.int & Prims.int) +private +let _ = Prims.int & (Prims.int & (Prims.int & Prims.int)) +private +let _ = Prims.int & (Prims.int & (Prims.int & (Prims.int & Prims.int))) +private +let _ = FStar.Pervasives.Native.tuple4 Prims.int Prims.bool +private +let _ = Prims.int & Prims.bool & Prims.string & Prims.unit +private +let _ = FStar.Pervasives.dtuple4 Prims.int (fun _ -> Prims.bool) +let t1 = 1, 2, 3 +let t2 = (1, 2), 3 +let t3 = 1, (2, 3) +let d1 = (| 1, 2, 3 |) +let d2 = (| (1, 2), 3 |) +let d3 = (| 1, (2, 3) |) +let dd2 = (| (| 1, 2 |), 3 |) +let dd3 = (| 1, (| 2, 3 |) |) +] + +Module before type checking: +module DTuples +Declarations: [ +private +let _ = a: Prims.int & b: Prims.int{a > b} & c: Prims.int{c == a + b} +private +let _ = a: Prims.int & b: Prims.int{a > b} & c: Prims.int{c == a + b} +private +let _ = a: Prims.int & z: b: Prims.int{a > b} & c: Prims.int{c == a + z} +private +let _ = a: Prims.int & b: Prims.int{a > b} & Prims.unit +private +let _ = a: Prims.int & b: Prims.int & Prims.int +private +let _ = a: Prims.int & b: Prims.int & c: Prims.int & Prims.int +private +let _ = a: Prims.int & b: Prims.int & c: Prims.int & d: Prims.int & Prims.int +private +let _ = Prims.int & Prims.int +private +let _ = Prims.int & Prims.int & Prims.int +private +let _ = Prims.int & Prims.int & Prims.int & Prims.int +private +let _ = Prims.int & Prims.int & Prims.int & Prims.int & Prims.int +private +let _ = Prims.int & (Prims.int & Prims.int) +private +let _ = Prims.int & (Prims.int & (Prims.int & Prims.int)) +private +let _ = Prims.int & (Prims.int & (Prims.int & (Prims.int & Prims.int))) +private +let _ = FStar.Pervasives.Native.tuple4 Prims.int Prims.bool +private +let _ = Prims.int & Prims.bool & Prims.string & Prims.unit +private +let _ = FStar.Pervasives.dtuple4 Prims.int (fun _ -> Prims.bool) +let t1 = 1, 2, 3 +let t2 = (1, 2), 3 +let t3 = 1, (2, 3) +let d1 = (| 1, 2, 3 |) +let d2 = (| (1, 2), 3 |) +let d3 = (| 1, (2, 3) |) +let dd2 = (| (| 1, 2 |), 3 |) +let dd3 = (| 1, (| 2, 3 |) |) +] + +Module after type checking: +module DTuples +Declarations: [ +private +let _ = a: Prims.int & b: Prims.int{a > b} & c: Prims.int{c == a + b} +private +let _ = a: Prims.int & b: Prims.int{a > b} & c: Prims.int{c == a + b} +private +let _ = a: Prims.int & z: b: Prims.int{a > b} & c: Prims.int{c == a + z} +private +let _ = a: Prims.int & b: Prims.int{a > b} & Prims.unit +private +let _ = a: Prims.int & b: Prims.int & Prims.int +private +let _ = a: Prims.int & b: Prims.int & c: Prims.int & Prims.int +private +let _ = a: Prims.int & b: Prims.int & c: Prims.int & d: Prims.int & Prims.int +private +let _ = Prims.int & Prims.int +private +let _ = Prims.int & Prims.int & Prims.int +private +let _ = Prims.int & Prims.int & Prims.int & Prims.int +private +let _ = Prims.int & Prims.int & Prims.int & Prims.int & Prims.int +private +let _ = Prims.int & (Prims.int & Prims.int) +private +let _ = Prims.int & (Prims.int & (Prims.int & Prims.int)) +private +let _ = Prims.int & (Prims.int & (Prims.int & (Prims.int & Prims.int))) +private +let _ = FStar.Pervasives.Native.tuple4 Prims.int Prims.bool +private +let _ = Prims.int & Prims.bool & Prims.string & Prims.unit +private +let _ = FStar.Pervasives.dtuple4 Prims.int (fun _ -> Prims.bool) +let t1 = 1, 2, 3 +let t2 = (1, 2), 3 +let t3 = 1, (2, 3) +let d1 = (| 1, 2, 3 |) +let d2 = (| (1, 2), 3 |) +let d3 = (| 1, (2, 3) |) +let dd2 = (| (| 1, 2 |), 3 |) +let dd3 = (| 1, (| 2, 3 |) |) +] + +Verified module: DTuples +All verification conditions discharged successfully diff --git a/tests/error-messages/DecreasesTypeWarning.fst.json_expected b/tests/error-messages/DecreasesTypeWarning.fst.json_expected new file mode 100644 index 00000000000..24cec7c5ef0 --- /dev/null +++ b/tests/error-messages/DecreasesTypeWarning.fst.json_expected @@ -0,0 +1,4 @@ +{"msg":["In the decreases clause for this function, the SMT solver may not be able to\nprove that the types of\n x (bound in DecreasesTypeWarning.fst(3,11-3,12))\nand xs (bound in DecreasesTypeWarning.fst(5,7-5,9))\nare equal.","The type of the first term is: Prims.nat","The type of the second term is: Prims.list Prims.nat","If the proof fails, try annotating these with the same type."],"level":"Warning","range":{"def":{"file_name":"DecreasesTypeWarning.fst","start_pos":{"line":3,"col":11},"end_pos":{"line":3,"col":12}},"use":{"file_name":"DecreasesTypeWarning.fst","start_pos":{"line":3,"col":11},"end_pos":{"line":3,"col":12}}},"number":290,"ctx":["While typechecking the top-level declaration `let rec f and g`"]} +{"msg":["In the decreases clause for this function, the SMT solver may not be able to\nprove that the types of\n xs (bound in DecreasesTypeWarning.fst(5,7-5,9))\nand x (bound in DecreasesTypeWarning.fst(3,11-3,12))\nare equal.","The type of the first term is: Prims.list Prims.nat","The type of the second term is: Prims.nat","If the proof fails, try annotating these with the same type."],"level":"Warning","range":{"def":{"file_name":"DecreasesTypeWarning.fst","start_pos":{"line":5,"col":7},"end_pos":{"line":5,"col":9}},"use":{"file_name":"DecreasesTypeWarning.fst","start_pos":{"line":5,"col":7},"end_pos":{"line":5,"col":9}}},"number":290,"ctx":["While typechecking the top-level declaration `let rec f and g`"]} +Verified module: DecreasesTypeWarning +All verification conditions discharged successfully diff --git a/tests/error-messages/DumpUvarsOf.fst.json_expected b/tests/error-messages/DumpUvarsOf.fst.json_expected new file mode 100644 index 00000000000..27641e26fca --- /dev/null +++ b/tests/error-messages/DumpUvarsOf.fst.json_expected @@ -0,0 +1,15 @@ +proof-state: State dump @ depth 1 (1): +Location: DumpUvarsOf.fst(13,4-13,12) +Goal 1/1: +(_: Prims.unit) |- _ : Prims.squash (DumpUvarsOf.p (*?u89*) _ (*?u90*) _) + +proof-state: State dump @ depth 1 (2): +Location: DumpUvarsOf.fst(14,4-14,36) +Goal 1/2: +(_: Prims.unit) |- _ : Prims.bool + +Goal 2/2: +(_: Prims.unit) |- _ : Prims.int + +Verified module: DumpUvarsOf +All verification conditions discharged successfully diff --git a/tests/error-messages/Erasable.fst.json_expected b/tests/error-messages/Erasable.fst.json_expected new file mode 100644 index 00000000000..08c75ccad90 --- /dev/null +++ b/tests/error-messages/Erasable.fst.json_expected @@ -0,0 +1,23 @@ +>> Got issues: [ +{"msg":["Incompatible attributes and qualifiers: erasable types do not support decidable\nequality and must be marked `noeq`."],"level":"Error","range":{"def":{"file_name":"Erasable.fst","start_pos":{"line":6,"col":0},"end_pos":{"line":8,"col":17}},"use":{"file_name":"Erasable.fst","start_pos":{"line":6,"col":0},"end_pos":{"line":8,"col":17}}},"number":162,"ctx":["While typechecking the top-level declaration `type Erasable.t0`","While typechecking the top-level declaration `[@@expect_failure] type Erasable.t0`"]} +>>] +>> Got issues: [ +{"msg":["Computed type Prims.int\nand effect Prims.GHOST\nis not compatible with the annotated type Prims.int\nand effect Tot"],"level":"Error","range":{"def":{"file_name":"Erasable.fst","start_pos":{"line":18,"col":2},"end_pos":{"line":20,"col":15}},"use":{"file_name":"Erasable.fst","start_pos":{"line":18,"col":2},"end_pos":{"line":20,"col":15}}},"number":34,"ctx":["While typechecking the top-level declaration `let test0_fail`","While typechecking the top-level declaration `[@@expect_failure] let test0_fail`"]} +>>] +>> Got issues: [ +{"msg":["Computed type Prims.int\nand effect GTot\nis not compatible with the annotated type Prims.int\nand effect Tot"],"level":"Error","range":{"def":{"file_name":"Erasable.fst","start_pos":{"line":28,"col":42},"end_pos":{"line":28,"col":52}},"use":{"file_name":"Erasable.fst","start_pos":{"line":28,"col":42},"end_pos":{"line":28,"col":52}}},"number":34,"ctx":["While typechecking the top-level declaration `let test1_fail`","While typechecking the top-level declaration `[@@expect_failure] let test1_fail`"]} +>>] +>> Got issues: [ +{"msg":["Illegal attribute: the `erasable` attribute is only permitted on inductive type\ndefinitions and abbreviations for non-informative types.","The term\nPrims.nat\nis considered informative."],"level":"Error","range":{"def":{"file_name":"Erasable.fst","start_pos":{"line":41,"col":12},"end_pos":{"line":41,"col":15}},"use":{"file_name":"Erasable.fst","start_pos":{"line":41,"col":12},"end_pos":{"line":41,"col":15}}},"number":162,"ctx":["While typechecking the top-level declaration `let e_nat`","While typechecking the top-level declaration `[@@expect_failure] let e_nat`"]} +>>] +>> Got issues: [ +{"msg":["Mismatch of attributes between declaration and definition.","Declaration is marked `erasable` but the definition is not."],"level":"Error","range":{"def":{"file_name":"Erasable.fst","start_pos":{"line":52,"col":0},"end_pos":{"line":52,"col":17}},"use":{"file_name":"Erasable.fst","start_pos":{"line":52,"col":0},"end_pos":{"line":52,"col":17}}},"number":162,"ctx":["While typechecking the top-level declaration `let e_nat_2`","While typechecking the top-level declaration `[@@expect_failure] let e_nat_2`"]} +>>] +>> Got issues: [ +{"msg":["Mismatch of attributes between declaration and definition.","Declaration is marked `erasable` but the definition is not."],"level":"Error","range":{"def":{"file_name":"Erasable.fst","start_pos":{"line":59,"col":0},"end_pos":{"line":59,"col":29}},"use":{"file_name":"Erasable.fst","start_pos":{"line":59,"col":0},"end_pos":{"line":59,"col":29}}},"number":162,"ctx":["While typechecking the top-level declaration `type Erasable.e_nat_3`","While typechecking the top-level declaration `[@@expect_failure] type Erasable.e_nat_3`"]} +>>] +{"msg":["Erasable.e_nat_2\nis declared but no definition was found","Add an 'assume' if this is intentional"],"level":"Warning","range":{"def":{"file_name":"Erasable.fst","start_pos":{"line":49,"col":4},"end_pos":{"line":49,"col":11}},"use":{"file_name":"Erasable.fst","start_pos":{"line":49,"col":4},"end_pos":{"line":49,"col":11}}},"number":240,"ctx":["While desugaring module Erasable"]} +{"msg":["Erasable.e_nat_3\nis declared but no definition was found","Add an 'assume' if this is intentional"],"level":"Warning","range":{"def":{"file_name":"Erasable.fst","start_pos":{"line":56,"col":4},"end_pos":{"line":56,"col":11}},"use":{"file_name":"Erasable.fst","start_pos":{"line":56,"col":4},"end_pos":{"line":56,"col":11}}},"number":240,"ctx":["While desugaring module Erasable"]} +{"msg":["Missing definitions in module Erasable:\n e_nat_2\n e_nat_3"],"level":"Warning","range":{"def":{"file_name":"dummy","start_pos":{"line":0,"col":0},"end_pos":{"line":0,"col":0}},"use":{"file_name":"Erasable.fst","start_pos":{"line":73,"col":0},"end_pos":{"line":73,"col":19}}},"number":240,"ctx":[]} +Verified module: Erasable +All verification conditions discharged successfully diff --git a/tests/error-messages/ExpectFailure.fst.json_expected b/tests/error-messages/ExpectFailure.fst.json_expected new file mode 100644 index 00000000000..e867cfa3742 --- /dev/null +++ b/tests/error-messages/ExpectFailure.fst.json_expected @@ -0,0 +1,23 @@ +>> Got issues: [ +{"msg":["Expected expression of type Prims.int\ngot expression 'a'\nof type FStar.Char.char"],"level":"Error","range":{"def":{"file_name":"ExpectFailure.fst","start_pos":{"line":20,"col":12},"end_pos":{"line":20,"col":15}},"use":{"file_name":"ExpectFailure.fst","start_pos":{"line":20,"col":12},"end_pos":{"line":20,"col":15}}},"number":189,"ctx":["While typechecking the top-level declaration `let uu___0`","While typechecking the top-level declaration `[@@expect_failure] let uu___0`"]} +>>] +>> Got issues: [ +{"msg":["Expected expression of type Prims.int\ngot expression 'a'\nof type FStar.Char.char"],"level":"Error","range":{"def":{"file_name":"ExpectFailure.fst","start_pos":{"line":24,"col":12},"end_pos":{"line":24,"col":15}},"use":{"file_name":"ExpectFailure.fst","start_pos":{"line":24,"col":12},"end_pos":{"line":24,"col":15}}},"number":189,"ctx":["While typechecking the top-level declaration `let uu___1`","While typechecking the top-level declaration `[@@expect_failure] let uu___1`"]} +>>] +>> Got issues: [ +{"msg":["Assertion failed","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"ExpectFailure.fst","start_pos":{"line":28,"col":15},"end_pos":{"line":28,"col":20}},"use":{"file_name":"ExpectFailure.fst","start_pos":{"line":28,"col":8},"end_pos":{"line":28,"col":14}}},"number":19,"ctx":["While checking for top-level effects","While typechecking the top-level declaration `let uu___2`","While typechecking the top-level declaration `[@@expect_failure] let uu___2`"]} +>>] +>> Got issues: [ +{"msg":["Assertion failed","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"ExpectFailure.fst","start_pos":{"line":30,"col":15},"end_pos":{"line":30,"col":20}},"use":{"file_name":"ExpectFailure.fst","start_pos":{"line":30,"col":8},"end_pos":{"line":30,"col":14}}},"number":19,"ctx":["While checking for top-level effects","While typechecking the top-level declaration `let uu___3`","While typechecking the top-level declaration `[@@expect_failure] let uu___3`"]} +>>] +>> Got issues: [ +{"msg":["Expected expression of type Prims.int\ngot expression 'a'\nof type FStar.Char.char"],"level":"Error","range":{"def":{"file_name":"ExpectFailure.fst","start_pos":{"line":43,"col":12},"end_pos":{"line":43,"col":15}},"use":{"file_name":"ExpectFailure.fst","start_pos":{"line":43,"col":12},"end_pos":{"line":43,"col":15}}},"number":189,"ctx":["While typechecking the top-level declaration `let uu___6`","While typechecking the top-level declaration `[@@expect_failure] let uu___6`"]} +>>] +>> Got issues: [ +{"msg":["Expected expression of type Prims.int\ngot expression 'a'\nof type FStar.Char.char"],"level":"Error","range":{"def":{"file_name":"ExpectFailure.fst","start_pos":{"line":50,"col":12},"end_pos":{"line":50,"col":15}},"use":{"file_name":"ExpectFailure.fst","start_pos":{"line":50,"col":12},"end_pos":{"line":50,"col":15}}},"number":189,"ctx":["While typechecking the top-level declaration `let uu___7`","While typechecking the top-level declaration `[@@expect_failure] let uu___7`"]} +>>] +>> Got issues: [ +{"msg":["Expected expression of type Prims.int\ngot expression 'a'\nof type FStar.Char.char"],"level":"Error","range":{"def":{"file_name":"ExpectFailure.fst","start_pos":{"line":61,"col":12},"end_pos":{"line":61,"col":15}},"use":{"file_name":"ExpectFailure.fst","start_pos":{"line":61,"col":12},"end_pos":{"line":61,"col":15}}},"number":189,"ctx":["While typechecking the top-level declaration `let uu___8`","While typechecking the top-level declaration `[@@expect_failure] let uu___8`"]} +>>] +Verified module: ExpectFailure +All verification conditions discharged successfully diff --git a/tests/error-messages/GhostImplicits.fst.json_expected b/tests/error-messages/GhostImplicits.fst.json_expected new file mode 100644 index 00000000000..e4970e81d41 --- /dev/null +++ b/tests/error-messages/GhostImplicits.fst.json_expected @@ -0,0 +1,5 @@ +>> Got issues: [ +{"msg":["Computed type Prims.nat\nand effect Prims.GHOST\nis not compatible with the annotated type Prims.nat\nand effect Tot"],"level":"Error","range":{"def":{"file_name":"GhostImplicits.fst","start_pos":{"line":25,"col":54},"end_pos":{"line":25,"col":57}},"use":{"file_name":"GhostImplicits.fst","start_pos":{"line":25,"col":54},"end_pos":{"line":25,"col":57}}},"number":34,"ctx":["While typechecking the top-level declaration `let test3`","While typechecking the top-level declaration `[@@expect_failure] let test3`"]} +>>] +Verified module: GhostImplicits +All verification conditions discharged successfully diff --git a/tests/error-messages/IfCond.fst.json_expected b/tests/error-messages/IfCond.fst.json_expected new file mode 100644 index 00000000000..ee8e7bcc5f3 --- /dev/null +++ b/tests/error-messages/IfCond.fst.json_expected @@ -0,0 +1,5 @@ +>> Got issues: [ +{"msg":["Expected expression of type Prims.bool\ngot expression 2\nof type Prims.int"],"level":"Error","range":{"def":{"file_name":"IfCond.fst","start_pos":{"line":4,"col":17},"end_pos":{"line":4,"col":18}},"use":{"file_name":"IfCond.fst","start_pos":{"line":4,"col":17},"end_pos":{"line":4,"col":18}}},"number":189,"ctx":["While typechecking the top-level declaration `let test`","While typechecking the top-level declaration `[@@expect_failure] let test`"]} +>>] +Verified module: IfCond +All verification conditions discharged successfully diff --git a/tests/error-messages/IfThen.fst.json_expected b/tests/error-messages/IfThen.fst.json_expected new file mode 100644 index 00000000000..62146ad901a --- /dev/null +++ b/tests/error-messages/IfThen.fst.json_expected @@ -0,0 +1,5 @@ +>> Got issues: [ +{"msg":["Expected expression of type Prims.int\ngot expression ()\nof type Prims.unit"],"level":"Error","range":{"def":{"file_name":"IfThen.fst","start_pos":{"line":5,"col":2},"end_pos":{"line":6,"col":6}},"use":{"file_name":"IfThen.fst","start_pos":{"line":5,"col":2},"end_pos":{"line":6,"col":6}}},"number":189,"ctx":["While typechecking the top-level declaration `let test`","While typechecking the top-level declaration `[@@expect_failure] let test`"]} +>>] +Verified module: IfThen +All verification conditions discharged successfully diff --git a/tests/error-messages/Inference.fst.json_expected b/tests/error-messages/Inference.fst.json_expected new file mode 100644 index 00000000000..bd6d4c0098f --- /dev/null +++ b/tests/error-messages/Inference.fst.json_expected @@ -0,0 +1,6 @@ +>> Got issues: [ +{"msg":["Subtyping check failed","Expected type Prims.eqtype\ngot type Type0","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"Prims.fst","start_pos":{"line":73,"col":23},"end_pos":{"line":73,"col":30}},"use":{"file_name":"Inference.fst","start_pos":{"line":20,"col":14},"end_pos":{"line":20,"col":15}}},"number":19,"ctx":["While typechecking the top-level declaration `let alloc_fails`","While typechecking the top-level declaration `[@@expect_failure] let alloc_fails`"]} +{"msg":["Subtyping check failed","Expected type Prims.eqtype\ngot type Type0","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"Prims.fst","start_pos":{"line":73,"col":23},"end_pos":{"line":73,"col":30}},"use":{"file_name":"Inference.fst","start_pos":{"line":20,"col":14},"end_pos":{"line":20,"col":15}}},"number":19,"ctx":["While checking for top-level effects","While typechecking the top-level declaration `let alloc_fails`","While typechecking the top-level declaration `[@@expect_failure] let alloc_fails`"]} +>>] +Verified module: Inference +All verification conditions discharged successfully diff --git a/tests/error-messages/Makefile b/tests/error-messages/Makefile index 0ea77d42015..0a916b2a11a 100644 --- a/tests/error-messages/Makefile +++ b/tests/error-messages/Makefile @@ -1,4 +1,5 @@ FSTAR_HOME=../.. +all: check-all FSTAR_FILES=$(wildcard *.fst) @@ -18,43 +19,63 @@ OTHERFLAGS+=--already_cached 'Prims FStar' OTHERFLAGS := $(filter-out --query_stats, $(OTHERFLAGS)) OTHERFLAGS := $(filter-out --hint_info, $(OTHERFLAGS)) -check-all: $(addsuffix .check, $(FSTAR_FILES)) -all: check-all +check-all: $(addsuffix .human_check, $(FSTAR_FILES)) \ + $(addsuffix .json_check, $(FSTAR_FILES)) # For these tests, we check that the resugared output -# matches the expected file. -Bug1997.fst.output: OTHERFLAGS+=--dump_module Bug1997 -Bug2820.fst.output: OTHERFLAGS+=--dump_module Bug2820 -Bug3145.fst.output: OTHERFLAGS+=--dump_module Bug3145 -Bug3227.fst.output: OTHERFLAGS+=--dump_module Bug3227 -Bug3292.fst.output: OTHERFLAGS+=--dump_module Bug3292 -CalcImpl.fst.output: OTHERFLAGS+=--dump_module CalcImpl -DTuples.fst.output: OTHERFLAGS+=--dump_module DTuples -SeqLit.fst.output: OTHERFLAGS+=--dump_module SeqLit +# matches the expected file. We have to repeat the lines for +# json_output, sadly. +Bug1997.fst.json_output : OTHERFLAGS+=--dump_module Bug1997 +Bug1997.fst.output : OTHERFLAGS+=--dump_module Bug1997 +Bug2820.fst.json_output : OTHERFLAGS+=--dump_module Bug2820 +Bug2820.fst.output : OTHERFLAGS+=--dump_module Bug2820 +Bug3145.fst.json_output : OTHERFLAGS+=--dump_module Bug3145 +Bug3145.fst.output : OTHERFLAGS+=--dump_module Bug3145 +Bug3227.fst.json_output : OTHERFLAGS+=--dump_module Bug3227 +Bug3227.fst.output : OTHERFLAGS+=--dump_module Bug3227 +Bug3292.fst.json_output : OTHERFLAGS+=--dump_module Bug3292 +Bug3292.fst.output : OTHERFLAGS+=--dump_module Bug3292 +CalcImpl.fst.json_output : OTHERFLAGS+=--dump_module CalcImpl +CalcImpl.fst.output : OTHERFLAGS+=--dump_module CalcImpl +DTuples.fst.json_output : OTHERFLAGS+=--dump_module DTuples +DTuples.fst.output : OTHERFLAGS+=--dump_module DTuples +SeqLit.fst.json_output : OTHERFLAGS+=--dump_module SeqLit +SeqLit.fst.output : OTHERFLAGS+=--dump_module SeqLit include $(FSTAR_HOME)/examples/Makefile.common -%.check: %.expected %.output +.PHONY: %.human_check +%.human_check: %.expected %.output + $(Q)diff -u --strip-trailing-cr $^ + $(Q)touch $@ + +.PHONY: %.json_check +%.json_check: %.json_expected %.json_output $(Q)diff -u --strip-trailing-cr $^ $(Q)touch $@ -%.accept: %.output - $(Q)cp $< $(patsubst %.output,%.expected,$<) +.PHONY: %.accept +%.accept: %.output %.json_output + $(Q)R=$(patsubst %.output,%,$<) && \ + cp $$R.output $$R.expected && \ + cp $$R.json_output $$R.json_expected +.PHONY: %.clean clean: $(call msg, "CLEAN", $<) $(Q)rm -f .depend - $(Q)rm -f *.check $(Q)rm -f *.output + $(Q)rm -f *.json_output $(Q)rm -rf _output $(Q)rm -rf _cache +.PHONY: accept accept: $(addsuffix .accept, $(FSTAR_FILES)) -.PHONY: %.check - # Re-do all tests +.PHONY: re re: clean all # Keep the output files so we can look at them easily .SECONDARY: $(patsubst %,%.output,$(FSTAR_FILES)) +.SECONDARY: $(patsubst %,%.json_output,$(FSTAR_FILES)) diff --git a/tests/error-messages/NegativeTests.BST.fst.json_expected b/tests/error-messages/NegativeTests.BST.fst.json_expected new file mode 100644 index 00000000000..b74b518fc33 --- /dev/null +++ b/tests/error-messages/NegativeTests.BST.fst.json_expected @@ -0,0 +1,11 @@ +>> Got issues: [ +{"msg":["Subtyping check failed","Expected type\n right:\n FStar.Pervasives.Native.option (tree 2)\n { 0 <= 1 /\\ 1 <= 2 /\\ None? right == (1 = 2) /\\\n None? FStar.Pervasives.Native.None == (1 = 0) }\ngot type FStar.Pervasives.Native.option (tree 2)","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"NegativeTests.BST.fst","start_pos":{"line":27,"col":36},"end_pos":{"line":27,"col":58}},"use":{"file_name":"NegativeTests.BST.fst","start_pos":{"line":37,"col":38},"end_pos":{"line":37,"col":42}}},"number":19,"ctx":["While typechecking the top-level declaration `let bad_node_1`","While typechecking the top-level declaration `[@@expect_failure] let bad_node_1`"]} +>>] +>> Got issues: [ +{"msg":["Subtyping check failed","Expected type\n right:\n FStar.Pervasives.Native.option (tree (l + 1))\n { l <= l /\\ l <= l + 1 /\\ None? right == (l = l + 1) /\\\n None? (FStar.Pervasives.Native.Some t) == (l = l) }\ngot type FStar.Pervasives.Native.option (tree (l + 1))","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"NegativeTests.BST.fst","start_pos":{"line":27,"col":36},"end_pos":{"line":27,"col":58}},"use":{"file_name":"NegativeTests.BST.fst","start_pos":{"line":40,"col":61},"end_pos":{"line":40,"col":65}}},"number":19,"ctx":["While typechecking the top-level declaration `let bad_node_2`","While typechecking the top-level declaration `[@@expect_failure] let bad_node_2`"]} +>>] +>> Got issues: [ +{"msg":["Subtyping check failed","Expected type\n right:\n FStar.Pervasives.Native.option (tree (l + 1))\n { l <= l + 1 /\\ l + 1 <= l + 1 /\\ None? right == (l + 1 = l + 1) /\\\n None? (FStar.Pervasives.Native.Some t1) == (l + 1 = l) }\ngot type FStar.Pervasives.Native.option (tree (l + 1))","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"NegativeTests.BST.fst","start_pos":{"line":27,"col":36},"end_pos":{"line":27,"col":58}},"use":{"file_name":"NegativeTests.BST.fst","start_pos":{"line":43,"col":78},"end_pos":{"line":43,"col":87}}},"number":19,"ctx":["While typechecking the top-level declaration `let bad_node_3`","While typechecking the top-level declaration `[@@expect_failure] let bad_node_3`"]} +>>] +Verified module: NegativeTests.BST +All verification conditions discharged successfully diff --git a/tests/error-messages/NegativeTests.Bug260.fst.json_expected b/tests/error-messages/NegativeTests.Bug260.fst.json_expected new file mode 100644 index 00000000000..8fa52c910cf --- /dev/null +++ b/tests/error-messages/NegativeTests.Bug260.fst.json_expected @@ -0,0 +1,7 @@ +>> Got issues: [ +{"msg":["Subtyping check failed","Expected type validity (S (S t))\ngot type validity (S t)","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"NegativeTests.Bug260.fst","start_pos":{"line":23,"col":37},"end_pos":{"line":26,"col":9}},"use":{"file_name":"NegativeTests.Bug260.fst","start_pos":{"line":26,"col":12},"end_pos":{"line":26,"col":19}}},"number":19,"ctx":["While typechecking the top-level declaration `let bad`","While typechecking the top-level declaration `[@@expect_failure] let bad`"]} +>>] +{"msg":["NegativeTests.Bug260.bad\nis declared but no definition was found","Add an 'assume' if this is intentional"],"level":"Warning","range":{"def":{"file_name":"NegativeTests.Bug260.fst","start_pos":{"line":23,"col":4},"end_pos":{"line":23,"col":7}},"use":{"file_name":"NegativeTests.Bug260.fst","start_pos":{"line":23,"col":4},"end_pos":{"line":23,"col":7}}},"number":240,"ctx":["While desugaring module NegativeTests.Bug260"]} +{"msg":["Missing definitions in module NegativeTests.Bug260: bad"],"level":"Warning","range":{"def":{"file_name":"dummy","start_pos":{"line":0,"col":0},"end_pos":{"line":0,"col":0}},"use":{"file_name":"NegativeTests.Bug260.fst","start_pos":{"line":36,"col":0},"end_pos":{"line":36,"col":34}}},"number":240,"ctx":[]} +Verified module: NegativeTests.Bug260 +All verification conditions discharged successfully diff --git a/tests/error-messages/NegativeTests.False.fst.json_expected b/tests/error-messages/NegativeTests.False.fst.json_expected new file mode 100644 index 00000000000..c5c93b87420 --- /dev/null +++ b/tests/error-messages/NegativeTests.False.fst.json_expected @@ -0,0 +1,12 @@ +>> Got issues: [ +{"msg":["Assertion failed","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"Prims.fst","start_pos":{"line":459,"col":77},"end_pos":{"line":459,"col":89}},"use":{"file_name":"NegativeTests.False.fst","start_pos":{"line":23,"col":13},"end_pos":{"line":23,"col":33}}},"number":19,"ctx":["While typechecking the top-level declaration `let bar`","While typechecking the top-level declaration `[@@expect_failure] let bar`"]} +>>] +>> Got issues: [ +{"msg":["Expected type Prims.l_True \\/ Prims.l_True\nbut Prims.Left Prims.T\nhas type Prims.sum (*?u1*) _ Prims.l_True"],"level":"Error","range":{"def":{"file_name":"NegativeTests.False.fst","start_pos":{"line":30,"col":18},"end_pos":{"line":30,"col":41}},"use":{"file_name":"NegativeTests.False.fst","start_pos":{"line":30,"col":18},"end_pos":{"line":30,"col":41}}},"number":12,"ctx":["While typechecking the top-level declaration `let absurd`","While typechecking the top-level declaration `[@@expect_failure] let absurd`"]} +{"msg":["Expected type Prims.l_True \\/ Prims.l_True\nbut Prims.Right Prims.T\nhas type Prims.sum Prims.l_True (*?u6*) _"],"level":"Error","range":{"def":{"file_name":"NegativeTests.False.fst","start_pos":{"line":30,"col":42},"end_pos":{"line":30,"col":66}},"use":{"file_name":"NegativeTests.False.fst","start_pos":{"line":30,"col":42},"end_pos":{"line":30,"col":66}}},"number":12,"ctx":["While typechecking the top-level declaration `let absurd`","While typechecking the top-level declaration `[@@expect_failure] let absurd`"]} +>>] +{"msg":["NegativeTests.False.bar\nis declared but no definition was found","Add an 'assume' if this is intentional"],"level":"Warning","range":{"def":{"file_name":"NegativeTests.False.fst","start_pos":{"line":21,"col":4},"end_pos":{"line":21,"col":7}},"use":{"file_name":"NegativeTests.False.fst","start_pos":{"line":21,"col":4},"end_pos":{"line":21,"col":7}}},"number":240,"ctx":["While desugaring module NegativeTests.False"]} +{"msg":["NegativeTests.False.absurd\nis declared but no definition was found","Add an 'assume' if this is intentional"],"level":"Warning","range":{"def":{"file_name":"NegativeTests.False.fst","start_pos":{"line":28,"col":4},"end_pos":{"line":28,"col":10}},"use":{"file_name":"NegativeTests.False.fst","start_pos":{"line":28,"col":4},"end_pos":{"line":28,"col":10}}},"number":240,"ctx":["While desugaring module NegativeTests.False"]} +{"msg":["Missing definitions in module NegativeTests.False:\n absurd\n bar"],"level":"Warning","range":{"def":{"file_name":"dummy","start_pos":{"line":0,"col":0},"end_pos":{"line":0,"col":0}},"use":{"file_name":"NegativeTests.False.fst","start_pos":{"line":30,"col":0},"end_pos":{"line":30,"col":66}}},"number":240,"ctx":[]} +Verified module: NegativeTests.False +All verification conditions discharged successfully diff --git a/tests/error-messages/NegativeTests.Heap.fst.json_expected b/tests/error-messages/NegativeTests.Heap.fst.json_expected new file mode 100644 index 00000000000..d2483620e64 --- /dev/null +++ b/tests/error-messages/NegativeTests.Heap.fst.json_expected @@ -0,0 +1,5 @@ +>> Got issues: [ +{"msg":["Assertion failed","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"NegativeTests.Heap.fst","start_pos":{"line":26,"col":21},"end_pos":{"line":26,"col":54}},"use":{"file_name":"NegativeTests.Heap.fst","start_pos":{"line":26,"col":21},"end_pos":{"line":26,"col":54}}},"number":19,"ctx":["While typechecking the top-level declaration `let test2`","While typechecking the top-level declaration `[@@expect_failure] let test2`"]} +>>] +Verified module: NegativeTests.Heap +All verification conditions discharged successfully diff --git a/tests/error-messages/NegativeTests.Neg.fst.json_expected b/tests/error-messages/NegativeTests.Neg.fst.json_expected new file mode 100644 index 00000000000..7525f0ac3d6 --- /dev/null +++ b/tests/error-messages/NegativeTests.Neg.fst.json_expected @@ -0,0 +1,41 @@ +>> Got issues: [ +{"msg":["Subtyping check failed","Expected type Prims.nat\ngot type Prims.int","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"Prims.fst","start_pos":{"line":680,"col":18},"end_pos":{"line":680,"col":24}},"use":{"file_name":"NegativeTests.Neg.fst","start_pos":{"line":20,"col":8},"end_pos":{"line":20,"col":10}}},"number":19,"ctx":["While checking for top-level effects","While typechecking the top-level declaration `let x`","While typechecking the top-level declaration `[@@expect_failure] let x`"]} +>>] +>> Got issues: [ +{"msg":["Subtyping check failed","Expected type Prims.nat\ngot type Prims.int","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"Prims.fst","start_pos":{"line":680,"col":18},"end_pos":{"line":680,"col":24}},"use":{"file_name":"NegativeTests.Neg.fst","start_pos":{"line":24,"col":8},"end_pos":{"line":24,"col":10}}},"number":19,"ctx":["While checking for top-level effects","While typechecking the top-level declaration `let y`","While typechecking the top-level declaration `[@@expect_failure] let y`"]} +>>] +>> Got issues: [ +{"msg":["Assertion failed","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"NegativeTests.Neg.fst","start_pos":{"line":27,"col":30},"end_pos":{"line":27,"col":35}},"use":{"file_name":"NegativeTests.Neg.fst","start_pos":{"line":27,"col":30},"end_pos":{"line":27,"col":35}}},"number":19,"ctx":["While typechecking the top-level declaration `let assert_0_eq_1`","While typechecking the top-level declaration `[@@expect_failure] let assert_0_eq_1`"]} +>>] +>> Got issues: [ +{"msg":["Patterns are incomplete","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"dummy","start_pos":{"line":0,"col":0},"end_pos":{"line":0,"col":0}},"use":{"file_name":"NegativeTests.Neg.fst","start_pos":{"line":30,"col":28},"end_pos":{"line":31,"col":15}}},"number":19,"ctx":["While typechecking the top-level declaration `let hd_int_inexhaustive`","While typechecking the top-level declaration `[@@expect_failure] let hd_int_inexhaustive`"]} +>>] +>> Got issues: [ +{"msg":["Assertion failed","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"NegativeTests.Neg.fst","start_pos":{"line":33,"col":44},"end_pos":{"line":33,"col":57}},"use":{"file_name":"NegativeTests.Neg.fst","start_pos":{"line":38,"col":32},"end_pos":{"line":38,"col":42}}},"number":19,"ctx":["While typechecking the top-level declaration `let test_precondition_label`","While typechecking the top-level declaration `[@@expect_failure] let test_precondition_label`"]} +>>] +>> Got issues: [ +{"msg":["Could not prove post-condition","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"NegativeTests.Neg.fst","start_pos":{"line":40,"col":83},"end_pos":{"line":40,"col":88}},"use":{"file_name":"NegativeTests.Neg.fst","start_pos":{"line":42,"col":33},"end_pos":{"line":42,"col":34}}},"number":19,"ctx":["While typechecking the top-level declaration `let test_postcondition_label`","While typechecking the top-level declaration `[@@expect_failure] let test_postcondition_label`"]} +>>] +>> Got issues: [ +{"msg":["Subtyping check failed","Expected type _: FStar.Pervasives.Native.option 'a {Some? _}\ngot type FStar.Pervasives.Native.option 'a","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"FStar.Pervasives.Native.fst","start_pos":{"line":33,"col":4},"end_pos":{"line":33,"col":8}},"use":{"file_name":"NegativeTests.Neg.fst","start_pos":{"line":46,"col":30},"end_pos":{"line":46,"col":31}}},"number":19,"ctx":["While typechecking the top-level declaration `let bad_projector`","While typechecking the top-level declaration `[@@expect_failure] let bad_projector`"]} +>>] +>> Got issues: [ +{"msg":["Subtyping check failed","Expected type _: FStar.Pervasives.result Prims.int {V? _}\ngot type FStar.Pervasives.result Prims.int","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"FStar.Pervasives.fsti","start_pos":{"line":519,"col":4},"end_pos":{"line":519,"col":5}},"use":{"file_name":"NegativeTests.Neg.fst","start_pos":{"line":50,"col":45},"end_pos":{"line":50,"col":47}}},"number":19,"ctx":["While typechecking the top-level declaration `val NegativeTests.Neg.test`","While typechecking the top-level declaration `[@@expect_failure] val NegativeTests.Neg.test`"]} +>>] +>> Got issues: [ +{"msg":["Subtyping check failed","Expected type Prims.nat\ngot type Prims.int","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"Prims.fst","start_pos":{"line":680,"col":18},"end_pos":{"line":680,"col":24}},"use":{"file_name":"NegativeTests.Neg.fst","start_pos":{"line":55,"col":25},"end_pos":{"line":55,"col":26}}},"number":19,"ctx":["While checking for top-level effects","While typechecking the top-level declaration `let h1`","While typechecking the top-level declaration `[@@expect_failure] let h1`"]} +>>] +>> Got issues: [ +{"msg":["Type annotation _: Type0{NegativeTests.Neg.phi_1510} for inductive\nNegativeTests.Neg.t is not Type or eqtype, or it is eqtype but contains\nnoeq/unopteq qualifiers"],"level":"Error","range":{"def":{"file_name":"NegativeTests.Neg.fst","start_pos":{"line":59,"col":5},"end_pos":{"line":59,"col":6}},"use":{"file_name":"NegativeTests.Neg.fst","start_pos":{"line":59,"col":5},"end_pos":{"line":59,"col":6}}},"number":309,"ctx":["While typechecking the top-level declaration `type NegativeTests.Neg.t`","While typechecking the top-level declaration `[@@expect_failure] type NegativeTests.Neg.t`"]} +>>] +>> Got issues: [ +{"msg":["Type annotation _: Type{Prims.l_False} for inductive NegativeTests.Neg.t2 is not\nType or eqtype, or it is eqtype but contains noeq/unopteq qualifiers"],"level":"Error","range":{"def":{"file_name":"NegativeTests.Neg.fst","start_pos":{"line":63,"col":5},"end_pos":{"line":63,"col":7}},"use":{"file_name":"NegativeTests.Neg.fst","start_pos":{"line":63,"col":5},"end_pos":{"line":63,"col":7}}},"number":309,"ctx":["While typechecking the top-level declaration `type NegativeTests.Neg.t2`","While typechecking the top-level declaration `[@@expect_failure] type NegativeTests.Neg.t2`"]} +>>] +{"msg":["NegativeTests.Neg.x\nis declared but no definition was found","Add an 'assume' if this is intentional"],"level":"Warning","range":{"def":{"file_name":"NegativeTests.Neg.fst","start_pos":{"line":18,"col":4},"end_pos":{"line":18,"col":5}},"use":{"file_name":"NegativeTests.Neg.fst","start_pos":{"line":18,"col":4},"end_pos":{"line":18,"col":5}}},"number":240,"ctx":["While desugaring module NegativeTests.Neg"]} +{"msg":["NegativeTests.Neg.y\nis declared but no definition was found","Add an 'assume' if this is intentional"],"level":"Warning","range":{"def":{"file_name":"NegativeTests.Neg.fst","start_pos":{"line":22,"col":4},"end_pos":{"line":22,"col":5}},"use":{"file_name":"NegativeTests.Neg.fst","start_pos":{"line":22,"col":4},"end_pos":{"line":22,"col":5}}},"number":240,"ctx":["While desugaring module NegativeTests.Neg"]} +{"msg":["NegativeTests.Neg.test_precondition_label\nis declared but no definition was found","Add an 'assume' if this is intentional"],"level":"Warning","range":{"def":{"file_name":"NegativeTests.Neg.fst","start_pos":{"line":36,"col":4},"end_pos":{"line":36,"col":27}},"use":{"file_name":"NegativeTests.Neg.fst","start_pos":{"line":36,"col":4},"end_pos":{"line":36,"col":27}}},"number":240,"ctx":["While desugaring module NegativeTests.Neg"]} +{"msg":["NegativeTests.Neg.test_postcondition_label\nis declared but no definition was found","Add an 'assume' if this is intentional"],"level":"Warning","range":{"def":{"file_name":"NegativeTests.Neg.fst","start_pos":{"line":40,"col":4},"end_pos":{"line":40,"col":28}},"use":{"file_name":"NegativeTests.Neg.fst","start_pos":{"line":40,"col":4},"end_pos":{"line":40,"col":28}}},"number":240,"ctx":["While desugaring module NegativeTests.Neg"]} +{"msg":["NegativeTests.Neg.bad_projector\nis declared but no definition was found","Add an 'assume' if this is intentional"],"level":"Warning","range":{"def":{"file_name":"NegativeTests.Neg.fst","start_pos":{"line":44,"col":4},"end_pos":{"line":44,"col":17}},"use":{"file_name":"NegativeTests.Neg.fst","start_pos":{"line":44,"col":4},"end_pos":{"line":44,"col":17}}},"number":240,"ctx":["While desugaring module NegativeTests.Neg"]} +{"msg":["Missing definitions in module NegativeTests.Neg:\n bad_projector\n test_postcondition_label\n test_precondition_label\n x\n y"],"level":"Warning","range":{"def":{"file_name":"dummy","start_pos":{"line":0,"col":0},"end_pos":{"line":0,"col":0}},"use":{"file_name":"NegativeTests.Neg.fst","start_pos":{"line":63,"col":0},"end_pos":{"line":63,"col":32}}},"number":240,"ctx":[]} +Verified module: NegativeTests.Neg +All verification conditions discharged successfully diff --git a/tests/error-messages/NegativeTests.Positivity.fst.json_expected b/tests/error-messages/NegativeTests.Positivity.fst.json_expected new file mode 100644 index 00000000000..1672d7ba66c --- /dev/null +++ b/tests/error-messages/NegativeTests.Positivity.fst.json_expected @@ -0,0 +1,20 @@ +>> Got issues: [ +{"msg":["Inductive type NegativeTests.Positivity.t1 does not satisfy the strict positivity condition"],"level":"Error","range":{"def":{"file_name":"NegativeTests.Positivity.fst","start_pos":{"line":22,"col":10},"end_pos":{"line":22,"col":12}},"use":{"file_name":"NegativeTests.Positivity.fst","start_pos":{"line":22,"col":10},"end_pos":{"line":22,"col":12}}},"number":3,"ctx":["While typechecking the top-level declaration `type NegativeTests.Positivity.t1`","While typechecking the top-level declaration `[@@expect_failure] type NegativeTests.Positivity.t1`"]} +>>] +>> Got issues: [ +{"msg":["Inductive type NegativeTests.Positivity.t5 does not satisfy the strict positivity condition"],"level":"Error","range":{"def":{"file_name":"NegativeTests.Positivity.fst","start_pos":{"line":37,"col":10},"end_pos":{"line":37,"col":12}},"use":{"file_name":"NegativeTests.Positivity.fst","start_pos":{"line":37,"col":10},"end_pos":{"line":37,"col":12}}},"number":3,"ctx":["While typechecking the top-level declaration `type NegativeTests.Positivity.t5`","While typechecking the top-level declaration `[@@expect_failure] type NegativeTests.Positivity.t5`"]} +>>] +>> Got issues: [ +{"msg":["Type NegativeTests.Positivity.t7 is not strictly positive since it instantiates a non-uniformly recursive parameter or index NegativeTests.Positivity.t7 of NegativeTests.Positivity.t6"],"level":"Error","range":{"def":{"file_name":"NegativeTests.Positivity.fst","start_pos":{"line":49,"col":12},"end_pos":{"line":49,"col":14}},"use":{"file_name":"NegativeTests.Positivity.fst","start_pos":{"line":49,"col":4},"end_pos":{"line":49,"col":7}}},"number":3,"ctx":["While typechecking the top-level declaration `type NegativeTests.Positivity.t7`","While typechecking the top-level declaration `[@@expect_failure] type NegativeTests.Positivity.t7`"]} +>>] +>> Got issues: [ +{"msg":["Type NegativeTests.Positivity.t8 is not strictly positive since it instantiates a non-uniformly recursive parameter or index NegativeTests.Positivity.t8 of NegativeTests.Positivity.t6"],"level":"Error","range":{"def":{"file_name":"NegativeTests.Positivity.fst","start_pos":{"line":54,"col":16},"end_pos":{"line":54,"col":18}},"use":{"file_name":"NegativeTests.Positivity.fst","start_pos":{"line":54,"col":4},"end_pos":{"line":54,"col":7}}},"number":3,"ctx":["While typechecking the top-level declaration `type NegativeTests.Positivity.t8`","While typechecking the top-level declaration `[@@expect_failure] type NegativeTests.Positivity.t8`"]} +>>] +>> Got issues: [ +{"msg":["Inductive type NegativeTests.Positivity.t10 does not satisfy the strict positivity condition"],"level":"Error","range":{"def":{"file_name":"NegativeTests.Positivity.fst","start_pos":{"line":61,"col":10},"end_pos":{"line":61,"col":13}},"use":{"file_name":"NegativeTests.Positivity.fst","start_pos":{"line":61,"col":10},"end_pos":{"line":61,"col":13}}},"number":3,"ctx":["While typechecking the top-level declaration `type NegativeTests.Positivity.t10`","While typechecking the top-level declaration `[@@expect_failure] type NegativeTests.Positivity.t10`"]} +>>] +>> Got issues: [ +{"msg":["Inductive type NegativeTests.Positivity.t11 does not satisfy the strict positivity condition"],"level":"Error","range":{"def":{"file_name":"NegativeTests.Positivity.fst","start_pos":{"line":66,"col":10},"end_pos":{"line":66,"col":13}},"use":{"file_name":"NegativeTests.Positivity.fst","start_pos":{"line":66,"col":10},"end_pos":{"line":66,"col":13}}},"number":3,"ctx":["While typechecking the top-level declaration `type NegativeTests.Positivity.t11`","While typechecking the top-level declaration `[@@expect_failure] type NegativeTests.Positivity.t11`"]} +>>] +Verified module: NegativeTests.Positivity +All verification conditions discharged successfully diff --git a/tests/error-messages/NegativeTests.Set.fst.json_expected b/tests/error-messages/NegativeTests.Set.fst.json_expected new file mode 100644 index 00000000000..59580de7b6e --- /dev/null +++ b/tests/error-messages/NegativeTests.Set.fst.json_expected @@ -0,0 +1,15 @@ +>> Got issues: [ +{"msg":["Assertion failed","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"NegativeTests.Set.fst","start_pos":{"line":28,"col":9},"end_pos":{"line":28,"col":30}},"use":{"file_name":"NegativeTests.Set.fst","start_pos":{"line":28,"col":9},"end_pos":{"line":28,"col":30}}},"number":19,"ctx":["While typechecking the top-level declaration `let should_fail1`","While typechecking the top-level declaration `[@@expect_failure] let should_fail1`"]} +>>] +>> Got issues: [ +{"msg":["Assertion failed","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"NegativeTests.Set.fst","start_pos":{"line":33,"col":9},"end_pos":{"line":33,"col":67}},"use":{"file_name":"NegativeTests.Set.fst","start_pos":{"line":33,"col":9},"end_pos":{"line":33,"col":67}}},"number":19,"ctx":["While typechecking the top-level declaration `let should_fail2`","While typechecking the top-level declaration `[@@expect_failure] let should_fail2`"]} +>>] +>> Got issues: [ +{"msg":["Assertion failed","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"NegativeTests.Set.fst","start_pos":{"line":38,"col":9},"end_pos":{"line":38,"col":52}},"use":{"file_name":"NegativeTests.Set.fst","start_pos":{"line":38,"col":9},"end_pos":{"line":38,"col":52}}},"number":19,"ctx":["While typechecking the top-level declaration `let should_fail3`","While typechecking the top-level declaration `[@@expect_failure] let should_fail3`"]} +>>] +{"msg":["NegativeTests.Set.should_fail1\nis declared but no definition was found","Add an 'assume' if this is intentional"],"level":"Warning","range":{"def":{"file_name":"NegativeTests.Set.fst","start_pos":{"line":25,"col":4},"end_pos":{"line":25,"col":16}},"use":{"file_name":"NegativeTests.Set.fst","start_pos":{"line":25,"col":4},"end_pos":{"line":25,"col":16}}},"number":240,"ctx":["While desugaring module NegativeTests.Set"]} +{"msg":["NegativeTests.Set.should_fail2\nis declared but no definition was found","Add an 'assume' if this is intentional"],"level":"Warning","range":{"def":{"file_name":"NegativeTests.Set.fst","start_pos":{"line":30,"col":4},"end_pos":{"line":30,"col":16}},"use":{"file_name":"NegativeTests.Set.fst","start_pos":{"line":30,"col":4},"end_pos":{"line":30,"col":16}}},"number":240,"ctx":["While desugaring module NegativeTests.Set"]} +{"msg":["NegativeTests.Set.should_fail3\nis declared but no definition was found","Add an 'assume' if this is intentional"],"level":"Warning","range":{"def":{"file_name":"NegativeTests.Set.fst","start_pos":{"line":35,"col":4},"end_pos":{"line":35,"col":16}},"use":{"file_name":"NegativeTests.Set.fst","start_pos":{"line":35,"col":4},"end_pos":{"line":35,"col":16}}},"number":240,"ctx":["While desugaring module NegativeTests.Set"]} +{"msg":["Missing definitions in module NegativeTests.Set:\n should_fail1\n should_fail2\n should_fail3"],"level":"Warning","range":{"def":{"file_name":"dummy","start_pos":{"line":0,"col":0},"end_pos":{"line":0,"col":0}},"use":{"file_name":"NegativeTests.Set.fst","start_pos":{"line":37,"col":0},"end_pos":{"line":38,"col":52}}},"number":240,"ctx":[]} +Verified module: NegativeTests.Set +All verification conditions discharged successfully diff --git a/tests/error-messages/NegativeTests.ShortCircuiting.fst.json_expected b/tests/error-messages/NegativeTests.ShortCircuiting.fst.json_expected new file mode 100644 index 00000000000..276cbe6aed4 --- /dev/null +++ b/tests/error-messages/NegativeTests.ShortCircuiting.fst.json_expected @@ -0,0 +1,11 @@ +>> Got issues: [ +{"msg":["Subtyping check failed","Expected type b: Prims.bool{bad_p b}\ngot type Prims.bool","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"NegativeTests.ShortCircuiting.fst","start_pos":{"line":19,"col":31},"end_pos":{"line":19,"col":38}},"use":{"file_name":"NegativeTests.ShortCircuiting.fst","start_pos":{"line":21,"col":16},"end_pos":{"line":21,"col":33}}},"number":19,"ctx":["While typechecking the top-level declaration `let rec bad`","While typechecking the top-level declaration `[@@expect_failure] let rec bad`"]} +>>] +>> Got issues: [ +{"msg":["Assertion failed","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"Prims.fst","start_pos":{"line":451,"col":90},"end_pos":{"line":451,"col":102}},"use":{"file_name":"NegativeTests.ShortCircuiting.fst","start_pos":{"line":25,"col":11},"end_pos":{"line":25,"col":36}}},"number":19,"ctx":["While typechecking the top-level declaration `let ff`","While typechecking the top-level declaration `[@@expect_failure] let ff`"]} +>>] +{"msg":["NegativeTests.ShortCircuiting.bad\nis declared but no definition was found","Add an 'assume' if this is intentional"],"level":"Warning","range":{"def":{"file_name":"NegativeTests.ShortCircuiting.fst","start_pos":{"line":19,"col":4},"end_pos":{"line":19,"col":7}},"use":{"file_name":"NegativeTests.ShortCircuiting.fst","start_pos":{"line":19,"col":4},"end_pos":{"line":19,"col":7}}},"number":240,"ctx":["While desugaring module NegativeTests.ShortCircuiting"]} +{"msg":["NegativeTests.ShortCircuiting.ff\nis declared but no definition was found","Add an 'assume' if this is intentional"],"level":"Warning","range":{"def":{"file_name":"NegativeTests.ShortCircuiting.fst","start_pos":{"line":23,"col":4},"end_pos":{"line":23,"col":6}},"use":{"file_name":"NegativeTests.ShortCircuiting.fst","start_pos":{"line":23,"col":4},"end_pos":{"line":23,"col":6}}},"number":240,"ctx":["While desugaring module NegativeTests.ShortCircuiting"]} +{"msg":["Missing definitions in module NegativeTests.ShortCircuiting:\n bad\n ff"],"level":"Warning","range":{"def":{"file_name":"dummy","start_pos":{"line":0,"col":0},"end_pos":{"line":0,"col":0}},"use":{"file_name":"NegativeTests.ShortCircuiting.fst","start_pos":{"line":25,"col":0},"end_pos":{"line":25,"col":36}}},"number":240,"ctx":[]} +Verified module: NegativeTests.ShortCircuiting +All verification conditions discharged successfully diff --git a/tests/error-messages/NegativeTests.Termination.fst.json_expected b/tests/error-messages/NegativeTests.Termination.fst.json_expected new file mode 100644 index 00000000000..1f3fce16960 --- /dev/null +++ b/tests/error-messages/NegativeTests.Termination.fst.json_expected @@ -0,0 +1,43 @@ +>> Got issues: [ +{"msg":["Could not prove termination of this recursive call","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"NegativeTests.Termination.fst","start_pos":{"line":18,"col":12},"end_pos":{"line":21,"col":29}},"use":{"file_name":"NegativeTests.Termination.fst","start_pos":{"line":21,"col":28},"end_pos":{"line":21,"col":29}}},"number":19,"ctx":["While typechecking the top-level declaration `let rec bug15`","While typechecking the top-level declaration `[@@expect_failure] let rec bug15`"]} +>>] +>> Got issues: [ +{"msg":["Could not prove termination of this recursive call","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"NegativeTests.Termination.fst","start_pos":{"line":25,"col":23},"end_pos":{"line":28,"col":35}},"use":{"file_name":"NegativeTests.Termination.fst","start_pos":{"line":28,"col":26},"end_pos":{"line":28,"col":35}}},"number":19,"ctx":["While typechecking the top-level declaration `let rec repeat_diverge`","While typechecking the top-level declaration `[@@expect_failure] let rec repeat_diverge`"]} +>>] +>> Got issues: [ +{"msg":["Could not prove termination of this recursive call","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"NegativeTests.Termination.fst","start_pos":{"line":31,"col":19},"end_pos":{"line":36,"col":54}},"use":{"file_name":"NegativeTests.Termination.fst","start_pos":{"line":35,"col":43},"end_pos":{"line":35,"col":44}}},"number":19,"ctx":["While typechecking the top-level declaration `let rec ackermann_bad`","While typechecking the top-level declaration `[@@expect_failure] let rec ackermann_bad`"]} +>>] +>> Got issues: [ +{"msg":["Could not prove termination of this recursive call","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"NegativeTests.Termination.fst","start_pos":{"line":40,"col":23},"end_pos":{"line":42,"col":29}},"use":{"file_name":"NegativeTests.Termination.fst","start_pos":{"line":42,"col":28},"end_pos":{"line":42,"col":29}}},"number":19,"ctx":["While typechecking the top-level declaration `let rec length_bad`","While typechecking the top-level declaration `[@@expect_failure] let rec length_bad`"]} +>>] +>> Got issues: [ +{"msg":["Could not prove termination of this recursive call","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"NegativeTests.Termination.fst","start_pos":{"line":53,"col":2},"end_pos":{"line":55,"col":29}},"use":{"file_name":"NegativeTests.Termination.fst","start_pos":{"line":55,"col":15},"end_pos":{"line":55,"col":29}}},"number":19,"ctx":["While typechecking the top-level declaration `let rec strangeZeroBad`","While typechecking the top-level declaration `[@@expect_failure] let rec strangeZeroBad`"]} +>>] +>> Got issues: [ +{"msg":["Could not prove termination of this recursive call","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"NegativeTests.Termination.fst","start_pos":{"line":64,"col":2},"end_pos":{"line":67,"col":29}},"use":{"file_name":"NegativeTests.Termination.fst","start_pos":{"line":67,"col":19},"end_pos":{"line":67,"col":29}}},"number":19,"ctx":["While typechecking the top-level declaration `let rec t1`","While typechecking the top-level declaration `[@@expect_failure] let rec t1`"]} +>>] +>> Got issues: [ +{"msg":["Could not prove termination of this recursive call","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"NegativeTests.Termination.fst","start_pos":{"line":71,"col":13},"end_pos":{"line":75,"col":35}},"use":{"file_name":"NegativeTests.Termination.fst","start_pos":{"line":75,"col":34},"end_pos":{"line":75,"col":35}}},"number":19,"ctx":["While typechecking the top-level declaration `let rec plus`","While typechecking the top-level declaration `[@@expect_failure] let rec plus`"]} +>>] +>> Got issues: [ +{"msg":["Patterns are incomplete","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"dummy","start_pos":{"line":0,"col":0},"end_pos":{"line":0,"col":0}},"use":{"file_name":"NegativeTests.Termination.fst","start_pos":{"line":80,"col":2},"end_pos":{"line":82,"col":14}}},"number":19,"ctx":["While typechecking the top-level declaration `let plus'`","While typechecking the top-level declaration `[@@expect_failure] let plus'`"]} +>>] +>> Got issues: [ +{"msg":["Could not prove termination of this recursive call","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"NegativeTests.Termination.fst","start_pos":{"line":86,"col":15},"end_pos":{"line":89,"col":31}},"use":{"file_name":"NegativeTests.Termination.fst","start_pos":{"line":89,"col":29},"end_pos":{"line":89,"col":31}}},"number":19,"ctx":["While typechecking the top-level declaration `let rec minus`","While typechecking the top-level declaration `[@@expect_failure] let rec minus`"]} +>>] +>> Got issues: [ +{"msg":["Could not prove termination of this recursive call","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"NegativeTests.Termination.fst","start_pos":{"line":94,"col":2},"end_pos":{"line":96,"col":26}},"use":{"file_name":"NegativeTests.Termination.fst","start_pos":{"line":96,"col":20},"end_pos":{"line":96,"col":26}}},"number":19,"ctx":["While typechecking the top-level declaration `let rec xxx`","While typechecking the top-level declaration `[@@expect_failure] let rec xxx`"]} +>>] +{"msg":["NegativeTests.Termination.bug15\nis declared but no definition was found","Add an 'assume' if this is intentional"],"level":"Warning","range":{"def":{"file_name":"NegativeTests.Termination.fst","start_pos":{"line":18,"col":4},"end_pos":{"line":18,"col":9}},"use":{"file_name":"NegativeTests.Termination.fst","start_pos":{"line":18,"col":4},"end_pos":{"line":18,"col":9}}},"number":240,"ctx":["While desugaring module NegativeTests.Termination"]} +{"msg":["NegativeTests.Termination.repeat_diverge\nis declared but no definition was found","Add an 'assume' if this is intentional"],"level":"Warning","range":{"def":{"file_name":"NegativeTests.Termination.fst","start_pos":{"line":23,"col":4},"end_pos":{"line":23,"col":18}},"use":{"file_name":"NegativeTests.Termination.fst","start_pos":{"line":23,"col":4},"end_pos":{"line":23,"col":18}}},"number":240,"ctx":["While desugaring module NegativeTests.Termination"]} +{"msg":["NegativeTests.Termination.ackermann_bad\nis declared but no definition was found","Add an 'assume' if this is intentional"],"level":"Warning","range":{"def":{"file_name":"NegativeTests.Termination.fst","start_pos":{"line":31,"col":4},"end_pos":{"line":31,"col":17}},"use":{"file_name":"NegativeTests.Termination.fst","start_pos":{"line":31,"col":4},"end_pos":{"line":31,"col":17}}},"number":240,"ctx":["While desugaring module NegativeTests.Termination"]} +{"msg":["NegativeTests.Termination.length_bad\nis declared but no definition was found","Add an 'assume' if this is intentional"],"level":"Warning","range":{"def":{"file_name":"NegativeTests.Termination.fst","start_pos":{"line":38,"col":4},"end_pos":{"line":38,"col":14}},"use":{"file_name":"NegativeTests.Termination.fst","start_pos":{"line":38,"col":4},"end_pos":{"line":38,"col":14}}},"number":240,"ctx":["While desugaring module NegativeTests.Termination"]} +{"msg":["NegativeTests.Termination.strangeZeroBad\nis declared but no definition was found","Add an 'assume' if this is intentional"],"level":"Warning","range":{"def":{"file_name":"NegativeTests.Termination.fst","start_pos":{"line":50,"col":4},"end_pos":{"line":50,"col":18}},"use":{"file_name":"NegativeTests.Termination.fst","start_pos":{"line":50,"col":4},"end_pos":{"line":50,"col":18}}},"number":240,"ctx":["While desugaring module NegativeTests.Termination"]} +{"msg":["NegativeTests.Termination.t1\nis declared but no definition was found","Add an 'assume' if this is intentional"],"level":"Warning","range":{"def":{"file_name":"NegativeTests.Termination.fst","start_pos":{"line":61,"col":4},"end_pos":{"line":61,"col":6}},"use":{"file_name":"NegativeTests.Termination.fst","start_pos":{"line":61,"col":4},"end_pos":{"line":61,"col":6}}},"number":240,"ctx":["While desugaring module NegativeTests.Termination"]} +{"msg":["NegativeTests.Termination.plus\nis declared but no definition was found","Add an 'assume' if this is intentional"],"level":"Warning","range":{"def":{"file_name":"NegativeTests.Termination.fst","start_pos":{"line":69,"col":4},"end_pos":{"line":69,"col":8}},"use":{"file_name":"NegativeTests.Termination.fst","start_pos":{"line":69,"col":4},"end_pos":{"line":69,"col":8}}},"number":240,"ctx":["While desugaring module NegativeTests.Termination"]} +{"msg":["NegativeTests.Termination.plus'\nis declared but no definition was found","Add an 'assume' if this is intentional"],"level":"Warning","range":{"def":{"file_name":"NegativeTests.Termination.fst","start_pos":{"line":77,"col":4},"end_pos":{"line":77,"col":9}},"use":{"file_name":"NegativeTests.Termination.fst","start_pos":{"line":77,"col":4},"end_pos":{"line":77,"col":9}}},"number":240,"ctx":["While desugaring module NegativeTests.Termination"]} +{"msg":["NegativeTests.Termination.minus\nis declared but no definition was found","Add an 'assume' if this is intentional"],"level":"Warning","range":{"def":{"file_name":"NegativeTests.Termination.fst","start_pos":{"line":84,"col":4},"end_pos":{"line":84,"col":9}},"use":{"file_name":"NegativeTests.Termination.fst","start_pos":{"line":84,"col":4},"end_pos":{"line":84,"col":9}}},"number":240,"ctx":["While desugaring module NegativeTests.Termination"]} +{"msg":["NegativeTests.Termination.xxx\nis declared but no definition was found","Add an 'assume' if this is intentional"],"level":"Warning","range":{"def":{"file_name":"NegativeTests.Termination.fst","start_pos":{"line":91,"col":4},"end_pos":{"line":91,"col":7}},"use":{"file_name":"NegativeTests.Termination.fst","start_pos":{"line":91,"col":4},"end_pos":{"line":91,"col":7}}},"number":240,"ctx":["While desugaring module NegativeTests.Termination"]} +{"msg":["Missing definitions in module NegativeTests.Termination:\n ackermann_bad\n bug15\n length_bad\n minus\n plus\n plus'\n repeat_diverge\n strangeZeroBad\n t1\n xxx"],"level":"Warning","range":{"def":{"file_name":"dummy","start_pos":{"line":0,"col":0},"end_pos":{"line":0,"col":0}},"use":{"file_name":"NegativeTests.Termination.fst","start_pos":{"line":93,"col":0},"end_pos":{"line":96,"col":26}}},"number":240,"ctx":[]} +Verified module: NegativeTests.Termination +All verification conditions discharged successfully diff --git a/tests/error-messages/NegativeTests.ZZImplicitFalse.fst.json_expected b/tests/error-messages/NegativeTests.ZZImplicitFalse.fst.json_expected new file mode 100644 index 00000000000..c2287841254 --- /dev/null +++ b/tests/error-messages/NegativeTests.ZZImplicitFalse.fst.json_expected @@ -0,0 +1,7 @@ +>> Got issues: [ +{"msg":["Subtyping check failed","Expected type Prims.l_False\ngot type Prims.unit","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"Prims.fst","start_pos":{"line":159,"col":29},"end_pos":{"line":159,"col":34}},"use":{"file_name":"NegativeTests.ZZImplicitFalse.fst","start_pos":{"line":20,"col":27},"end_pos":{"line":20,"col":28}}},"number":19,"ctx":["While typechecking the top-level declaration `let test`","While typechecking the top-level declaration `[@@expect_failure] let test`"]} +>>] +{"msg":["NegativeTests.ZZImplicitFalse.test\nis declared but no definition was found","Add an 'assume' if this is intentional"],"level":"Warning","range":{"def":{"file_name":"NegativeTests.ZZImplicitFalse.fst","start_pos":{"line":18,"col":4},"end_pos":{"line":18,"col":8}},"use":{"file_name":"NegativeTests.ZZImplicitFalse.fst","start_pos":{"line":18,"col":4},"end_pos":{"line":18,"col":8}}},"number":240,"ctx":["While desugaring module NegativeTests.ZZImplicitFalse"]} +{"msg":["Missing definitions in module NegativeTests.ZZImplicitFalse: test"],"level":"Warning","range":{"def":{"file_name":"dummy","start_pos":{"line":0,"col":0},"end_pos":{"line":0,"col":0}},"use":{"file_name":"NegativeTests.ZZImplicitFalse.fst","start_pos":{"line":20,"col":0},"end_pos":{"line":20,"col":34}}},"number":240,"ctx":[]} +Verified module: NegativeTests.ZZImplicitFalse +All verification conditions discharged successfully diff --git a/tests/error-messages/OccursCheckOnArrows.fst.json_expected b/tests/error-messages/OccursCheckOnArrows.fst.json_expected new file mode 100644 index 00000000000..9b99c70d220 --- /dev/null +++ b/tests/error-messages/OccursCheckOnArrows.fst.json_expected @@ -0,0 +1,5 @@ +>> Got issues: [ +{"msg":["_: (*?u5*) _ -> (*?u4*) _ is not a subtype of the expected type (*?u5*) _"],"level":"Error","range":{"def":{"file_name":"OccursCheckOnArrows.fst","start_pos":{"line":18,"col":9},"end_pos":{"line":18,"col":10}},"use":{"file_name":"OccursCheckOnArrows.fst","start_pos":{"line":18,"col":15},"end_pos":{"line":18,"col":16}}},"number":54,"ctx":["While solving deferred constraints","solve_non_tactic_deferred_constraints","While typechecking the top-level declaration `let oops`","While typechecking the top-level declaration `[@@expect_failure] let oops`"]} +>>] +Verified module: OccursCheckOnArrows +All verification conditions discharged successfully diff --git a/tests/error-messages/OptionStack.fst.json_expected b/tests/error-messages/OptionStack.fst.json_expected new file mode 100644 index 00000000000..268463b8b3a --- /dev/null +++ b/tests/error-messages/OptionStack.fst.json_expected @@ -0,0 +1,14 @@ +>> Got issues: [ +{"msg":["Assertion failed","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"OptionStack.fst","start_pos":{"line":21,"col":16},"end_pos":{"line":21,"col":21}},"use":{"file_name":"OptionStack.fst","start_pos":{"line":21,"col":9},"end_pos":{"line":21,"col":15}}},"number":19,"ctx":["While checking for top-level effects","While typechecking the top-level declaration `let t0`","While typechecking the top-level declaration `[@@expect_failure] let t0`"]} +>>] +>> Got issues: [ +{"msg":["Assertion failed","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"OptionStack.fst","start_pos":{"line":24,"col":16},"end_pos":{"line":24,"col":21}},"use":{"file_name":"OptionStack.fst","start_pos":{"line":24,"col":9},"end_pos":{"line":24,"col":15}}},"number":19,"ctx":["While checking for top-level effects","While typechecking the top-level declaration `let t1`","While typechecking the top-level declaration `[@@expect_failure] let t1`"]} +>>] +>> Got issues: [ +{"msg":["Assertion failed","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"OptionStack.fst","start_pos":{"line":27,"col":16},"end_pos":{"line":27,"col":21}},"use":{"file_name":"OptionStack.fst","start_pos":{"line":27,"col":9},"end_pos":{"line":27,"col":15}}},"number":19,"ctx":["While checking for top-level effects","While typechecking the top-level declaration `let t2`","While typechecking the top-level declaration `[@@expect_failure] let t2`"]} +>>] +>> Got issues: [ +{"msg":["Assertion failed","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"OptionStack.fst","start_pos":{"line":39,"col":16},"end_pos":{"line":39,"col":21}},"use":{"file_name":"OptionStack.fst","start_pos":{"line":39,"col":9},"end_pos":{"line":39,"col":15}}},"number":19,"ctx":["While checking for top-level effects","While typechecking the top-level declaration `let t7`","While typechecking the top-level declaration `[@@expect_failure] let t7`"]} +>>] +Verified module: OptionStack +All verification conditions discharged successfully diff --git a/tests/error-messages/PatAnnot.fst.json_expected b/tests/error-messages/PatAnnot.fst.json_expected new file mode 100644 index 00000000000..81144835ab0 --- /dev/null +++ b/tests/error-messages/PatAnnot.fst.json_expected @@ -0,0 +1,20 @@ +>> Got issues: [ +{"msg":["Subtyping check failed","Expected type Prims.squash Prims.l_False\ngot type Prims.unit","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"PatAnnot.fst","start_pos":{"line":25,"col":19},"end_pos":{"line":25,"col":24}},"use":{"file_name":"PatAnnot.fst","start_pos":{"line":25,"col":8},"end_pos":{"line":25,"col":9}}},"number":19,"ctx":["While checking for top-level effects","While typechecking the top-level declaration `let whoops`","While typechecking the top-level declaration `[@@expect_failure] let whoops`"]} +>>] +>> Got issues: [ +{"msg":["Assertion failed","The SMT solver could not prove the query. Use --query_stats for more details.","Also see: Prims.fst(459,77-459,89)",""],"level":"Error","range":{"def":{"file_name":"PatAnnot.fst","start_pos":{"line":28,"col":0},"end_pos":{"line":30,"col":14}},"use":{"file_name":"PatAnnot.fst","start_pos":{"line":28,"col":0},"end_pos":{"line":30,"col":14}}},"number":19,"ctx":["While checking for top-level effects","While typechecking the top-level declaration `let whoops2`","While typechecking the top-level declaration `[@@expect_failure] let whoops2`"]} +>>] +>> Got issues: [ +{"msg":["Assertion failed","The SMT solver could not prove the query. Use --query_stats for more details.","Also see: Prims.fst(459,77-459,89)",""],"level":"Error","range":{"def":{"file_name":"PatAnnot.fst","start_pos":{"line":33,"col":0},"end_pos":{"line":35,"col":14}},"use":{"file_name":"PatAnnot.fst","start_pos":{"line":33,"col":0},"end_pos":{"line":35,"col":14}}},"number":19,"ctx":["While checking for top-level effects","While typechecking the top-level declaration `let sub_bv`","While typechecking the top-level declaration `[@@expect_failure] let sub_bv`"]} +>>] +>> Got issues: [ +{"msg":["Subtyping check failed","Expected type Prims.squash Prims.l_False\ngot type Prims.unit","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"PatAnnot.fst","start_pos":{"line":40,"col":26},"end_pos":{"line":40,"col":31}},"use":{"file_name":"PatAnnot.fst","start_pos":{"line":39,"col":10},"end_pos":{"line":39,"col":12}}},"number":19,"ctx":["While checking for top-level effects","While typechecking the top-level declaration `let s`","While typechecking the top-level declaration `[@@expect_failure] let s`"]} +>>] +>> Got issues: [ +{"msg":["Subtyping check failed","Expected type Prims.nat\ngot type Prims.int","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"Prims.fst","start_pos":{"line":680,"col":18},"end_pos":{"line":680,"col":24}},"use":{"file_name":"PatAnnot.fst","start_pos":{"line":46,"col":10},"end_pos":{"line":46,"col":11}}},"number":19,"ctx":["While typechecking the top-level declaration `let test1`","While typechecking the top-level declaration `[@@expect_failure] let test1`"]} +>>] +>> Got issues: [ +{"msg":["Type annotation on parameter incompatible with the expected type","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"Prims.fst","start_pos":{"line":680,"col":18},"end_pos":{"line":680,"col":24}},"use":{"file_name":"PatAnnot.fst","start_pos":{"line":55,"col":36},"end_pos":{"line":55,"col":39}}},"number":19,"ctx":["While checking for top-level effects","While typechecking the top-level declaration `let test3`","While typechecking the top-level declaration `[@@expect_failure] let test3`"]} +>>] +Verified module: PatAnnot +All verification conditions discharged successfully diff --git a/tests/error-messages/PatCoerce.fst.json_expected b/tests/error-messages/PatCoerce.fst.json_expected new file mode 100644 index 00000000000..1fcd79b55ff --- /dev/null +++ b/tests/error-messages/PatCoerce.fst.json_expected @@ -0,0 +1,7 @@ +>> Got issues: [ +{"msg":["Type of pattern (Prims.bool) does not match type of scrutinee (Type0)"],"level":"Error","range":{"def":{"file_name":"PatCoerce.fst","start_pos":{"line":25,"col":4},"end_pos":{"line":25,"col":8}},"use":{"file_name":"PatCoerce.fst","start_pos":{"line":25,"col":4},"end_pos":{"line":25,"col":8}}},"number":114,"ctx":["While typechecking the top-level declaration `let bla`","While typechecking the top-level declaration `[@@expect_failure] let bla`"]} +>>] +{"msg":["PatCoerce.bla\nis declared but no definition was found","Add an 'assume' if this is intentional"],"level":"Warning","range":{"def":{"file_name":"PatCoerce.fst","start_pos":{"line":20,"col":4},"end_pos":{"line":20,"col":7}},"use":{"file_name":"PatCoerce.fst","start_pos":{"line":20,"col":4},"end_pos":{"line":20,"col":7}}},"number":240,"ctx":["While desugaring module PatCoerce"]} +{"msg":["Missing definitions in module PatCoerce: bla"],"level":"Warning","range":{"def":{"file_name":"dummy","start_pos":{"line":0,"col":0},"end_pos":{"line":0,"col":0}},"use":{"file_name":"PatCoerce.fst","start_pos":{"line":23,"col":0},"end_pos":{"line":26,"col":14}}},"number":240,"ctx":[]} +Verified module: PatCoerce +All verification conditions discharged successfully diff --git a/tests/error-messages/PatternMatch.fst.json_expected b/tests/error-messages/PatternMatch.fst.json_expected new file mode 100644 index 00000000000..f4e47fbfced --- /dev/null +++ b/tests/error-messages/PatternMatch.fst.json_expected @@ -0,0 +1,44 @@ +>> Got issues: [ +{"msg":["Type ascriptions within patterns are only allowed on variables"],"level":"Error","range":{"def":{"file_name":"PatternMatch.fst","start_pos":{"line":15,"col":27},"end_pos":{"line":15,"col":34}},"use":{"file_name":"PatternMatch.fst","start_pos":{"line":15,"col":27},"end_pos":{"line":15,"col":34}}},"number":178,"ctx":["While desugaring module PatternMatch"]} +>>] +>> Got issues: [ +{"msg":["Type ascriptions within patterns are only allowed on variables"],"level":"Error","range":{"def":{"file_name":"PatternMatch.fst","start_pos":{"line":18,"col":29},"end_pos":{"line":18,"col":48}},"use":{"file_name":"PatternMatch.fst","start_pos":{"line":18,"col":29},"end_pos":{"line":18,"col":48}}},"number":178,"ctx":["While desugaring module PatternMatch"]} +>>] +>> Got issues: [ +{"msg":["Type ascriptions within patterns are only allowed on variables"],"level":"Error","range":{"def":{"file_name":"PatternMatch.fst","start_pos":{"line":41,"col":3},"end_pos":{"line":41,"col":24}},"use":{"file_name":"PatternMatch.fst","start_pos":{"line":41,"col":3},"end_pos":{"line":41,"col":24}}},"number":178,"ctx":["While desugaring module PatternMatch"]} +>>] +>> Got issues: [ +{"msg":["Patterns are incomplete","The SMT solver could not prove the query. Use --query_stats for more details.","Also see: PatternMatch.fst(21,4-21,8)",""],"level":"Error","range":{"def":{"file_name":"dummy","start_pos":{"line":0,"col":0},"end_pos":{"line":0,"col":0}},"use":{"file_name":"dummy","start_pos":{"line":0,"col":0},"end_pos":{"line":0,"col":0}}},"number":19,"ctx":["While checking for top-level effects","While typechecking the top-level declaration `let uu___7`","While typechecking the top-level declaration `[@@expect_failure] let uu___6`"]} +>>] +>> Got issues: [ +{"msg":["Patterns are incomplete","The SMT solver could not prove the query. Use --query_stats for more details.","Also see: PatternMatch.fst(24,4-24,11)",""],"level":"Error","range":{"def":{"file_name":"dummy","start_pos":{"line":0,"col":0},"end_pos":{"line":0,"col":0}},"use":{"file_name":"dummy","start_pos":{"line":0,"col":0},"end_pos":{"line":0,"col":0}}},"number":19,"ctx":["While checking for top-level effects","While typechecking the top-level declaration `let uu___10`","While typechecking the top-level declaration `[@@expect_failure] let uu___9`"]} +>>] +>> Got issues: [ +{"msg":["Patterns are incomplete","The SMT solver could not prove the query. Use --query_stats for more details.","Also see: PatternMatch.fst(29,4-29,8)",""],"level":"Error","range":{"def":{"file_name":"dummy","start_pos":{"line":0,"col":0},"end_pos":{"line":0,"col":0}},"use":{"file_name":"dummy","start_pos":{"line":0,"col":0},"end_pos":{"line":0,"col":0}}},"number":19,"ctx":["While checking for top-level effects","While typechecking the top-level declaration `let uu___17`","While typechecking the top-level declaration `[@@expect_failure] let uu___16`"]} +>>] +>> Got issues: [ +{"msg":["Type of pattern (Prims.bool) does not match type of scrutinee (Prims.int)"],"level":"Error","range":{"def":{"file_name":"PatternMatch.fst","start_pos":{"line":32,"col":4},"end_pos":{"line":32,"col":9}},"use":{"file_name":"PatternMatch.fst","start_pos":{"line":32,"col":4},"end_pos":{"line":32,"col":9}}},"number":114,"ctx":["While typechecking the top-level declaration `let uu___20`","While typechecking the top-level declaration `[@@expect_failure] let uu___19`"]} +>>] +>> Got issues: [ +{"msg":["Type of pattern PatternMatch.ab\ndoes not match type of scrutinee Prims.int","Head mismatch PatternMatch.ab vs Prims.int"],"level":"Error","range":{"def":{"file_name":"PatternMatch.fst","start_pos":{"line":35,"col":4},"end_pos":{"line":35,"col":5}},"use":{"file_name":"PatternMatch.fst","start_pos":{"line":35,"col":4},"end_pos":{"line":35,"col":5}}},"number":114,"ctx":["While typechecking the top-level declaration `let uu___23`","While typechecking the top-level declaration `[@@expect_failure] let uu___22`"]} +>>] +>> Got issues: [ +{"msg":["Patterns are incomplete","The SMT solver could not prove the query. Use --query_stats for more details.","Also see: PatternMatch.fst(38,4-38,5)",""],"level":"Error","range":{"def":{"file_name":"dummy","start_pos":{"line":0,"col":0},"end_pos":{"line":0,"col":0}},"use":{"file_name":"dummy","start_pos":{"line":0,"col":0},"end_pos":{"line":0,"col":0}}},"number":19,"ctx":["While checking for top-level effects","While typechecking the top-level declaration `let uu___26`","While typechecking the top-level declaration `[@@expect_failure] let uu___25`"]} +>>] +>> Got issues: [ +{"msg":["Type of pattern (Prims.bool) does not match type of scrutinee (Prims.int)"],"level":"Error","range":{"def":{"file_name":"PatternMatch.fst","start_pos":{"line":55,"col":4},"end_pos":{"line":55,"col":8}},"use":{"file_name":"PatternMatch.fst","start_pos":{"line":55,"col":4},"end_pos":{"line":55,"col":8}}},"number":114,"ctx":["While typechecking the top-level declaration `let uu___54`","While typechecking the top-level declaration `[@@expect_failure] let uu___53`"]} +>>] +>> Got issues: [ +{"msg":["Type of pattern (Prims.bool) does not match type of scrutinee (Prims.int)"],"level":"Error","range":{"def":{"file_name":"PatternMatch.fst","start_pos":{"line":58,"col":4},"end_pos":{"line":58,"col":16}},"use":{"file_name":"PatternMatch.fst","start_pos":{"line":58,"col":4},"end_pos":{"line":58,"col":16}}},"number":114,"ctx":["While typechecking the top-level declaration `let uu___57`","While typechecking the top-level declaration `[@@expect_failure] let uu___56`"]} +>>] +>> Got issues: [ +{"msg":["Patterns are incomplete","The SMT solver could not prove the query. Use --query_stats for more details.","Also see: PatternMatch.fst(61,5-61,12)",""],"level":"Error","range":{"def":{"file_name":"dummy","start_pos":{"line":0,"col":0},"end_pos":{"line":0,"col":0}},"use":{"file_name":"dummy","start_pos":{"line":0,"col":0},"end_pos":{"line":0,"col":0}}},"number":19,"ctx":["While checking for top-level effects","While typechecking the top-level declaration `let x`","While typechecking the top-level declaration `[@@expect_failure] let uu___60`"]} +>>] +>> Got issues: [ +{"msg":["Patterns are incomplete","The SMT solver could not prove the query. Use --query_stats for more details.","Also see: PatternMatch.fst(64,5-64,20)",""],"level":"Error","range":{"def":{"file_name":"dummy","start_pos":{"line":0,"col":0},"end_pos":{"line":0,"col":0}},"use":{"file_name":"dummy","start_pos":{"line":0,"col":0},"end_pos":{"line":0,"col":0}}},"number":19,"ctx":["While checking for top-level effects","While typechecking the top-level declaration `let x`","While typechecking the top-level declaration `[@@expect_failure] let uu___64`"]} +>>] +>> Got issues: [ +{"msg":["Type of pattern (Prims.bool) does not match type of scrutinee (Prims.int)"],"level":"Error","range":{"def":{"file_name":"PatternMatch.fst","start_pos":{"line":69,"col":5},"end_pos":{"line":69,"col":15}},"use":{"file_name":"PatternMatch.fst","start_pos":{"line":69,"col":5},"end_pos":{"line":69,"col":15}}},"number":114,"ctx":["While typechecking the top-level declaration `let uu___74`","While typechecking the top-level declaration `[@@expect_failure] let uu___73`"]} +>>] +Verified module: PatternMatch +All verification conditions discharged successfully diff --git a/tests/error-messages/QuickTest.fst.json_expected b/tests/error-messages/QuickTest.fst.json_expected new file mode 100644 index 00000000000..ef157695432 --- /dev/null +++ b/tests/error-messages/QuickTest.fst.json_expected @@ -0,0 +1,5 @@ +>> Got issues: [ +{"msg":["","The SMT solver could not prove the query. Use --query_stats for more details.","Also see: QuickTest(1,2-3,4)","Other related locations: QuickTest.fst(116,12-116,20)"],"level":"Error","range":{"def":{"file_name":"QuickTest.fst","start_pos":{"line":131,"col":2},"end_pos":{"line":134,"col":34}},"use":{"file_name":"QuickTest.fst","start_pos":{"line":131,"col":2},"end_pos":{"line":134,"col":34}}},"number":19,"ctx":["While typechecking the top-level declaration `let va_lemma_Test2`","While typechecking the top-level declaration `[@@expect_failure] let va_lemma_Test2`"]} +>>] +Verified module: QuickTest +All verification conditions discharged successfully diff --git a/tests/error-messages/QuickTestNBE.fst.json_expected b/tests/error-messages/QuickTestNBE.fst.json_expected new file mode 100644 index 00000000000..3571367a50b --- /dev/null +++ b/tests/error-messages/QuickTestNBE.fst.json_expected @@ -0,0 +1,5 @@ +>> Got issues: [ +{"msg":["","The SMT solver could not prove the query. Use --query_stats for more details.","Also see: QuickTestNBE(1,2-3,4)","Other related locations: QuickTestNBE.fst(116,16-116,18)"],"level":"Error","range":{"def":{"file_name":"QuickTestNBE.fst","start_pos":{"line":131,"col":2},"end_pos":{"line":134,"col":34}},"use":{"file_name":"QuickTestNBE.fst","start_pos":{"line":131,"col":2},"end_pos":{"line":134,"col":34}}},"number":19,"ctx":["While typechecking the top-level declaration `let va_lemma_Test2`","While typechecking the top-level declaration `[@@expect_failure] let va_lemma_Test2`"]} +>>] +Verified module: QuickTestNBE +All verification conditions discharged successfully diff --git a/tests/error-messages/RecordFields.fst.json_expected b/tests/error-messages/RecordFields.fst.json_expected new file mode 100644 index 00000000000..10b41ba28ce --- /dev/null +++ b/tests/error-messages/RecordFields.fst.json_expected @@ -0,0 +1,11 @@ +>> Got issues: [ +{"msg":["Field 'd' is redundant for type RecordFields.r",""],"level":"Error","range":{"def":{"file_name":"RecordFields.fst","start_pos":{"line":8,"col":29},"end_pos":{"line":8,"col":30}},"use":{"file_name":"RecordFields.fst","start_pos":{"line":8,"col":29},"end_pos":{"line":8,"col":30}}},"number":118,"ctx":["While typechecking the top-level declaration `let t2`","While typechecking the top-level declaration `[@@expect_failure] let t2`"]} +>>] +>> Got issues: [ +{"msg":["Missing fields for record type 'RecordFields.r': 'c'"],"level":"Error","range":{"def":{"file_name":"RecordFields.fst","start_pos":{"line":11,"col":14},"end_pos":{"line":11,"col":22}},"use":{"file_name":"RecordFields.fst","start_pos":{"line":11,"col":14},"end_pos":{"line":11,"col":22}}},"number":118,"ctx":["While typechecking the top-level declaration `let t3`","While typechecking the top-level declaration `[@@expect_failure] let t3`"]} +>>] +>> Got issues: [ +{"msg":["Field 'd' is redundant for type RecordFields.r","Missing fields: 'c'"],"level":"Error","range":{"def":{"file_name":"RecordFields.fst","start_pos":{"line":14,"col":24},"end_pos":{"line":14,"col":25}},"use":{"file_name":"RecordFields.fst","start_pos":{"line":14,"col":24},"end_pos":{"line":14,"col":25}}},"number":118,"ctx":["While typechecking the top-level declaration `let t4`","While typechecking the top-level declaration `[@@expect_failure] let t4`"]} +>>] +Verified module: RecordFields +All verification conditions discharged successfully diff --git a/tests/error-messages/ResolveImplicitsErrorPos.fst.json_expected b/tests/error-messages/ResolveImplicitsErrorPos.fst.json_expected new file mode 100644 index 00000000000..8e20e660adb --- /dev/null +++ b/tests/error-messages/ResolveImplicitsErrorPos.fst.json_expected @@ -0,0 +1,10 @@ +proof-state: State dump @ depth 0 (at the time of failure): +Location: ResolveImplicitsErrorPos.fst(14,13-16,7) +Goal 1/1 (Instantiating meta argument): +(_: Prims.unit) |- _ : Prims.int + +>> Got issues: [ +{"msg":["Tactic failed","oops"],"level":"Error","range":{"def":{"file_name":"ResolveImplicitsErrorPos.fst","start_pos":{"line":15,"col":10},"end_pos":{"line":15,"col":11}},"use":{"file_name":"ResolveImplicitsErrorPos.fst","start_pos":{"line":15,"col":10},"end_pos":{"line":15,"col":11}}},"number":228,"ctx":["While solving implicits with a tactic","While solving deferred constraints","While typechecking the top-level declaration `let test`","While typechecking the top-level declaration `[@@expect_failure] let test`"]} +>>] +Verified module: ResolveImplicitsErrorPos +All verification conditions discharged successfully diff --git a/tests/error-messages/SMTPatSymbols.fst.json_expected b/tests/error-messages/SMTPatSymbols.fst.json_expected new file mode 100644 index 00000000000..257fd50b486 --- /dev/null +++ b/tests/error-messages/SMTPatSymbols.fst.json_expected @@ -0,0 +1,3 @@ +{"msg":["Pattern uses these theory symbols or terms that should not be in an SMT pattern:\n Prims.op_Addition, Prims.op_Subtraction"],"level":"Warning","range":{"def":{"file_name":"SMTPatSymbols.fst","start_pos":{"line":3,"col":34},"end_pos":{"line":3,"col":52}},"use":{"file_name":"SMTPatSymbols.fst","start_pos":{"line":3,"col":34},"end_pos":{"line":3,"col":52}}},"number":271,"ctx":["While typechecking the top-level declaration `val SMTPatSymbols.lem`"]} +Verified module: SMTPatSymbols +All verification conditions discharged successfully diff --git a/tests/error-messages/SeqLit.fst.json_expected b/tests/error-messages/SeqLit.fst.json_expected new file mode 100644 index 00000000000..62ff1d1e16d --- /dev/null +++ b/tests/error-messages/SeqLit.fst.json_expected @@ -0,0 +1,78 @@ +Module after desugaring: +module SeqLit +Declarations: [ +let x = seq![1; 2; 3] +let y = [1; 2; 3] +private +let _ = assert (FStar.Seq.Base.seq_to_list SeqLit.x == SeqLit.y) +private +let _ = [1; 2] @ [3] +private +let _ = [1; 2; 3] +private +let _ = [1; 2; 3] +private +let _ = [1; 2; 3] +private +let _ = FStar.Seq.Base.append seq![1; 2] seq![3] +private +let _ = seq![1; 2; 3] +private +let _ = seq![1; 2; 3] +private +let _ = seq![1; 2; 3] +] + +Module before type checking: +module SeqLit +Declarations: [ +let x = seq![1; 2; 3] +let y = [1; 2; 3] +private +let _ = assert (FStar.Seq.Base.seq_to_list SeqLit.x == SeqLit.y) +private +let _ = [1; 2] @ [3] +private +let _ = [1; 2; 3] +private +let _ = [1; 2; 3] +private +let _ = [1; 2; 3] +private +let _ = FStar.Seq.Base.append seq![1; 2] seq![3] +private +let _ = seq![1; 2; 3] +private +let _ = seq![1; 2; 3] +private +let _ = seq![1; 2; 3] +] + +Module after type checking: +module SeqLit +Declarations: [ +let x = seq![1; 2; 3] +let y = [1; 2; 3] +private +let _ = assert (FStar.Seq.Base.seq_to_list SeqLit.x == SeqLit.y) +private +let _ = [1; 2] @ [3] +private +let _ = [1; 2; 3] +private +let _ = [1; 2; 3] +private +let _ = [1; 2; 3] +private +let _ = FStar.Seq.Base.append seq![1; 2] seq![3] +private +let _ = seq![1; 2; 3] +private +let _ = seq![1; 2; 3] +private +let _ = seq![1; 2; 3] +] + +{"msg":["The operator '@' has been resolved to FStar.List.Tot.append even though\nFStar.List.Tot is not in scope. Please add an 'open FStar.List.Tot' to stop\nrelying on this deprecated, special treatment of '@'."],"level":"Warning","range":{"def":{"file_name":"SeqLit.fst","start_pos":{"line":8,"col":17},"end_pos":{"line":8,"col":18}},"use":{"file_name":"SeqLit.fst","start_pos":{"line":8,"col":17},"end_pos":{"line":8,"col":18}}},"number":337,"ctx":["While desugaring module SeqLit"]} +Verified module: SeqLit +All verification conditions discharged successfully diff --git a/tests/error-messages/StableErr.fst.json_expected b/tests/error-messages/StableErr.fst.json_expected new file mode 100644 index 00000000000..b3a9827d236 --- /dev/null +++ b/tests/error-messages/StableErr.fst.json_expected @@ -0,0 +1,5 @@ +>> Got issues: [ +{"msg":["A query could not be solved internally, and --no_smt was given.","Query =\nPrims.l_False"],"level":"Error","range":{"def":{"file_name":"StableErr.fst","start_pos":{"line":26,"col":0},"end_pos":{"line":26,"col":20}},"use":{"file_name":"StableErr.fst","start_pos":{"line":26,"col":0},"end_pos":{"line":26,"col":20}}},"number":298,"ctx":["While checking for top-level effects","While typechecking the top-level declaration `let uu___0`","While typechecking the top-level declaration `[@@expect_failure] let uu___0`"]} +>>] +Verified module: StableErr +All verification conditions discharged successfully diff --git a/tests/error-messages/StrictUnfolding.fst.json_expected b/tests/error-messages/StrictUnfolding.fst.json_expected new file mode 100644 index 00000000000..dfaeeb98326 --- /dev/null +++ b/tests/error-messages/StrictUnfolding.fst.json_expected @@ -0,0 +1,25 @@ +proof-state: State dump @ depth 1 (A): +Location: StrictUnfolding.fst(28,8-28,16) +Goal 1/1: +(x: FStar.Pervasives.Native.option (x: FStar.Integers.int{x = 0}) {Some? x}) |- _ : Prims.squash (StrictUnfolding.project x == 0) + +proof-state: State dump @ depth 1 (B): +Location: StrictUnfolding.fst(30,8-30,16) +Goal 1/1: +(x: FStar.Pervasives.Native.option (x: FStar.Integers.int{x = 0}) {Some? x}) |- _ : Prims.squash (StrictUnfolding.project x == 0) + +proof-state: State dump @ depth 1 (A): +Location: StrictUnfolding.fst(35,8-35,16) +Goal 1/1: +(_: Prims.unit), (return_val: x: FStar.Pervasives.Native.option Prims.int {Some? x}), (_: return_val == FStar.Pervasives.Native.Some 0), (any_result: Prims.int), (_: StrictUnfolding.project (FStar.Pervasives.Native.Some 0) == any_result), (any_result'0: Prims.logical), (_: StrictUnfolding.project (FStar.Pervasives.Native.Some 0) == 0 == any_result'0) |- _ : Prims.squash (StrictUnfolding.project (FStar.Pervasives.Native.Some 0) == 0) + +proof-state: State dump @ depth 1 (B): +Location: StrictUnfolding.fst(37,8-37,16) +Goal 1/1: +(_: Prims.unit), (return_val: x: FStar.Pervasives.Native.option Prims.int {Some? x}), (_: return_val == FStar.Pervasives.Native.Some 0), (any_result: Prims.int), (_: StrictUnfolding.project (FStar.Pervasives.Native.Some 0) == any_result), (any_result'0: Prims.logical), (_: StrictUnfolding.project (FStar.Pervasives.Native.Some 0) == 0 == any_result'0) |- _ : Prims.squash (0 == 0) + +>> Got issues: [ +{"msg":["Could not prove goal #1\n","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"StrictUnfolding.fst","start_pos":{"line":50,"col":48},"end_pos":{"line":50,"col":72}},"use":{"file_name":"StrictUnfolding.fst","start_pos":{"line":50,"col":2},"end_pos":{"line":50,"col":8}}},"number":19,"ctx":["While typechecking the top-level declaration `let test_integer_generic_wo_fstar_integers`","While typechecking the top-level declaration `[@@expect_failure] let test_integer_generic_wo_fstar_integers`"]} +>>] +Verified module: StrictUnfolding +All verification conditions discharged successfully diff --git a/tests/error-messages/StringNormalization.fst.json_expected b/tests/error-messages/StringNormalization.fst.json_expected new file mode 100644 index 00000000000..5507d606414 --- /dev/null +++ b/tests/error-messages/StringNormalization.fst.json_expected @@ -0,0 +1,5 @@ +>> Got issues: [ +{"msg":["Assertion failed","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"StringNormalization.fst","start_pos":{"line":83,"col":9},"end_pos":{"line":83,"col":58}},"use":{"file_name":"StringNormalization.fst","start_pos":{"line":83,"col":2},"end_pos":{"line":83,"col":8}}},"number":19,"ctx":["While checking for top-level effects","While typechecking the top-level declaration `let uu___34`","While typechecking the top-level declaration `[@@expect_failure] let uu___34`"]} +>>] +Verified module: StringNormalization +All verification conditions discharged successfully diff --git a/tests/error-messages/Test.FunctionalExtensionality.fst.json_expected b/tests/error-messages/Test.FunctionalExtensionality.fst.json_expected new file mode 100644 index 00000000000..8cf2987cd6c --- /dev/null +++ b/tests/error-messages/Test.FunctionalExtensionality.fst.json_expected @@ -0,0 +1,14 @@ +>> Got issues: [ +{"msg":["Subtyping check failed","Expected type Prims.nat ^-> Prims.int\ngot type Prims.int ^-> Prims.int","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"FStar.FunctionalExtensionality.fsti","start_pos":{"line":102,"col":60},"end_pos":{"line":102,"col":77}},"use":{"file_name":"Test.FunctionalExtensionality.fst","start_pos":{"line":36,"col":49},"end_pos":{"line":36,"col":50}}},"number":19,"ctx":["While typechecking the top-level declaration `let sub_fails`","While typechecking the top-level declaration `[@@expect_failure] let sub_fails`"]} +>>] +>> Got issues: [ +{"msg":["Assertion failed","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"Test.FunctionalExtensionality.fst","start_pos":{"line":80,"col":9},"end_pos":{"line":80,"col":43}},"use":{"file_name":"Test.FunctionalExtensionality.fst","start_pos":{"line":80,"col":2},"end_pos":{"line":80,"col":8}}},"number":19,"ctx":["While checking for top-level effects","While typechecking the top-level declaration `let unable_to_extend_equality_to_larger_domains_1`","While typechecking the top-level declaration `[@@expect_failure] let unable_to_extend_equality_to_larger_domains_1`"]} +>>] +>> Got issues: [ +{"msg":["Subtyping check failed","Expected type _: Prims.int -> Prims.int\ngot type Prims.nat ^-> Prims.int","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"Prims.fst","start_pos":{"line":680,"col":18},"end_pos":{"line":680,"col":24}},"use":{"file_name":"Test.FunctionalExtensionality.fst","start_pos":{"line":92,"col":36},"end_pos":{"line":92,"col":47}}},"number":19,"ctx":["While checking for top-level effects","While typechecking the top-level declaration `let unable_to_extend_equality_to_larger_domains_2`","While typechecking the top-level declaration `[@@expect_failure] let unable_to_extend_equality_to_larger_domains_2`"]} +>>] +>> Got issues: [ +{"msg":["Subtyping check failed","Expected type Prims.int ^-> Prims.int\ngot type Prims.int ^-> Prims.nat","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"FStar.FunctionalExtensionality.fsti","start_pos":{"line":102,"col":60},"end_pos":{"line":102,"col":77}},"use":{"file_name":"Test.FunctionalExtensionality.fst","start_pos":{"line":142,"col":57},"end_pos":{"line":142,"col":58}}},"number":19,"ctx":["While typechecking the top-level declaration `let sub_currently_not`","While typechecking the top-level declaration `[@@expect_failure] let sub_currently_not`"]} +>>] +Verified module: Test.FunctionalExtensionality +All verification conditions discharged successfully diff --git a/tests/error-messages/TestErrorLocations.fst.json_expected b/tests/error-messages/TestErrorLocations.fst.json_expected new file mode 100644 index 00000000000..af547dee013 --- /dev/null +++ b/tests/error-messages/TestErrorLocations.fst.json_expected @@ -0,0 +1,51 @@ +>> Got issues: [ +{"msg":["Assertion failed","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"TestErrorLocations.fst","start_pos":{"line":20,"col":21},"end_pos":{"line":20,"col":23}},"use":{"file_name":"TestErrorLocations.fst","start_pos":{"line":20,"col":21},"end_pos":{"line":20,"col":23}}},"number":19,"ctx":["While checking for top-level effects","While typechecking the top-level declaration `let test0`","While typechecking the top-level declaration `[@@expect_failure] let test0`"]} +>>] +>> Got issues: [ +{"msg":["Assertion failed","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"TestErrorLocations.fst","start_pos":{"line":24,"col":21},"end_pos":{"line":24,"col":23}},"use":{"file_name":"TestErrorLocations.fst","start_pos":{"line":24,"col":21},"end_pos":{"line":24,"col":23}}},"number":19,"ctx":["While checking for top-level effects","While typechecking the top-level declaration `let test1`","While typechecking the top-level declaration `[@@expect_failure] let test1`"]} +>>] +>> Got issues: [ +{"msg":["Assertion failed","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"TestErrorLocations.fst","start_pos":{"line":27,"col":50},"end_pos":{"line":27,"col":58}},"use":{"file_name":"TestErrorLocations.fst","start_pos":{"line":31,"col":10},"end_pos":{"line":31,"col":19}}},"number":19,"ctx":["While typechecking the top-level declaration `let test2`","While typechecking the top-level declaration `[@@expect_failure] let test2`"]} +>>] +>> Got issues: [ +{"msg":["Assertion failed","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"TestErrorLocations.fst","start_pos":{"line":27,"col":50},"end_pos":{"line":27,"col":58}},"use":{"file_name":"TestErrorLocations.fst","start_pos":{"line":37,"col":10},"end_pos":{"line":37,"col":11}}},"number":19,"ctx":["While typechecking the top-level declaration `let test3`","While typechecking the top-level declaration `[@@expect_failure] let test3`"]} +>>] +>> Got issues: [ +{"msg":["Subtyping check failed","Expected type Prims.nat\ngot type Prims.int","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"Prims.fst","start_pos":{"line":680,"col":18},"end_pos":{"line":680,"col":24}},"use":{"file_name":"TestErrorLocations.fst","start_pos":{"line":43,"col":20},"end_pos":{"line":43,"col":21}}},"number":19,"ctx":["While typechecking the top-level declaration `let test4`","While typechecking the top-level declaration `[@@expect_failure] let test4`"]} +>>] +>> Got issues: [ +{"msg":["Could not prove post-condition","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"TestErrorLocations.fst","start_pos":{"line":46,"col":84},"end_pos":{"line":46,"col":90}},"use":{"file_name":"TestErrorLocations.fst","start_pos":{"line":48,"col":14},"end_pos":{"line":48,"col":19}}},"number":19,"ctx":["While typechecking the top-level declaration `let test5`","While typechecking the top-level declaration `[@@expect_failure] let test5`"]} +>>] +>> Got issues: [ +{"msg":["Assertion failed","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"Prims.fst","start_pos":{"line":451,"col":90},"end_pos":{"line":451,"col":102}},"use":{"file_name":"TestErrorLocations.fst","start_pos":{"line":58,"col":15},"end_pos":{"line":58,"col":17}}},"number":19,"ctx":["While typechecking the top-level declaration `let test6`","While typechecking the top-level declaration `[@@expect_failure] let test6`"]} +>>] +>> Got issues: [ +{"msg":["Assertion failed","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"Prims.fst","start_pos":{"line":451,"col":90},"end_pos":{"line":451,"col":102}},"use":{"file_name":"TestErrorLocations.fst","start_pos":{"line":62,"col":15},"end_pos":{"line":62,"col":17}}},"number":19,"ctx":["While typechecking the top-level declaration `let test7`","While typechecking the top-level declaration `[@@expect_failure] let test7`"]} +>>] +>> Got issues: [ +{"msg":["Subtyping check failed","Expected type Prims.nat\ngot type Prims.int","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"Prims.fst","start_pos":{"line":680,"col":18},"end_pos":{"line":680,"col":24}},"use":{"file_name":"TestErrorLocations.fst","start_pos":{"line":66,"col":27},"end_pos":{"line":66,"col":28}}},"number":19,"ctx":["While checking for top-level effects","While typechecking the top-level declaration `let test8`","While typechecking the top-level declaration `[@@expect_failure] let test8`"]} +>>] +>> Got issues: [ +{"msg":["Subtyping check failed","Expected type Type0\ngot type Type0","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"TestErrorLocations.fst","start_pos":{"line":68,"col":52},"end_pos":{"line":68,"col":66}},"use":{"file_name":"TestErrorLocations.fst","start_pos":{"line":70,"col":25},"end_pos":{"line":70,"col":34}}},"number":19,"ctx":["While typechecking the top-level declaration `val TestErrorLocations.test9`","While typechecking the top-level declaration `[@@expect_failure] val TestErrorLocations.test9`"]} +>>] +>> Got issues: [ +{"msg":["Subtyping check failed","Expected type Prims.squash (exists (x: Prims.nat). x = 0)\ngot type Prims.unit","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"FStar.Classical.Sugar.fsti","start_pos":{"line":66,"col":22},"end_pos":{"line":66,"col":41}},"use":{"file_name":"TestErrorLocations.fst","start_pos":{"line":89,"col":20},"end_pos":{"line":89,"col":36}}},"number":19,"ctx":["While typechecking the top-level declaration `let test_elim_exists`","While typechecking the top-level declaration `[@@expect_failure] let test_elim_exists`"]} +>>] +>> Got issues: [ +{"msg":["Subtyping check failed","Expected type Prims.squash (forall (x: Prims.nat). x = 0)\ngot type Prims.unit","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"TestErrorLocations.fst","start_pos":{"line":97,"col":28},"end_pos":{"line":97,"col":33}},"use":{"file_name":"TestErrorLocations.fst","start_pos":{"line":97,"col":28},"end_pos":{"line":97,"col":33}}},"number":19,"ctx":["While typechecking the top-level declaration `let test_elim_forall`","While typechecking the top-level declaration `[@@expect_failure] let test_elim_forall`"]} +>>] +>> Got issues: [ +{"msg":["Subtyping check failed","Expected type Prims.squash (p /\\ q)\ngot type Prims.unit","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"TestErrorLocations.fst","start_pos":{"line":101,"col":19},"end_pos":{"line":101,"col":20}},"use":{"file_name":"TestErrorLocations.fst","start_pos":{"line":102,"col":12},"end_pos":{"line":102,"col":13}}},"number":19,"ctx":["While typechecking the top-level declaration `let test_elim_and`","While typechecking the top-level declaration `[@@expect_failure] let test_elim_and`"]} +>>] +>> Got issues: [ +{"msg":["Subtyping check failed","Expected type Prims.squash (p /\\ q)\ngot type Prims.unit","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"TestErrorLocations.fst","start_pos":{"line":107,"col":21},"end_pos":{"line":107,"col":22}},"use":{"file_name":"TestErrorLocations.fst","start_pos":{"line":108,"col":17},"end_pos":{"line":108,"col":18}}},"number":19,"ctx":["While typechecking the top-level declaration `let test_elim_and`","While typechecking the top-level declaration `[@@expect_failure] let test_elim_and`"]} +>>] +>> Got issues: [ +{"msg":["Subtyping check failed","Expected type Prims.squash (p \\/ q)\ngot type Prims.unit","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"FStar.Classical.Sugar.fsti","start_pos":{"line":88,"col":21},"end_pos":{"line":88,"col":31}},"use":{"file_name":"TestErrorLocations.fst","start_pos":{"line":114,"col":12},"end_pos":{"line":114,"col":18}}},"number":19,"ctx":["While typechecking the top-level declaration `let test_elim_or`","While typechecking the top-level declaration `[@@expect_failure] let test_elim_or`"]} +>>] +{"msg":["TestErrorLocations.test5\nis declared but no definition was found","Add an 'assume' if this is intentional"],"level":"Warning","range":{"def":{"file_name":"TestErrorLocations.fst","start_pos":{"line":46,"col":4},"end_pos":{"line":46,"col":9}},"use":{"file_name":"TestErrorLocations.fst","start_pos":{"line":46,"col":4},"end_pos":{"line":46,"col":9}}},"number":240,"ctx":["While desugaring module TestErrorLocations"]} +{"msg":["TestErrorLocations.test6\nis declared but no definition was found","Add an 'assume' if this is intentional"],"level":"Warning","range":{"def":{"file_name":"TestErrorLocations.fst","start_pos":{"line":56,"col":4},"end_pos":{"line":56,"col":9}},"use":{"file_name":"TestErrorLocations.fst","start_pos":{"line":56,"col":4},"end_pos":{"line":56,"col":9}}},"number":240,"ctx":["While desugaring module TestErrorLocations"]} +{"msg":["TestErrorLocations.test7\nis declared but no definition was found","Add an 'assume' if this is intentional"],"level":"Warning","range":{"def":{"file_name":"TestErrorLocations.fst","start_pos":{"line":60,"col":4},"end_pos":{"line":60,"col":9}},"use":{"file_name":"TestErrorLocations.fst","start_pos":{"line":60,"col":4},"end_pos":{"line":60,"col":9}}},"number":240,"ctx":["While desugaring module TestErrorLocations"]} +{"msg":["Missing definitions in module TestErrorLocations:\n test5\n test6\n test7"],"level":"Warning","range":{"def":{"file_name":"dummy","start_pos":{"line":0,"col":0},"end_pos":{"line":0,"col":0}},"use":{"file_name":"TestErrorLocations.fst","start_pos":{"line":113,"col":0},"end_pos":{"line":117,"col":17}}},"number":240,"ctx":[]} +Verified module: TestErrorLocations +All verification conditions discharged successfully diff --git a/tests/error-messages/TestHasEq.fst.json_expected b/tests/error-messages/TestHasEq.fst.json_expected new file mode 100644 index 00000000000..6ee1378b13a --- /dev/null +++ b/tests/error-messages/TestHasEq.fst.json_expected @@ -0,0 +1,11 @@ +>> Got issues: [ +{"msg":["Failed to prove that the type\n'TestHasEq.t3'\nsupports decidable equality because of this argument.","Add either the 'noeq' or 'unopteq' qualifier","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"TestHasEq.fst","start_pos":{"line":57,"col":0},"end_pos":{"line":58,"col":19}},"use":{"file_name":"TestHasEq.fst","start_pos":{"line":58,"col":10},"end_pos":{"line":58,"col":11}}},"number":19,"ctx":["While typechecking the top-level declaration `type TestHasEq.t3`","While typechecking the top-level declaration `[@@expect_failure] type TestHasEq.t3`"]} +>>] +>> Got issues: [ +{"msg":["Subtyping check failed","Expected type Prims.eqtype\ngot type Type0","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"TestHasEq.fst","start_pos":{"line":84,"col":12},"end_pos":{"line":84,"col":22}},"use":{"file_name":"TestHasEq.fst","start_pos":{"line":84,"col":10},"end_pos":{"line":84,"col":70}}},"number":19,"ctx":["While typechecking the top-level declaration `let test`","While typechecking the top-level declaration `[@@expect_failure] let test`"]} +>>] +>> Got issues: [ +{"msg":["The qualifier list\n[unopteq]\nis not permissible for this element","The `unopteq` qualifier is not allowed on erasable inductives since they don't\nhave decidable equality."],"level":"Error","range":{"def":{"file_name":"TestHasEq.fst","start_pos":{"line":88,"col":8},"end_pos":{"line":89,"col":30}},"use":{"file_name":"TestHasEq.fst","start_pos":{"line":88,"col":8},"end_pos":{"line":89,"col":30}}},"number":162,"ctx":["While typechecking the top-level declaration `type TestHasEq.erasable_t2`","While typechecking the top-level declaration `[@@expect_failure] type TestHasEq.erasable_t2`"]} +>>] +Verified module: TestHasEq +All verification conditions discharged successfully diff --git a/tests/error-messages/UnboundOp.fst.json_expected b/tests/error-messages/UnboundOp.fst.json_expected new file mode 100644 index 00000000000..39e8b8263b2 --- /dev/null +++ b/tests/error-messages/UnboundOp.fst.json_expected @@ -0,0 +1,5 @@ +>> Got issues: [ +{"msg":["Unexpected or unbound operator: ^%^"],"level":"Error","range":{"def":{"file_name":"UnboundOp.fst","start_pos":{"line":4,"col":10},"end_pos":{"line":4,"col":13}},"use":{"file_name":"UnboundOp.fst","start_pos":{"line":4,"col":10},"end_pos":{"line":4,"col":13}}},"number":180,"ctx":["While desugaring module UnboundOp"]} +>>] +Verified module: UnboundOp +All verification conditions discharged successfully diff --git a/tests/error-messages/Unit2.fst.json_expected b/tests/error-messages/Unit2.fst.json_expected new file mode 100644 index 00000000000..57a530df15b --- /dev/null +++ b/tests/error-messages/Unit2.fst.json_expected @@ -0,0 +1,5 @@ +>> Got issues: [ +{"msg":["Assertion failed","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"Unit2.fst","start_pos":{"line":37,"col":21},"end_pos":{"line":38,"col":60}},"use":{"file_name":"Unit2.fst","start_pos":{"line":37,"col":21},"end_pos":{"line":38,"col":60}}},"number":19,"ctx":["While typechecking the top-level declaration `let test6`","While typechecking the top-level declaration `[@@expect_failure] let test6`"]} +>>] +Verified module: Unit2 +All verification conditions discharged successfully diff --git a/tests/error-messages/UnresolvedFields.fst.json_expected b/tests/error-messages/UnresolvedFields.fst.json_expected new file mode 100644 index 00000000000..aa597bddf44 --- /dev/null +++ b/tests/error-messages/UnresolvedFields.fst.json_expected @@ -0,0 +1,11 @@ +>> Got issues: [ +{"msg":["Field name c could not be resolved."],"level":"Error","range":{"def":{"file_name":"UnresolvedFields.fst","start_pos":{"line":10,"col":4},"end_pos":{"line":10,"col":5}},"use":{"file_name":"UnresolvedFields.fst","start_pos":{"line":10,"col":4},"end_pos":{"line":10,"col":5}}},"number":360,"ctx":["While typechecking the top-level declaration `let test0`","While typechecking the top-level declaration `[@@expect_failure] let test0`"]} +>>] +>> Got issues: [ +{"msg":["Field 'c' is redundant for type UnresolvedFields.t",""],"level":"Error","range":{"def":{"file_name":"UnresolvedFields.fst","start_pos":{"line":17,"col":4},"end_pos":{"line":17,"col":5}},"use":{"file_name":"UnresolvedFields.fst","start_pos":{"line":17,"col":4},"end_pos":{"line":17,"col":5}}},"number":118,"ctx":["While typechecking the top-level declaration `let test1`","While typechecking the top-level declaration `[@@expect_failure] let test1`"]} +>>] +>> Got issues: [ +{"msg":["Field 'c' is redundant for type UnresolvedFields.t",""],"level":"Error","range":{"def":{"file_name":"UnresolvedFields.fst","start_pos":{"line":26,"col":4},"end_pos":{"line":26,"col":5}},"use":{"file_name":"UnresolvedFields.fst","start_pos":{"line":26,"col":4},"end_pos":{"line":26,"col":5}}},"number":118,"ctx":["While typechecking the top-level declaration `let test2`","While typechecking the top-level declaration `[@@expect_failure] let test2`"]} +>>] +Verified module: UnresolvedFields +All verification conditions discharged successfully diff --git a/tests/error-messages/WPExtensionality.fst.json_expected b/tests/error-messages/WPExtensionality.fst.json_expected new file mode 100644 index 00000000000..75c78a0d750 --- /dev/null +++ b/tests/error-messages/WPExtensionality.fst.json_expected @@ -0,0 +1,5 @@ +>> Got issues: [ +{"msg":["Assertion failed","The SMT solver could not prove the query. Use --query_stats for more details."],"level":"Error","range":{"def":{"file_name":"Prims.fst","start_pos":{"line":451,"col":90},"end_pos":{"line":451,"col":102}},"use":{"file_name":"WPExtensionality.fst","start_pos":{"line":61,"col":3},"end_pos":{"line":61,"col":34}}},"number":19,"ctx":["While typechecking the top-level declaration `let bug`","While typechecking the top-level declaration `[@@expect_failure] let bug`"]} +>>] +Verified module: WPExtensionality +All verification conditions discharged successfully