From 6352f69ff411168ef31a7e2c83a9bf45d9740767 Mon Sep 17 00:00:00 2001 From: Nikhil Swamy Date: Fri, 10 Mar 2023 13:11:31 -0800 Subject: [PATCH 01/48] Parsing declarations incrementally, in support of the vscode interactive mode --- ocaml/fstar-lib/FStar_Parser_ParseIt.ml | 60 ++++++++++++++++++++++-- ocaml/fstar-lib/FStar_Parser_ParseIt.mli | 38 +++++++++++++++ src/parser/FStar.Parser.AST.fst | 10 ++++ src/parser/FStar.Parser.Driver.fst | 2 + src/parser/FStar.Parser.ParseIt.fsti | 6 ++- src/parser/Makefile | 2 + src/parser/parse.mly | 13 +++-- src/tests/FStar.Tests.Pars.fst | 22 ++++++++- src/tests/boot/FStar.Tests.Test.fst | 1 + 9 files changed, 143 insertions(+), 11 deletions(-) create mode 100755 ocaml/fstar-lib/FStar_Parser_ParseIt.mli diff --git a/ocaml/fstar-lib/FStar_Parser_ParseIt.ml b/ocaml/fstar-lib/FStar_Parser_ParseIt.ml index e65603d191e..acc57abe58f 100644 --- a/ocaml/fstar-lib/FStar_Parser_ParseIt.ml +++ b/ocaml/fstar-lib/FStar_Parser_ParseIt.ml @@ -92,12 +92,16 @@ let check_extension fn = type parse_frag = | Filename of filename | Toplevel of input_frag + | Incremental of input_frag | Fragment of input_frag +type parse_error = (FStar_Errors.raw_error * string * FStar_Compiler_Range.range) + type parse_result = | ASTFragment of (FStar_Parser_AST.inputFragment * (string * FStar_Compiler_Range.range) list) + | IncrementalFragment of (FStar_Parser_AST.decl list * (string * FStar_Compiler_Range.range) list * parse_error option) | Term of FStar_Parser_AST.term - | ParseError of (FStar_Errors.raw_error * string * FStar_Compiler_Range.range) + | ParseError of parse_error let parse fn = FStar_Parser_Util.warningHandler := (function @@ -109,6 +113,7 @@ let parse fn = let f', contents = read_file f in (try create contents f' 1 0, f' with _ -> raise_err (Fatal_InvalidUTF8Encoding, U.format1 "File %s has invalid UTF-8 encoding.\n" f')) + | Incremental s | Toplevel s | Fragment s -> create s.frag_text s.frag_fname (Z.to_int s.frag_line) (Z.to_int s.frag_col), "" @@ -118,12 +123,54 @@ let parse fn = let tok = FStar_Parser_LexFStar.token lexbuf in (tok, lexbuf.start_p, lexbuf.cur_p) in + let err_of_parse_error () = + let pos = FStar_Parser_Util.pos_of_lexpos lexbuf.cur_p in + let r = FStar_Compiler_Range.mk_range filename pos pos in + Fatal_SyntaxError, "Syntax error", r + in + + let parse_incremental_decls () = + let parse_one_decl = MenhirLib.Convert.Simplified.traditional2revised FStar_Parser_Parse.oneDeclOrEOF in + let open FStar_Pervasives in + let rec parse decls = + let _, _, r = err_of_parse_error () in + let d = + try + Inl (parse_one_decl lexer) + with + | FStar_Errors.Error(e, msg, r, _ctx) -> + Inr (e, msg, r) + + | Parsing.Parse_error as _e -> + Inr (err_of_parse_error ()) + in + match d with + | Inl None -> List.rev decls, None + | Inl (Some d) -> + if not (FStar_Parser_AST.decl_syntax_is_delimited d) + then rollback lexbuf; + parse (d::decls) + | Inr err -> List.rev decls, Some err + in + parse [] + in + let parse_incremental_fragment () = + let decls, err_opt = parse_incremental_decls () in + match err_opt with + | None -> + FStar_Parser_AST.as_frag decls + | Some (e, msg, r) -> + raise (FStar_Errors.Error(e, msg, r, [])) + in try match fn with | Filename _ | Toplevel _ -> begin - let fileOrFragment = MenhirLib.Convert.Simplified.traditional2revised FStar_Parser_Parse.inputFragment lexer in + let fileOrFragment = + if FStar_Options.debug_any() + then parse_incremental_fragment () + else MenhirLib.Convert.Simplified.traditional2revised FStar_Parser_Parse.inputFragment lexer in let frags = match fileOrFragment with | FStar_Pervasives.Inl modul -> if has_extension filename interface_extensions @@ -135,6 +182,11 @@ let parse fn = | _ -> fileOrFragment in ASTFragment (frags, FStar_Parser_Util.flush_comments ()) end + + | Incremental _ -> + let decls, err_opt = parse_incremental_decls () in + IncrementalFragment(decls, FStar_Parser_Util.flush_comments(), err_opt) + | Fragment _ -> Term (MenhirLib.Convert.Simplified.traditional2revised FStar_Parser_Parse.term lexer) with @@ -145,9 +197,7 @@ let parse fn = ParseError (e, msg, r) | Parsing.Parse_error as _e -> - let pos = FStar_Parser_Util.pos_of_lexpos lexbuf.cur_p in - let r = FStar_Compiler_Range.mk_range filename pos pos in - ParseError (Fatal_SyntaxError, "Syntax error", r) + ParseError (err_of_parse_error()) (** Parsing of command-line error/warning/silent flags. *) let parse_warn_error s = diff --git a/ocaml/fstar-lib/FStar_Parser_ParseIt.mli b/ocaml/fstar-lib/FStar_Parser_ParseIt.mli new file mode 100755 index 00000000000..8548fff73aa --- /dev/null +++ b/ocaml/fstar-lib/FStar_Parser_ParseIt.mli @@ -0,0 +1,38 @@ +module U = FStar_Compiler_Util +open FStar_Errors +open FStar_Syntax_Syntax +open Lexing +open FStar_Sedlexing + +type filename = string + +type input_frag = { + frag_fname:filename; + frag_text:string; + frag_line:Prims.int; + frag_col:Prims.int +} + +val read_vfs_entry : string -> (U.time * string) option +val add_vfs_entry: string -> string -> unit +val get_file_last_modification_time: string -> U.time + +type parse_frag = + | Filename of filename + | Toplevel of input_frag + | Incremental of input_frag + | Fragment of input_frag + +type parse_error = (FStar_Errors.raw_error * string * FStar_Compiler_Range.range) + +type parse_result = + | ASTFragment of (FStar_Parser_AST.inputFragment * (string * FStar_Compiler_Range.range) list) + | IncrementalFragment of (FStar_Parser_AST.decl list * (string * FStar_Compiler_Range.range) list * parse_error option) + | Term of FStar_Parser_AST.term + | ParseError of parse_error + +val parse: parse_frag -> parse_result + +val find_file: string -> string + +val parse_warn_error: string -> FStar_Errors.error_setting list diff --git a/src/parser/FStar.Parser.AST.fst b/src/parser/FStar.Parser.AST.fst index 9d324009a6c..79b8ecb3d3d 100644 --- a/src/parser/FStar.Parser.AST.fst +++ b/src/parser/FStar.Parser.AST.fst @@ -208,6 +208,7 @@ type lift = { msource: lid; mdest: lid; lift_op: lift_op; + braced: bool; //a detail: for incremental parsing, we need to know if it is delimited by bracces } type pragma = @@ -987,3 +988,12 @@ let ident_of_binder r b = let idents_of_binders bs r = bs |> List.map (ident_of_binder r) + +let decl_syntax_is_delimited (d:decl) = + match d.d with + | Pragma _ + | NewEffect (DefineEffect _) + | LayeredEffect (DefineEffect _) + | SubEffect {braced=true} -> true + | _ -> false + diff --git a/src/parser/FStar.Parser.Driver.fst b/src/parser/FStar.Parser.Driver.fst index d43e424ac61..ebda66f3f0b 100644 --- a/src/parser/FStar.Parser.Driver.fst +++ b/src/parser/FStar.Parser.Driver.fst @@ -35,6 +35,8 @@ let parse_fragment (frag: ParseIt.input_frag) : fragment = Empty | ASTFragment (Inr decls, _) -> //interactive mode: more decls Decls decls + | IncrementalFragment (decls, _, _) -> + Decls decls | ParseError (e, msg, r) -> raise_error (e, msg) r | Term _ -> diff --git a/src/parser/FStar.Parser.ParseIt.fsti b/src/parser/FStar.Parser.ParseIt.fsti index 4d8593573cf..9ad31c188d2 100644 --- a/src/parser/FStar.Parser.ParseIt.fsti +++ b/src/parser/FStar.Parser.ParseIt.fsti @@ -39,12 +39,16 @@ val get_file_last_modification_time: fname:string -> time type parse_frag = | Filename of filename | Toplevel of input_frag + | Incremental of input_frag | Fragment of input_frag +type parse_error = (Errors.raw_error * string * Range.range) + type parse_result = | ASTFragment of (AST.inputFragment * list (string * Range.range)) + | IncrementalFragment of (list AST.decl & list (string * Range.range) & option parse_error) | Term of AST.term - | ParseError of (Errors.raw_error * string * Range.range) + | ParseError of parse_error val parse: parse_frag -> parse_result // either (AST.inputFragment * list (string * Range.range)) , (string * Range.range) val find_file: string -> string diff --git a/src/parser/Makefile b/src/parser/Makefile index c7c971b893b..1bec20525c0 100644 --- a/src/parser/Makefile +++ b/src/parser/Makefile @@ -12,4 +12,6 @@ lex.fs: lex.fsl clean: rm lex.fs parse.fs +FSTAR_HOME ?= ../.. + include ../Makefile.boot.common diff --git a/src/parser/parse.mly b/src/parser/parse.mly index 0baa17c2fca..ea960c9a0e2 100644 --- a/src/parser/parse.mly +++ b/src/parser/parse.mly @@ -128,7 +128,9 @@ let none_to_empty_list x = %start inputFragment %start term %start warn_error_list +%start oneDeclOrEOF %type inputFragment +%type oneDeclOrEOF %type term %type lident %type <(FStar_Errors.flag * string) list> warn_error_list @@ -141,6 +143,9 @@ inputFragment: as_frag decls } +oneDeclOrEOF: + | EOF { None } + | d=decl { Some d } /******************************************************************************/ /* Top level declarations */ @@ -410,7 +415,7 @@ effectDecl: subEffect: | src_eff=quident SQUIGGLY_RARROW tgt_eff=quident EQUALS lift=simpleTerm - { { msource = src_eff; mdest = tgt_eff; lift_op = NonReifiableLift lift } } + { { msource = src_eff; mdest = tgt_eff; lift_op = NonReifiableLift lift; braced=false } } | src_eff=quident SQUIGGLY_RARROW tgt_eff=quident LBRACE lift1=separated_pair(IDENT, EQUALS, simpleTerm) @@ -422,9 +427,9 @@ subEffect: | None -> begin match lift1 with | ("lift", lift) -> - { msource = src_eff; mdest = tgt_eff; lift_op = LiftForFree lift } + { msource = src_eff; mdest = tgt_eff; lift_op = LiftForFree lift; braced=true } | ("lift_wp", lift_wp) -> - { msource = src_eff; mdest = tgt_eff; lift_op = NonReifiableLift lift_wp } + { msource = src_eff; mdest = tgt_eff; lift_op = NonReifiableLift lift_wp; braced=true } | _ -> raise_error (Fatal_UnexpectedIdentifier, "Unexpected identifier; expected {'lift', and possibly 'lift_wp'}") (lhs parseState) end @@ -435,7 +440,7 @@ subEffect: | "lift", "lift_wp" -> tm2, tm1 | _ -> raise_error (Fatal_UnexpectedIdentifier, "Unexpected identifier; expected {'lift', 'lift_wp'}") (lhs parseState) in - { msource = src_eff; mdest = tgt_eff; lift_op = ReifiableLift (lift, lift_wp) } + { msource = src_eff; mdest = tgt_eff; lift_op = ReifiableLift (lift, lift_wp); braced=true } } polymonadic_bind: diff --git a/src/tests/FStar.Tests.Pars.fst b/src/tests/FStar.Tests.Pars.fst index 2a5c5c9d925..c26c73f2250 100644 --- a/src/tests/FStar.Tests.Pars.fst +++ b/src/tests/FStar.Tests.Pars.fst @@ -164,7 +164,7 @@ let pars_and_tc_fragment (s:string) = let tcenv = init() in let frag = frag_of_text s in try - let test_mod', tcenv' = FStar.Universal.tc_one_fragment !test_mod_ref tcenv frag in + let test_mod', tcenv' = FStar.Universal.tc_one_fragment !test_mod_ref tcenv (frag, None) in test_mod_ref := test_mod'; tcenv_ref := Some tcenv'; let n = get_err_count () in @@ -196,3 +196,23 @@ let test_hashes () = aux 100; Options.init() +let parse_incremental_decls () = + let source = + "module Demo\n\ + let f x = match x with | Some x -> true | None -> false\n\ + let test y = if Some? y then f y else true\n\ + let some junk )(" + in + let open FStar.Parser.ParseIt in + let input = Incremental { frag_fname = "Demo.fst"; + frag_text = source; + frag_line = 0; + frag_col = 0 } in + let open FStar.Compiler.Range in + match parse input with + | IncrementalFragment ([d0;d1;d2], _, Some (_, _, rng)) + when (let p = start_of_range rng in + line_of_pos p = 3 && col_of_pos p = 15) -> + () + | _ -> + failwith "Incremental parsing failed" diff --git a/src/tests/boot/FStar.Tests.Test.fst b/src/tests/boot/FStar.Tests.Test.fst index 0dba79a3afe..bcb3a8cd98c 100644 --- a/src/tests/boot/FStar.Tests.Test.fst +++ b/src/tests/boot/FStar.Tests.Test.fst @@ -27,6 +27,7 @@ let main argv = try FStar.Main.setup_hooks(); Pars.init() |> ignore; + Pars.parse_incremental_decls(); Norm.run_all (); if Unif.run_all () then () else exit 1; exit 0 From 3eb4d7531ee4e913b76b9c1a5c70361a62be8f84 Mon Sep 17 00:00:00 2001 From: Nikhil Swamy Date: Fri, 10 Mar 2023 13:12:55 -0800 Subject: [PATCH 02/48] make boot; a quick way to rebuild just fstar.exe during development --- Makefile | 9 +++++++-- 1 file changed, 7 insertions(+), 2 deletions(-) diff --git a/Makefile b/Makefile index 82ed3db8d6d..68eb0035e1b 100644 --- a/Makefile +++ b/Makefile @@ -62,8 +62,13 @@ dune-bootstrap: .PHONY: boot boot: - +$(MAKE) dune - +$(MAKE) dune-bootstrap + +$(MAKE) dune-extract-all + $(Q)cp version.txt $(DUNE_SNAPSHOT)/ + @# Call Dune to build the snapshot. + @echo " DUNE BUILD" + $(Q)cd $(DUNE_SNAPSHOT) && dune build --profile release + @echo " RAW INSTALL" + $(Q)cp ocaml/_build/default/fstar/main.exe $(FSTAR_CURDIR)/bin/fstar.exe install: $(Q)+$(MAKE) -C src/ocaml-output install From db8bbbe109faa7377231e60860ebfcc9d3cbe8c6 Mon Sep 17 00:00:00 2001 From: Nikhil Swamy Date: Fri, 10 Mar 2023 13:14:30 -0800 Subject: [PATCH 03/48] Comparing AST.decl for equality, while ignoring ranges; used for incrementality is fstar-vscode-assistant --- .../generated/FStar_Parser_AST_Comparison.ml | 590 ++++++++++++++++++ src/parser/FStar.Parser.AST.Comparison.fst | 534 ++++++++++++++++ src/parser/FStar.Parser.AST.Comparison.fsti | 24 + 3 files changed, 1148 insertions(+) create mode 100755 ocaml/fstar-lib/generated/FStar_Parser_AST_Comparison.ml create mode 100755 src/parser/FStar.Parser.AST.Comparison.fst create mode 100755 src/parser/FStar.Parser.AST.Comparison.fsti diff --git a/ocaml/fstar-lib/generated/FStar_Parser_AST_Comparison.ml b/ocaml/fstar-lib/generated/FStar_Parser_AST_Comparison.ml new file mode 100755 index 00000000000..d22e09c9edc --- /dev/null +++ b/ocaml/fstar-lib/generated/FStar_Parser_AST_Comparison.ml @@ -0,0 +1,590 @@ +open Prims +let (eq_ident : FStar_Ident.ident -> FStar_Ident.ident -> Prims.bool) = + fun i1 -> fun i2 -> FStar_Ident.ident_equals i1 i2 +let eq_list : + 'a . + ('a -> 'a -> Prims.bool) -> 'a Prims.list -> 'a Prims.list -> Prims.bool + = + fun f -> + fun t1 -> + fun t2 -> + ((FStar_Compiler_List.length t1) = (FStar_Compiler_List.length t2)) + && (FStar_Compiler_List.forall2 f t1 t2) +let eq_option : + 'a . + ('a -> 'a -> Prims.bool) -> + 'a FStar_Pervasives_Native.option -> + 'a FStar_Pervasives_Native.option -> Prims.bool + = + fun f -> + fun t1 -> + fun t2 -> + match (t1, t2) with + | (FStar_Pervasives_Native.None, FStar_Pervasives_Native.None) -> + true + | (FStar_Pervasives_Native.Some t11, FStar_Pervasives_Native.Some + t21) -> f t11 t21 + | uu___ -> false +let (eq_sconst : FStar_Const.sconst -> FStar_Const.sconst -> Prims.bool) = + fun c1 -> fun c2 -> false +let rec (eq_term : + FStar_Parser_AST.term -> FStar_Parser_AST.term -> Prims.bool) = + fun t1 -> fun t2 -> eq_term' t1.FStar_Parser_AST.tm t2.FStar_Parser_AST.tm +and (eq_terms : + FStar_Parser_AST.term Prims.list -> + FStar_Parser_AST.term Prims.list -> Prims.bool) + = fun t1 -> fun t2 -> eq_list eq_term t1 t2 +and (eq_arg : + (FStar_Parser_AST.term * FStar_Parser_AST.imp) -> + (FStar_Parser_AST.term * FStar_Parser_AST.imp) -> Prims.bool) + = + fun t1 -> + fun t2 -> + let uu___ = t1 in + match uu___ with + | (t11, a1) -> + let uu___1 = t2 in + (match uu___1 with + | (t21, a2) -> (eq_term t11 t21) && (eq_imp a1 a2)) +and (eq_imp : FStar_Parser_AST.imp -> FStar_Parser_AST.imp -> Prims.bool) = + fun i1 -> + fun i2 -> + match (i1, i2) with + | (FStar_Parser_AST.FsTypApp, FStar_Parser_AST.FsTypApp) -> true + | (FStar_Parser_AST.Hash, FStar_Parser_AST.Hash) -> true + | (FStar_Parser_AST.UnivApp, FStar_Parser_AST.UnivApp) -> true + | (FStar_Parser_AST.Infix, FStar_Parser_AST.Infix) -> true + | (FStar_Parser_AST.Nothing, FStar_Parser_AST.Nothing) -> true + | (FStar_Parser_AST.HashBrace t1, FStar_Parser_AST.HashBrace t2) -> + eq_term t1 t2 + | uu___ -> false +and (eq_args : + (FStar_Parser_AST.term * FStar_Parser_AST.imp) Prims.list -> + (FStar_Parser_AST.term * FStar_Parser_AST.imp) Prims.list -> Prims.bool) + = fun t1 -> fun t2 -> eq_list eq_arg t1 t2 +and (eq_arg_qualifier : + FStar_Parser_AST.arg_qualifier -> + FStar_Parser_AST.arg_qualifier -> Prims.bool) + = + fun arg_qualifier1 -> + fun arg_qualifier2 -> + match (arg_qualifier1, arg_qualifier2) with + | (FStar_Parser_AST.Implicit, FStar_Parser_AST.Implicit) -> true + | (FStar_Parser_AST.Equality, FStar_Parser_AST.Equality) -> true + | (FStar_Parser_AST.Meta t1, FStar_Parser_AST.Meta t2) -> eq_term t1 t2 + | (FStar_Parser_AST.TypeClassArg, FStar_Parser_AST.TypeClassArg) -> + true + | uu___ -> false +and (eq_pattern : + FStar_Parser_AST.pattern -> FStar_Parser_AST.pattern -> Prims.bool) = + fun p1 -> + fun p2 -> eq_pattern' p1.FStar_Parser_AST.pat p2.FStar_Parser_AST.pat +and (eq_aqual : + FStar_Parser_AST.arg_qualifier FStar_Pervasives_Native.option -> + FStar_Parser_AST.arg_qualifier FStar_Pervasives_Native.option -> + Prims.bool) + = fun a1 -> fun a2 -> eq_option eq_arg_qualifier a1 a2 +and (eq_pattern' : + FStar_Parser_AST.pattern' -> FStar_Parser_AST.pattern' -> Prims.bool) = + fun p1 -> + fun p2 -> + match (p1, p2) with + | (FStar_Parser_AST.PatWild (q1, a1), FStar_Parser_AST.PatWild + (q2, a2)) -> (eq_aqual q1 q2) && (eq_terms a1 a2) + | (FStar_Parser_AST.PatConst s1, FStar_Parser_AST.PatConst s2) -> + eq_sconst s1 s2 + | (FStar_Parser_AST.PatApp (p11, ps1), FStar_Parser_AST.PatApp + (p21, ps2)) -> (eq_pattern p11 p21) && (eq_list eq_pattern ps1 ps2) + | (FStar_Parser_AST.PatTvar (i1, aq1, as1), FStar_Parser_AST.PatTvar + (i2, aq2, as2)) -> + ((FStar_Ident.ident_equals i1 i2) && (eq_aqual aq1 aq2)) && + (eq_terms as1 as2) + | (FStar_Parser_AST.PatVar (i1, aq1, as1), FStar_Parser_AST.PatVar + (i2, aq2, as2)) -> + ((FStar_Ident.ident_equals i1 i2) && (eq_aqual aq1 aq2)) && + (eq_terms as1 as2) + | (FStar_Parser_AST.PatName l1, FStar_Parser_AST.PatName l2) -> + FStar_Ident.lid_equals l1 l2 + | (FStar_Parser_AST.PatOr ps1, FStar_Parser_AST.PatOr ps2) -> + eq_list eq_pattern ps1 ps2 + | (FStar_Parser_AST.PatList ps1, FStar_Parser_AST.PatList ps2) -> + eq_list eq_pattern ps1 ps2 + | (FStar_Parser_AST.PatTuple (ps1, b1), FStar_Parser_AST.PatTuple + (ps2, b2)) -> (eq_list eq_pattern ps1 ps2) && (b1 = b2) + | (FStar_Parser_AST.PatRecord ps1, FStar_Parser_AST.PatRecord ps2) -> + eq_list + (fun uu___ -> + fun uu___1 -> + match (uu___, uu___1) with + | ((l1, p11), (l2, p21)) -> + (FStar_Ident.lid_equals l1 l2) && (eq_pattern p11 p21)) + ps1 ps2 + | (FStar_Parser_AST.PatAscribed (p11, (t1, topt1)), + FStar_Parser_AST.PatAscribed (p21, (t2, topt2))) -> + ((eq_pattern p11 p21) && (eq_term t1 t2)) && + (eq_option eq_term topt1 topt2) + | (FStar_Parser_AST.PatOp i1, FStar_Parser_AST.PatOp i2) -> + eq_ident i1 i2 + | (FStar_Parser_AST.PatVQuote t1, FStar_Parser_AST.PatVQuote t2) -> + eq_term t1 t2 + | uu___ -> false +and (eq_term' : + FStar_Parser_AST.term' -> FStar_Parser_AST.term' -> Prims.bool) = + fun t1 -> + fun t2 -> + match (t1, t2) with + | (FStar_Parser_AST.Wild, FStar_Parser_AST.Wild) -> true + | (FStar_Parser_AST.Const s1, FStar_Parser_AST.Const s2) -> + FStar_Const.eq_const s1 s2 + | (FStar_Parser_AST.Op (i1, ts1), FStar_Parser_AST.Op (i2, ts2)) -> + (eq_ident i1 i2) && (eq_terms ts1 ts2) + | (FStar_Parser_AST.Tvar i1, FStar_Parser_AST.Tvar i2) -> + eq_ident i1 i2 + | (FStar_Parser_AST.Uvar i1, FStar_Parser_AST.Uvar i2) -> + eq_ident i1 i2 + | (FStar_Parser_AST.Var l1, FStar_Parser_AST.Var l2) -> + FStar_Ident.lid_equals l1 l2 + | (FStar_Parser_AST.Name l1, FStar_Parser_AST.Name l2) -> + FStar_Ident.lid_equals l1 l2 + | (FStar_Parser_AST.Projector (l1, i1), FStar_Parser_AST.Projector + (l2, i2)) -> + (FStar_Ident.lid_equals l1 l2) && (FStar_Ident.ident_equals i1 i2) + | (FStar_Parser_AST.Construct (l1, args1), FStar_Parser_AST.Construct + (l2, args2)) -> + (FStar_Ident.lid_equals l1 l2) && (eq_args args1 args2) + | (FStar_Parser_AST.Abs (ps1, t11), FStar_Parser_AST.Abs (ps2, t21)) -> + (eq_list eq_pattern ps1 ps2) && (eq_term t11 t21) + | (FStar_Parser_AST.App (h1, t11, i1), FStar_Parser_AST.App + (h2, t21, i2)) -> + ((eq_term h1 h2) && (eq_term t11 t21)) && (eq_imp i1 i2) + | (FStar_Parser_AST.Let (lq1, defs1, t11), FStar_Parser_AST.Let + (lq2, defs2, t21)) -> + ((lq1 = lq2) && + (eq_list + (fun uu___ -> + fun uu___1 -> + match (uu___, uu___1) with + | ((o1, (p1, t12)), (o2, (p2, t22))) -> + ((eq_option eq_terms o1 o2) && (eq_pattern p1 p2)) + && (eq_term t12 t22)) defs1 defs2)) + && (eq_term t11 t21) + | (FStar_Parser_AST.LetOperator (defs1, t11), + FStar_Parser_AST.LetOperator (defs2, t21)) -> + (eq_list + (fun uu___ -> + fun uu___1 -> + match (uu___, uu___1) with + | ((i1, ps1, t12), (i2, ps2, t22)) -> + ((eq_ident i1 i2) && (eq_pattern ps1 ps2)) && + (eq_term t12 t22)) defs1 defs2) + && (eq_term t11 t21) + | (FStar_Parser_AST.LetOpen (l1, t11), FStar_Parser_AST.LetOpen + (l2, t21)) -> (FStar_Ident.lid_equals l1 l2) && (eq_term t11 t21) + | (FStar_Parser_AST.LetOpenRecord (t11, t21, t3), + FStar_Parser_AST.LetOpenRecord (t4, t5, t6)) -> + ((eq_term t11 t4) && (eq_term t21 t5)) && (eq_term t3 t6) + | (FStar_Parser_AST.Seq (t11, t21), FStar_Parser_AST.Seq (t3, t4)) -> + (eq_term t11 t3) && (eq_term t21 t4) + | (FStar_Parser_AST.Bind (i1, t11, t21), FStar_Parser_AST.Bind + (i2, t3, t4)) -> + ((FStar_Ident.ident_equals i1 i2) && (eq_term t11 t3)) && + (eq_term t21 t4) + | (FStar_Parser_AST.If (t11, i1, mra1, t21, t3), FStar_Parser_AST.If + (t4, i2, mra2, t5, t6)) -> + ((((eq_term t11 t4) && (eq_option eq_ident i1 i2)) && + (eq_option eq_match_returns_annotation mra1 mra2)) + && (eq_term t21 t5)) + && (eq_term t3 t6) + | (FStar_Parser_AST.Match (t11, i1, mra1, bs1), FStar_Parser_AST.Match + (t21, i2, mra2, bs2)) -> + (((eq_term t11 t21) && (eq_option eq_ident i1 i2)) && + (eq_option eq_match_returns_annotation mra1 mra2)) + && (eq_list eq_branch bs1 bs2) + | (FStar_Parser_AST.TryWith (t11, bs1), FStar_Parser_AST.TryWith + (t21, bs2)) -> (eq_term t11 t21) && (eq_list eq_branch bs1 bs2) + | (FStar_Parser_AST.Ascribed (t11, t21, topt1, b1), + FStar_Parser_AST.Ascribed (t3, t4, topt2, b2)) -> + (((eq_term t11 t3) && (eq_term t21 t4)) && + (eq_option eq_term topt1 topt2)) + && (b1 = b2) + | (FStar_Parser_AST.Record (topt1, fs1), FStar_Parser_AST.Record + (topt2, fs2)) -> + (eq_option eq_term topt1 topt2) && + (eq_list + (fun uu___ -> + fun uu___1 -> + match (uu___, uu___1) with + | ((l1, t11), (l2, t21)) -> + (FStar_Ident.lid_equals l1 l2) && (eq_term t11 t21)) + fs1 fs2) + | (FStar_Parser_AST.Project (t11, l1), FStar_Parser_AST.Project + (t21, l2)) -> (eq_term t11 t21) && (FStar_Ident.lid_equals l1 l2) + | (FStar_Parser_AST.Product (bs1, t11), FStar_Parser_AST.Product + (bs2, t21)) -> (eq_list eq_binder bs1 bs2) && (eq_term t11 t21) + | (FStar_Parser_AST.Sum (bs1, t11), FStar_Parser_AST.Sum (bs2, t21)) -> + (eq_list + (fun b1 -> + fun b2 -> + match (b1, b2) with + | (FStar_Pervasives.Inl b11, FStar_Pervasives.Inl b21) -> + eq_binder b11 b21 + | (FStar_Pervasives.Inr t12, FStar_Pervasives.Inr t22) -> + eq_term t12 t22 + | (FStar_Pervasives.Inl uu___, FStar_Pervasives.Inr uu___1) + -> false + | (FStar_Pervasives.Inr uu___, FStar_Pervasives.Inl uu___1) + -> false) bs1 bs2) + && (eq_term t11 t21) + | (FStar_Parser_AST.QForall (bs1, ps1, t11), FStar_Parser_AST.QForall + (bs2, ps2, t21)) -> + let eq_ps uu___ uu___1 = + match (uu___, uu___1) with + | ((is1, ts1), (is2, ts2)) -> + (eq_list eq_ident is1 is2) && + (eq_list (eq_list eq_term) ts1 ts2) in + ((eq_list eq_binder bs1 bs2) && (eq_ps ps1 ps2)) && + (eq_term t11 t21) + | (FStar_Parser_AST.QExists (bs1, ps1, t11), FStar_Parser_AST.QExists + (bs2, ps2, t21)) -> + let eq_ps uu___ uu___1 = + match (uu___, uu___1) with + | ((is1, ts1), (is2, ts2)) -> + (eq_list eq_ident is1 is2) && + (eq_list (eq_list eq_term) ts1 ts2) in + ((eq_list eq_binder bs1 bs2) && (eq_ps ps1 ps2)) && + (eq_term t11 t21) + | (FStar_Parser_AST.Refine (t11, t21), FStar_Parser_AST.Refine + (t3, t4)) -> (eq_binder t11 t3) && (eq_term t21 t4) + | (FStar_Parser_AST.NamedTyp (i1, t11), FStar_Parser_AST.NamedTyp + (i2, t21)) -> (eq_ident i1 i2) && (eq_term t11 t21) + | (FStar_Parser_AST.Paren t11, FStar_Parser_AST.Paren t21) -> + eq_term t11 t21 + | (FStar_Parser_AST.Requires (t11, s1), FStar_Parser_AST.Requires + (t21, s2)) -> (eq_term t11 t21) && (eq_option (=) s1 s2) + | (FStar_Parser_AST.Ensures (t11, s1), FStar_Parser_AST.Ensures + (t21, s2)) -> (eq_term t11 t21) && (eq_option (=) s1 s2) + | (FStar_Parser_AST.LexList ts1, FStar_Parser_AST.LexList ts2) -> + eq_list eq_term ts1 ts2 + | (FStar_Parser_AST.WFOrder (t11, t21), FStar_Parser_AST.WFOrder + (t3, t4)) -> (eq_term t11 t3) && (eq_term t21 t4) + | (FStar_Parser_AST.Decreases (t11, s1), FStar_Parser_AST.Decreases + (t21, s2)) -> (eq_term t11 t21) && (eq_option (=) s1 s2) + | (FStar_Parser_AST.Labeled (t11, s1, b1), FStar_Parser_AST.Labeled + (t21, s2, b2)) -> ((eq_term t11 t21) && (s1 = s2)) && (b1 = b2) + | (FStar_Parser_AST.Discrim l1, FStar_Parser_AST.Discrim l2) -> + FStar_Ident.lid_equals l1 l2 + | (FStar_Parser_AST.Attributes ts1, FStar_Parser_AST.Attributes ts2) -> + eq_list eq_term ts1 ts2 + | (FStar_Parser_AST.Antiquote t11, FStar_Parser_AST.Antiquote t21) -> + eq_term t11 t21 + | (FStar_Parser_AST.Quote (t11, k1), FStar_Parser_AST.Quote (t21, k2)) + -> (eq_term t11 t21) && (k1 = k2) + | (FStar_Parser_AST.VQuote t11, FStar_Parser_AST.VQuote t21) -> + eq_term t11 t21 + | (FStar_Parser_AST.CalcProof (t11, t21, cs1), + FStar_Parser_AST.CalcProof (t3, t4, cs2)) -> + ((eq_term t11 t3) && (eq_term t21 t4)) && + (eq_list eq_calc_step cs1 cs2) + | (FStar_Parser_AST.IntroForall (bs1, t11, t21), + FStar_Parser_AST.IntroForall (bs2, t3, t4)) -> + ((eq_list eq_binder bs1 bs2) && (eq_term t11 t3)) && + (eq_term t21 t4) + | (FStar_Parser_AST.IntroExists (bs1, t11, ts1, t21), + FStar_Parser_AST.IntroExists (bs2, t3, ts2, t4)) -> + (((eq_list eq_binder bs1 bs2) && (eq_term t11 t3)) && + (eq_list eq_term ts1 ts2)) + && (eq_term t21 t4) + | (FStar_Parser_AST.IntroImplies (t11, t21, b1, t3), + FStar_Parser_AST.IntroImplies (t4, t5, b2, t6)) -> + (((eq_term t11 t4) && (eq_term t21 t5)) && (eq_binder b1 b2)) && + (eq_term t3 t6) + | (FStar_Parser_AST.IntroOr (b1, t11, t21, t3), + FStar_Parser_AST.IntroOr (b2, t4, t5, t6)) -> + (((b1 = b2) && (eq_term t11 t4)) && (eq_term t21 t5)) && + (eq_term t3 t6) + | (FStar_Parser_AST.IntroAnd (t11, t21, t3, t4), + FStar_Parser_AST.IntroAnd (t5, t6, t7, t8)) -> + (((eq_term t11 t5) && (eq_term t21 t6)) && (eq_term t3 t7)) && + (eq_term t4 t8) + | (FStar_Parser_AST.ElimForall (bs1, t11, ts1), + FStar_Parser_AST.ElimForall (bs2, t21, ts2)) -> + ((eq_list eq_binder bs1 bs2) && (eq_term t11 t21)) && + (eq_list eq_term ts1 ts2) + | (FStar_Parser_AST.ElimExists (bs1, t11, t21, b1, t3), + FStar_Parser_AST.ElimExists (bs2, t4, t5, b2, t6)) -> + ((((eq_list eq_binder bs1 bs2) && (eq_term t11 t4)) && + (eq_term t21 t5)) + && (eq_binder b1 b2)) + && (eq_term t3 t6) + | (FStar_Parser_AST.ElimImplies (t11, t21, t3), + FStar_Parser_AST.ElimImplies (t4, t5, t6)) -> + ((eq_term t11 t4) && (eq_term t21 t5)) && (eq_term t3 t6) + | (FStar_Parser_AST.ElimOr (t11, t21, t3, b1, t4, b2, t5), + FStar_Parser_AST.ElimOr (t6, t7, t8, b3, t9, b4, t10)) -> + ((((((eq_term t11 t6) && (eq_term t21 t7)) && (eq_term t3 t8)) && + (eq_binder b1 b3)) + && (eq_term t4 t9)) + && (eq_binder b2 b4)) + && (eq_term t5 t10) + | (FStar_Parser_AST.ElimAnd (t11, t21, t3, b1, b2, t4), + FStar_Parser_AST.ElimAnd (t5, t6, t7, b3, b4, t8)) -> + (((((eq_term t11 t5) && (eq_term t21 t6)) && (eq_term t3 t7)) && + (eq_binder b1 b3)) + && (eq_binder b2 b4)) + && (eq_term t4 t8) + | uu___ -> false +and (eq_calc_step : + FStar_Parser_AST.calc_step -> FStar_Parser_AST.calc_step -> Prims.bool) = + fun uu___ -> + fun uu___1 -> + match (uu___, uu___1) with + | (FStar_Parser_AST.CalcStep (t1, t2, t3), FStar_Parser_AST.CalcStep + (t4, t5, t6)) -> + ((eq_term t1 t4) && (eq_term t2 t5)) && (eq_term t3 t6) +and (eq_binder : + FStar_Parser_AST.binder -> FStar_Parser_AST.binder -> Prims.bool) = + fun b1 -> fun b2 -> eq_binder' b1.FStar_Parser_AST.b b2.FStar_Parser_AST.b +and (eq_binder' : + FStar_Parser_AST.binder' -> FStar_Parser_AST.binder' -> Prims.bool) = + fun b1 -> + fun b2 -> + match (b1, b2) with + | (FStar_Parser_AST.Variable i1, FStar_Parser_AST.Variable i2) -> + eq_ident i1 i2 + | (FStar_Parser_AST.TVariable i1, FStar_Parser_AST.TVariable i2) -> + eq_ident i1 i2 + | (FStar_Parser_AST.Annotated (i1, t1), FStar_Parser_AST.Annotated + (i2, t2)) -> (eq_ident i1 i2) && (eq_term t1 t2) + | (FStar_Parser_AST.TAnnotated (i1, t1), FStar_Parser_AST.TAnnotated + (i2, t2)) -> (eq_ident i1 i2) && (eq_term t1 t2) + | (FStar_Parser_AST.NoName t1, FStar_Parser_AST.NoName t2) -> + eq_term t1 t2 + | uu___ -> false +and (eq_match_returns_annotation : + (FStar_Ident.ident FStar_Pervasives_Native.option * FStar_Parser_AST.term * + Prims.bool) -> + (FStar_Ident.ident FStar_Pervasives_Native.option * FStar_Parser_AST.term + * Prims.bool) -> Prims.bool) + = + fun uu___ -> + fun uu___1 -> + match (uu___, uu___1) with + | ((i1, t1, b1), (i2, t2, b2)) -> + ((eq_option eq_ident i1 i2) && (eq_term t1 t2)) && (b1 = b2) +and (eq_branch : + (FStar_Parser_AST.pattern * FStar_Parser_AST.term + FStar_Pervasives_Native.option * FStar_Parser_AST.term) -> + (FStar_Parser_AST.pattern * FStar_Parser_AST.term + FStar_Pervasives_Native.option * FStar_Parser_AST.term) -> Prims.bool) + = + fun uu___ -> + fun uu___1 -> + match (uu___, uu___1) with + | ((p1, o1, t1), (p2, o2, t2)) -> + ((eq_pattern p1 p2) && (eq_option eq_term o1 o2)) && + (eq_term t1 t2) +let (eq_tycon_record : + FStar_Parser_AST.tycon_record -> + FStar_Parser_AST.tycon_record -> Prims.bool) + = + fun t1 -> + fun t2 -> + eq_list + (fun uu___ -> + fun uu___1 -> + match (uu___, uu___1) with + | ((i1, a1, a2, t11), (i2, a3, a4, t21)) -> + (((eq_ident i1 i2) && (eq_aqual a1 a3)) && + (eq_list eq_term a2 a4)) + && (eq_term t11 t21)) t1 t2 +let (eq_constructor_payload : + FStar_Parser_AST.constructor_payload -> + FStar_Parser_AST.constructor_payload -> Prims.bool) + = + fun t1 -> + fun t2 -> + match (t1, t2) with + | (FStar_Parser_AST.VpOfNotation t11, FStar_Parser_AST.VpOfNotation + t21) -> eq_term t11 t21 + | (FStar_Parser_AST.VpArbitrary t11, FStar_Parser_AST.VpArbitrary t21) + -> eq_term t11 t21 + | (FStar_Parser_AST.VpRecord (r1, k1), FStar_Parser_AST.VpRecord + (r2, k2)) -> (eq_tycon_record r1 r2) && (eq_option eq_term k1 k2) + | uu___ -> false +let (eq_tycon : + FStar_Parser_AST.tycon -> FStar_Parser_AST.tycon -> Prims.bool) = + fun t1 -> + fun t2 -> + match (t1, t2) with + | (FStar_Parser_AST.TyconAbstract (i1, bs1, k1), + FStar_Parser_AST.TyconAbstract (i2, bs2, k2)) -> + ((eq_ident i1 i2) && (eq_list eq_binder bs1 bs2)) && + (eq_option eq_term k1 k2) + | (FStar_Parser_AST.TyconAbbrev (i1, bs1, k1, t11), + FStar_Parser_AST.TyconAbbrev (i2, bs2, k2, t21)) -> + (((eq_ident i1 i2) && (eq_list eq_binder bs1 bs2)) && + (eq_option eq_term k1 k2)) + && (eq_term t11 t21) + | (FStar_Parser_AST.TyconRecord (i1, bs1, k1, a1, r1), + FStar_Parser_AST.TyconRecord (i2, bs2, k2, a2, r2)) -> + ((((eq_ident i1 i2) && (eq_list eq_binder bs1 bs2)) && + (eq_option eq_term k1 k2)) + && (eq_list eq_term a1 a2)) + && (eq_tycon_record r1 r2) + | (FStar_Parser_AST.TyconVariant (i1, bs1, k1, cs1), + FStar_Parser_AST.TyconVariant (i2, bs2, k2, cs2)) -> + (((eq_ident i1 i2) && (eq_list eq_binder bs1 bs2)) && + (eq_option eq_term k1 k2)) + && + (eq_list + (fun uu___ -> + fun uu___1 -> + match (uu___, uu___1) with + | ((i11, o1, a1), (i21, o2, a2)) -> + ((eq_ident i11 i21) && + (eq_option eq_constructor_payload o1 o2)) + && (eq_list eq_term a1 a2)) cs1 cs2) + | uu___ -> false +let (eq_lid : FStar_Ident.lident -> FStar_Ident.lident -> Prims.bool) = + FStar_Ident.lid_equals +let (eq_lift : FStar_Parser_AST.lift -> FStar_Parser_AST.lift -> Prims.bool) + = + fun t1 -> + fun t2 -> + ((eq_lid t1.FStar_Parser_AST.msource t2.FStar_Parser_AST.msource) && + (eq_lid t1.FStar_Parser_AST.mdest t2.FStar_Parser_AST.mdest)) + && + (match ((t1.FStar_Parser_AST.lift_op), (t2.FStar_Parser_AST.lift_op)) + with + | (FStar_Parser_AST.NonReifiableLift t11, + FStar_Parser_AST.NonReifiableLift t21) -> eq_term t11 t21 + | (FStar_Parser_AST.ReifiableLift (t11, t21), + FStar_Parser_AST.ReifiableLift (t3, t4)) -> + (eq_term t11 t3) && (eq_term t21 t4) + | (FStar_Parser_AST.LiftForFree t11, FStar_Parser_AST.LiftForFree + t21) -> eq_term t11 t21 + | uu___ -> false) +let (eq_pragma : + FStar_Parser_AST.pragma -> FStar_Parser_AST.pragma -> Prims.bool) = + fun t1 -> + fun t2 -> + match (t1, t2) with + | (FStar_Parser_AST.SetOptions s1, FStar_Parser_AST.SetOptions s2) -> + s1 = s2 + | (FStar_Parser_AST.ResetOptions s1, FStar_Parser_AST.ResetOptions s2) + -> eq_option (fun s11 -> fun s21 -> s11 = s21) s1 s2 + | (FStar_Parser_AST.PushOptions s1, FStar_Parser_AST.PushOptions s2) -> + eq_option (fun s11 -> fun s21 -> s11 = s21) s1 s2 + | (FStar_Parser_AST.PopOptions, FStar_Parser_AST.PopOptions) -> true + | (FStar_Parser_AST.RestartSolver, FStar_Parser_AST.RestartSolver) -> + true + | (FStar_Parser_AST.PrintEffectsGraph, + FStar_Parser_AST.PrintEffectsGraph) -> true + | uu___ -> false +let (eq_qualifier : + FStar_Parser_AST.qualifier -> FStar_Parser_AST.qualifier -> Prims.bool) = + fun t1 -> + fun t2 -> + match (t1, t2) with + | (FStar_Parser_AST.Private, FStar_Parser_AST.Private) -> true + | (FStar_Parser_AST.Noeq, FStar_Parser_AST.Noeq) -> true + | (FStar_Parser_AST.Unopteq, FStar_Parser_AST.Unopteq) -> true + | (FStar_Parser_AST.Assumption, FStar_Parser_AST.Assumption) -> true + | (FStar_Parser_AST.DefaultEffect, FStar_Parser_AST.DefaultEffect) -> + true + | (FStar_Parser_AST.TotalEffect, FStar_Parser_AST.TotalEffect) -> true + | (FStar_Parser_AST.Effect_qual, FStar_Parser_AST.Effect_qual) -> true + | (FStar_Parser_AST.New, FStar_Parser_AST.New) -> true + | (FStar_Parser_AST.Inline, FStar_Parser_AST.Inline) -> true + | (FStar_Parser_AST.Visible, FStar_Parser_AST.Visible) -> true + | (FStar_Parser_AST.Unfold_for_unification_and_vcgen, + FStar_Parser_AST.Unfold_for_unification_and_vcgen) -> true + | (FStar_Parser_AST.Inline_for_extraction, + FStar_Parser_AST.Inline_for_extraction) -> true + | (FStar_Parser_AST.Irreducible, FStar_Parser_AST.Irreducible) -> true + | (FStar_Parser_AST.NoExtract, FStar_Parser_AST.NoExtract) -> true + | (FStar_Parser_AST.Reifiable, FStar_Parser_AST.Reifiable) -> true + | (FStar_Parser_AST.Reflectable, FStar_Parser_AST.Reflectable) -> true + | (FStar_Parser_AST.Opaque, FStar_Parser_AST.Opaque) -> true + | (FStar_Parser_AST.Logic, FStar_Parser_AST.Logic) -> true + | uu___ -> false +let (eq_qualifiers : + FStar_Parser_AST.qualifiers -> FStar_Parser_AST.qualifiers -> Prims.bool) = + fun t1 -> fun t2 -> eq_list eq_qualifier t1 t2 +let rec (eq_decl' : + FStar_Parser_AST.decl' -> FStar_Parser_AST.decl' -> Prims.bool) = + fun d1 -> + fun d2 -> + match (d1, d2) with + | (FStar_Parser_AST.TopLevelModule lid1, + FStar_Parser_AST.TopLevelModule lid2) -> eq_lid lid1 lid2 + | (FStar_Parser_AST.Open lid1, FStar_Parser_AST.Open lid2) -> + eq_lid lid1 lid2 + | (FStar_Parser_AST.Friend lid1, FStar_Parser_AST.Friend lid2) -> + eq_lid lid1 lid2 + | (FStar_Parser_AST.Include lid1, FStar_Parser_AST.Include lid2) -> + eq_lid lid1 lid2 + | (FStar_Parser_AST.ModuleAbbrev (i1, lid1), + FStar_Parser_AST.ModuleAbbrev (i2, lid2)) -> + (eq_ident i1 i2) && (eq_lid lid1 lid2) + | (FStar_Parser_AST.TopLevelLet (lq1, pats1), + FStar_Parser_AST.TopLevelLet (lq2, pats2)) -> + (lq1 = lq2) && + (eq_list + (fun uu___ -> + fun uu___1 -> + match (uu___, uu___1) with + | ((p1, t1), (p2, t2)) -> + (eq_pattern p1 p2) && (eq_term t1 t2)) pats1 pats2) + | (FStar_Parser_AST.Tycon (b1, b2, tcs1), FStar_Parser_AST.Tycon + (b3, b4, tcs2)) -> + ((b1 = b3) && (b2 = b4)) && (eq_list eq_tycon tcs1 tcs2) + | (FStar_Parser_AST.Val (i1, t1), FStar_Parser_AST.Val (i2, t2)) -> + (eq_ident i1 i2) && (eq_term t1 t2) + | (FStar_Parser_AST.Exception (i1, t1), FStar_Parser_AST.Exception + (i2, t2)) -> (eq_ident i1 i2) && (eq_option eq_term t1 t2) + | (FStar_Parser_AST.NewEffect ed1, FStar_Parser_AST.NewEffect ed2) -> + eq_effect_decl ed1 ed2 + | (FStar_Parser_AST.LayeredEffect ed1, FStar_Parser_AST.LayeredEffect + ed2) -> eq_effect_decl ed1 ed2 + | (FStar_Parser_AST.SubEffect l1, FStar_Parser_AST.SubEffect l2) -> + eq_lift l1 l2 + | (FStar_Parser_AST.Polymonadic_bind (lid1, lid2, lid3, t1), + FStar_Parser_AST.Polymonadic_bind (lid4, lid5, lid6, t2)) -> + (((eq_lid lid1 lid4) && (eq_lid lid2 lid5)) && (eq_lid lid3 lid6)) + && (eq_term t1 t2) + | (FStar_Parser_AST.Polymonadic_subcomp (lid1, lid2, t1), + FStar_Parser_AST.Polymonadic_subcomp (lid3, lid4, t2)) -> + ((eq_lid lid1 lid3) && (eq_lid lid2 lid4)) && (eq_term t1 t2) + | (FStar_Parser_AST.Pragma p1, FStar_Parser_AST.Pragma p2) -> + eq_pragma p1 p2 + | (FStar_Parser_AST.Assume (i1, t1), FStar_Parser_AST.Assume (i2, t2)) + -> (eq_ident i1 i2) && (eq_term t1 t2) + | (FStar_Parser_AST.Splice (is1, t1), FStar_Parser_AST.Splice + (is2, t2)) -> (eq_list eq_ident is1 is2) && (eq_term t1 t2) + | uu___ -> false +and (eq_effect_decl : + FStar_Parser_AST.effect_decl -> FStar_Parser_AST.effect_decl -> Prims.bool) + = + fun t1 -> + fun t2 -> + match (t1, t2) with + | (FStar_Parser_AST.DefineEffect (i1, bs1, t11, ds1), + FStar_Parser_AST.DefineEffect (i2, bs2, t21, ds2)) -> + (((eq_ident i1 i2) && (eq_list eq_binder bs1 bs2)) && + (eq_term t11 t21)) + && (eq_list eq_decl ds1 ds2) + | (FStar_Parser_AST.RedefineEffect (i1, bs1, t11), + FStar_Parser_AST.RedefineEffect (i2, bs2, t21)) -> + ((eq_ident i1 i2) && (eq_list eq_binder bs1 bs2)) && + (eq_term t11 t21) + | uu___ -> false +and (eq_decl : FStar_Parser_AST.decl -> FStar_Parser_AST.decl -> Prims.bool) + = + fun d1 -> + fun d2 -> + ((eq_decl' d1.FStar_Parser_AST.d d2.FStar_Parser_AST.d) && + (eq_list eq_qualifier d1.FStar_Parser_AST.quals + d2.FStar_Parser_AST.quals)) + && + (eq_list eq_term d1.FStar_Parser_AST.attrs d2.FStar_Parser_AST.attrs) \ No newline at end of file diff --git a/src/parser/FStar.Parser.AST.Comparison.fst b/src/parser/FStar.Parser.AST.Comparison.fst new file mode 100755 index 00000000000..4e0e636fe84 --- /dev/null +++ b/src/parser/FStar.Parser.AST.Comparison.fst @@ -0,0 +1,534 @@ +(* + Copyright 2023 Microsoft Research + + Licensed under the Apache License, Version 2.0 (the "License"); + you may not use this file except in compliance with the License. + You may obtain a copy of the License at + + http://www.apache.org/licenses/LICENSE-2.0 + + Unless required by applicable law or agreed to in writing, software + distributed under the License is distributed on an "AS IS" BASIS, + WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or implied. + See the License for the specific language governing permissions and + limitations under the License. + + Authors: N. Swamy and Copilot +*) +module FStar.Parser.AST.Comparison +open FStar.Pervasives +open FStar.Compiler.Effect +open FStar.Compiler.List +open FStar.Errors +module C = FStar.Parser.Const +open FStar.Compiler.Range +open FStar.Ident +open FStar +open FStar.Compiler +open FStar.Compiler.Util +open FStar.Const +open FStar.Parser.AST + +let eq_ident (i1 i2:ident) = + Ident.ident_equals i1 i2 + +let eq_list (f: 'a -> 'a -> bool) (t1 t2:list 'a) + : bool + = List.length t1 = List.length t2 && + List.forall2 f t1 t2 + +// eq_opt: A equality on option a +let eq_option (f: 'a -> 'a -> bool) (t1 t2:option 'a) + : bool + = match t1, t2 with + | None, None -> true + | Some t1, Some t2 -> f t1 t2 + | _ -> false + +let eq_sconst (c1 c2: sconst) : bool = false + +let rec eq_term (t1 t2:term) + : bool + = eq_term' t1.tm t2.tm + +and eq_terms (t1 t2:list term) + : bool + = eq_list eq_term t1 t2 + +and eq_arg (t1 t2 : (term & imp)) + = let t1, a1 = t1 in + let t2, a2 = t2 in + eq_term t1 t2 && + eq_imp a1 a2 + +and eq_imp (i1 i2: imp) + = match i1, i2 with + | FsTypApp, FsTypApp + | Hash, Hash + | UnivApp, UnivApp + | Infix, Infix + | Nothing, Nothing -> true + | HashBrace t1, HashBrace t2 -> + eq_term t1 t2 + | _ -> false + +and eq_args (t1 t2: list (term & imp)) + : bool + = eq_list eq_arg t1 t2 + +and eq_arg_qualifier (arg_qualifier1:arg_qualifier) (arg_qualifier2:arg_qualifier) : bool = + match arg_qualifier1, arg_qualifier2 with + | Implicit, Implicit -> true + | Equality, Equality -> true + | Meta t1, Meta t2 -> eq_term t1 t2 + | TypeClassArg, TypeClassArg -> true + | _ -> false + +and eq_pattern (p1 p2:pattern) + : bool + = eq_pattern' p1.pat p2.pat + +and eq_aqual a1 a2 = eq_option eq_arg_qualifier a1 a2 + +and eq_pattern' (p1 p2:pattern') + : bool + = match p1, p2 with + | PatWild(q1, a1), PatWild(q2, a2) -> + eq_aqual q1 q2 && + eq_terms a1 a2 + | PatConst s1, PatConst s2 -> + eq_sconst s1 s2 + | PatApp (p1, ps1), PatApp(p2, ps2) -> + eq_pattern p1 p2 && + eq_list eq_pattern ps1 ps2 + | PatTvar (i1, aq1, as1), PatTvar(i2, aq2, as2) + | PatVar (i1, aq1, as1), PatVar(i2, aq2, as2) -> + Ident.ident_equals i1 i2 && + eq_aqual aq1 aq2 && + eq_terms as1 as2 + | PatName l1, PatName l2 -> + Ident.lid_equals l1 l2 + | PatOr ps1, PatOr ps2 + | PatList ps1, PatList ps2 -> + eq_list eq_pattern ps1 ps2 + | PatTuple(ps1, b1), PatTuple(ps2, b2) -> + eq_list eq_pattern ps1 ps2 && + b1 = b2 + | PatRecord ps1, PatRecord ps2 -> + eq_list (fun (l1, p1) (l2, p2) -> + Ident.lid_equals l1 l2 && + eq_pattern p1 p2) + ps1 ps2 + | PatAscribed (p1, (t1, topt1)), PatAscribed (p2, (t2, topt2)) -> + eq_pattern p1 p2 && + eq_term t1 t2 && + eq_option eq_term topt1 topt2 + | PatOp i1, PatOp i2 -> + eq_ident i1 i2 + | PatVQuote t1, PatVQuote t2 -> + eq_term t1 t2 + | _ -> false + +and eq_term' (t1 t2:term') + : bool + = match t1, t2 with + | Wild, Wild -> true + | Const s1, Const s2 -> eq_const s1 s2 + | Op (i1, ts1), Op (i2, ts2) -> + eq_ident i1 i2 && + eq_terms ts1 ts2 + | Tvar i1, Tvar i2 + | Uvar i1, Uvar i2 -> + eq_ident i1 i2 + | Var l1, Var l2 + | Name l1, Name l2 -> + Ident.lid_equals l1 l2 + | Projector (l1, i1), Projector (l2, i2) -> + Ident.lid_equals l1 l2 && + Ident.ident_equals i1 i2 + | Construct (l1, args1), Construct (l2, args2) -> + Ident.lid_equals l1 l2 && + eq_args args1 args2 + | Abs (ps1, t1), Abs (ps2, t2) -> + eq_list eq_pattern ps1 ps2 && + eq_term t1 t2 + | App (h1, t1, i1), App(h2, t2, i2) -> + eq_term h1 h2 && + eq_term t1 t2 && + eq_imp i1 i2 + | Let (lq1, defs1, t1), Let (lq2, defs2, t2) -> + lq1=lq2 && + eq_list (fun (o1, (p1, t1)) (o2, (p2, t2)) -> + eq_option eq_terms o1 o2 && + eq_pattern p1 p2 && + eq_term t1 t2) + defs1 defs2 && + eq_term t1 t2 + | LetOperator (defs1, t1), LetOperator (defs2, t2) -> + eq_list (fun (i1, ps1, t1) (i2, ps2, t2) -> + eq_ident i1 i2 && + eq_pattern ps1 ps2 && + eq_term t1 t2) + defs1 defs2 && + eq_term t1 t2 + | LetOpen (l1, t1), LetOpen (l2, t2) -> + Ident.lid_equals l1 l2 && + eq_term t1 t2 + // compare all the remaining cases of terms starting with LetOperator + | LetOpenRecord (t1, t2, t3), LetOpenRecord (t4, t5, t6) -> + eq_term t1 t4 && + eq_term t2 t5 && + eq_term t3 t6 + | Seq (t1, t2), Seq (t3, t4) -> + eq_term t1 t3 && + eq_term t2 t4 + | Bind (i1, t1, t2), Bind (i2, t3, t4) -> + Ident.ident_equals i1 i2 && + eq_term t1 t3 && + eq_term t2 t4 + | If (t1, i1, mra1, t2, t3), If (t4, i2, mra2, t5, t6) -> + eq_term t1 t4 && + eq_option eq_ident i1 i2 && + eq_option eq_match_returns_annotation mra1 mra2 && + eq_term t2 t5 && + eq_term t3 t6 + | Match (t1, i1, mra1, bs1), Match (t2, i2, mra2, bs2) -> + eq_term t1 t2 && + eq_option eq_ident i1 i2 && + eq_option eq_match_returns_annotation mra1 mra2 && + eq_list eq_branch bs1 bs2 + | TryWith (t1, bs1), TryWith (t2, bs2) -> + eq_term t1 t2 && + eq_list eq_branch bs1 bs2 + | Ascribed (t1, t2, topt1, b1), Ascribed (t3, t4, topt2, b2) -> + eq_term t1 t3 && + eq_term t2 t4 && + eq_option eq_term topt1 topt2 && + b1 = b2 + | Record (topt1, fs1), Record (topt2, fs2) -> + eq_option eq_term topt1 topt2 && + eq_list (fun (l1, t1) (l2, t2) -> + Ident.lid_equals l1 l2 && + eq_term t1 t2) + fs1 fs2 + | Project (t1, l1), Project (t2, l2) -> + eq_term t1 t2 && + Ident.lid_equals l1 l2 + | Product (bs1, t1), Product (bs2, t2) -> + eq_list eq_binder bs1 bs2 && + eq_term t1 t2 + | Sum (bs1, t1), Sum (bs2, t2) -> + eq_list (fun b1 b2 -> + match b1, b2 with + | Inl b1, Inl b2 -> + eq_binder b1 b2 + | Inr t1, Inr t2 -> + eq_term t1 t2 + | Inl _, Inr _ + | Inr _, Inl _ -> + false) + bs1 bs2 && + eq_term t1 t2 + | QForall (bs1, ps1, t1), QForall (bs2, ps2, t2) + | QExists (bs1, ps1, t1), QExists (bs2, ps2, t2) -> + //ps1 and ps2 have type list ident * list (list term) + // generate equality on ps1 p2 + let eq_ps (is1, ts1) (is2, ts2) = + eq_list eq_ident is1 is2 && + eq_list (eq_list eq_term) ts1 ts2 + in + eq_list eq_binder bs1 bs2 && + eq_ps ps1 ps2 && + eq_term t1 t2 + // continue comparing all the remaining cases of terms, starting with Refine + | Refine (t1, t2), Refine (t3, t4) -> + eq_binder t1 t3 && + eq_term t2 t4 + // continue comparing all the remaining cases of terms, starting with NamedTyp + | NamedTyp (i1, t1), NamedTyp (i2, t2) -> + eq_ident i1 i2 && + eq_term t1 t2 + | Paren t1, Paren t2 -> + eq_term t1 t2 + | Requires (t1, s1), Requires (t2, s2) -> + eq_term t1 t2 && + eq_option ( = ) s1 s2 + | Ensures (t1, s1), Ensures (t2, s2) -> + eq_term t1 t2 && + eq_option ( = ) s1 s2 + | LexList ts1, LexList ts2 -> + eq_list eq_term ts1 ts2 + | WFOrder (t1, t2), WFOrder (t3, t4) -> + eq_term t1 t3 && + eq_term t2 t4 + | Decreases (t1, s1), Decreases (t2, s2) -> + eq_term t1 t2 && + eq_option ( = ) s1 s2 + | Labeled (t1, s1, b1), Labeled (t2, s2, b2) -> + eq_term t1 t2 && + s1 = s2 && + b1 = b2 + | Discrim l1, Discrim l2 -> + Ident.lid_equals l1 l2 + | Attributes ts1, Attributes ts2 -> + eq_list eq_term ts1 ts2 + | Antiquote t1, Antiquote t2 -> + eq_term t1 t2 + | Quote (t1, k1), Quote (t2, k2) -> + eq_term t1 t2 && + k1 = k2 + | VQuote t1, VQuote t2 -> + eq_term t1 t2 + | CalcProof (t1, t2, cs1), CalcProof (t3, t4, cs2) -> + eq_term t1 t3 && + eq_term t2 t4 && + eq_list eq_calc_step cs1 cs2 + | IntroForall (bs1, t1, t2), IntroForall (bs2, t3, t4) -> + eq_list eq_binder bs1 bs2 && + eq_term t1 t3 && + eq_term t2 t4 + | IntroExists (bs1, t1, ts1, t2), IntroExists (bs2, t3, ts2, t4) -> + eq_list eq_binder bs1 bs2 && + eq_term t1 t3 && + eq_list eq_term ts1 ts2 && + eq_term t2 t4 + | IntroImplies (t1, t2, b1, t3), IntroImplies (t4, t5, b2, t6) -> + eq_term t1 t4 && + eq_term t2 t5 && + eq_binder b1 b2 && + eq_term t3 t6 + | IntroOr (b1, t1, t2, t3), IntroOr (b2, t4, t5, t6) -> + b1 = b2 && + eq_term t1 t4 && + eq_term t2 t5 && + eq_term t3 t6 + | IntroAnd (t1, t2, t3, t4), IntroAnd (t5, t6, t7, t8) -> + eq_term t1 t5 && + eq_term t2 t6 && + eq_term t3 t7 && + eq_term t4 t8 + | ElimForall (bs1, t1, ts1), ElimForall (bs2, t2, ts2) -> + eq_list eq_binder bs1 bs2 && + eq_term t1 t2 && + eq_list eq_term ts1 ts2 + | ElimExists (bs1, t1, t2, b1, t3), ElimExists (bs2, t4, t5, b2, t6) -> + eq_list eq_binder bs1 bs2 && + eq_term t1 t4 && + eq_term t2 t5 && + eq_binder b1 b2 && + eq_term t3 t6 + | ElimImplies (t1, t2, t3), ElimImplies (t4, t5, t6) -> + eq_term t1 t4 && + eq_term t2 t5 && + eq_term t3 t6 + | ElimOr (t1, t2, t3, b1, t4, b2, t5), ElimOr (t6, t7, t8, b3, t9, b4, t10) -> + eq_term t1 t6 && + eq_term t2 t7 && + eq_term t3 t8 && + eq_binder b1 b3 && + eq_term t4 t9 && + eq_binder b2 b4 && + eq_term t5 t10 + | ElimAnd (t1, t2, t3, b1, b2, t4), ElimAnd (t5, t6, t7, b3, b4, t8) -> + eq_term t1 t5 && + eq_term t2 t6 && + eq_term t3 t7 && + eq_binder b1 b3 && + eq_binder b2 b4 && + eq_term t4 t8 + | _ -> false + +and eq_calc_step (CalcStep (t1, t2, t3)) (CalcStep (t4, t5, t6)) = + eq_term t1 t4 && + eq_term t2 t5 && + eq_term t3 t6 + +and eq_binder (b1 b2:binder) = eq_binder' b1.b b2.b + +and eq_binder' (b1 b2:binder') = + match b1, b2 with + | Variable i1, Variable i2 -> eq_ident i1 i2 + | TVariable i1, TVariable i2 -> eq_ident i1 i2 + | Annotated (i1, t1), Annotated (i2, t2) -> + eq_ident i1 i2 && + eq_term t1 t2 + | TAnnotated (i1, t1), TAnnotated (i2, t2) -> + eq_ident i1 i2 && + eq_term t1 t2 + | NoName t1, NoName t2 -> + eq_term t1 t2 + | _ -> false + +and eq_match_returns_annotation (i1, t1, b1) (i2, t2, b2) = + eq_option eq_ident i1 i2 && + eq_term t1 t2 && + b1 = b2 + +and eq_branch (p1, o1, t1) (p2, o2, t2) = + eq_pattern p1 p2 && + eq_option eq_term o1 o2 && + eq_term t1 t2 + + +let eq_tycon_record (t1 t2: tycon_record) = + eq_list (fun (i1, a1, a2, t1) (i2, a3, a4, t2) -> + eq_ident i1 i2 && + eq_aqual a1 a3 && + eq_list eq_term a2 a4 && + eq_term t1 t2) t1 t2 + +let eq_constructor_payload (t1 t2: constructor_payload) = + match t1, t2 with + | VpOfNotation t1, VpOfNotation t2 -> eq_term t1 t2 + | VpArbitrary t1, VpArbitrary t2 -> eq_term t1 t2 + | VpRecord (r1, k1), VpRecord (r2, k2) -> + eq_tycon_record r1 r2 && + eq_option eq_term k1 k2 + | _ -> false + +let eq_tycon (t1 t2: tycon) = + match t1, t2 with + | TyconAbstract (i1, bs1, k1), TyconAbstract (i2, bs2, k2) -> + eq_ident i1 i2 && + eq_list eq_binder bs1 bs2 && + eq_option eq_term k1 k2 + | TyconAbbrev (i1, bs1, k1, t1), TyconAbbrev (i2, bs2, k2, t2) -> + eq_ident i1 i2 && + eq_list eq_binder bs1 bs2 && + eq_option eq_term k1 k2 && + eq_term t1 t2 + | TyconRecord (i1, bs1, k1, a1, r1), TyconRecord (i2, bs2, k2, a2, r2) -> + eq_ident i1 i2 && + eq_list eq_binder bs1 bs2 && + eq_option eq_term k1 k2 && + eq_list eq_term a1 a2 && + eq_tycon_record r1 r2 + | TyconVariant (i1, bs1, k1, cs1), TyconVariant (i2, bs2, k2, cs2) -> + eq_ident i1 i2 && + eq_list eq_binder bs1 bs2 && + eq_option eq_term k1 k2 && + eq_list (fun (i1, o1, a1) (i2, o2, a2) -> + eq_ident i1 i2 && + eq_option eq_constructor_payload o1 o2 && + eq_list eq_term a1 a2) cs1 cs2 + | _ -> false + +let eq_lid = Ident.lid_equals + +let eq_lift (t1 t2: lift) = + eq_lid t1.msource t2.msource && + eq_lid t1.mdest t2.mdest && + (match t1.lift_op, t2.lift_op with + | NonReifiableLift t1, NonReifiableLift t2 -> eq_term t1 t2 + | ReifiableLift (t1, t2), ReifiableLift (t3, t4) -> + eq_term t1 t3 && + eq_term t2 t4 + | LiftForFree t1, LiftForFree t2 -> eq_term t1 t2 + | _ -> false) + + +let eq_pragma (t1 t2: pragma) = + match t1, t2 with + | SetOptions s1, SetOptions s2 -> s1 = s2 + | ResetOptions s1, ResetOptions s2 -> eq_option (fun s1 s2 -> s1 = s2) s1 s2 + | PushOptions s1, PushOptions s2 -> eq_option (fun s1 s2 -> s1 = s2) s1 s2 + | PopOptions, PopOptions -> true + | RestartSolver, RestartSolver -> true + | PrintEffectsGraph, PrintEffectsGraph -> true + | _ -> false + + +let eq_qualifier (t1 t2: qualifier) = + match t1, t2 with + | Private, Private -> true + | Noeq, Noeq -> true + | Unopteq, Unopteq -> true + | Assumption, Assumption -> true + | DefaultEffect, DefaultEffect -> true + | TotalEffect, TotalEffect -> true + | Effect_qual, Effect_qual -> true + | New, New -> true + | Inline, Inline -> true + | Visible, Visible -> true + | Unfold_for_unification_and_vcgen, Unfold_for_unification_and_vcgen -> true + | Inline_for_extraction, Inline_for_extraction -> true + | Irreducible, Irreducible -> true + | NoExtract, NoExtract -> true + | Reifiable, Reifiable -> true + | Reflectable, Reflectable -> true + | Opaque, Opaque -> true + | Logic, Logic -> true + | _ -> false + +let eq_qualifiers (t1 t2: qualifiers) = + eq_list eq_qualifier t1 t2 + +let rec eq_decl' (d1 d2:decl') : bool = + //generate the cases of this comparison starting with TopLevelModule + match d1, d2 with + | TopLevelModule lid1, TopLevelModule lid2 -> + eq_lid lid1 lid2 + | Open lid1, Open lid2 -> + eq_lid lid1 lid2 + | Friend lid1, Friend lid2 -> + eq_lid lid1 lid2 + | Include lid1, Include lid2 -> + eq_lid lid1 lid2 + | ModuleAbbrev (i1, lid1), ModuleAbbrev (i2, lid2) -> + eq_ident i1 i2 && + eq_lid lid1 lid2 + | TopLevelLet (lq1, pats1), TopLevelLet (lq2, pats2) -> + lq1=lq2 && + eq_list (fun (p1, t1) (p2, t2) -> eq_pattern p1 p2 && eq_term t1 t2) pats1 pats2 + | Tycon (b1, b2, tcs1), Tycon (b3, b4, tcs2) -> + b1 = b3 && + b2 = b4 && + eq_list eq_tycon tcs1 tcs2 + | Val (i1, t1), Val (i2, t2) -> + eq_ident i1 i2 && + eq_term t1 t2 + | Exception (i1, t1), Exception (i2, t2) -> + eq_ident i1 i2 && + eq_option eq_term t1 t2 + | NewEffect ed1, NewEffect ed2 -> + eq_effect_decl ed1 ed2 + | LayeredEffect ed1, LayeredEffect ed2 -> + eq_effect_decl ed1 ed2 + | SubEffect l1, SubEffect l2 -> + eq_lift l1 l2 + | Polymonadic_bind (lid1, lid2, lid3, t1), Polymonadic_bind (lid4, lid5, lid6, t2) -> + eq_lid lid1 lid4 && + eq_lid lid2 lid5 && + eq_lid lid3 lid6 && + eq_term t1 t2 + | Polymonadic_subcomp (lid1, lid2, t1), Polymonadic_subcomp (lid3, lid4, t2) -> + eq_lid lid1 lid3 && + eq_lid lid2 lid4 && + eq_term t1 t2 + | Pragma p1, Pragma p2 -> + eq_pragma p1 p2 + | Assume (i1, t1), Assume (i2, t2) -> + eq_ident i1 i2 && + eq_term t1 t2 + | Splice (is1, t1), Splice (is2, t2) -> + eq_list eq_ident is1 is2 && + eq_term t1 t2 + | _ -> false + +and eq_effect_decl (t1 t2: effect_decl) = + match t1, t2 with + | DefineEffect (i1, bs1, t1, ds1), DefineEffect (i2, bs2, t2, ds2) -> + eq_ident i1 i2 && + eq_list eq_binder bs1 bs2 && + eq_term t1 t2 && + eq_list eq_decl ds1 ds2 + | RedefineEffect (i1, bs1, t1), RedefineEffect (i2, bs2, t2) -> + eq_ident i1 i2 && + eq_list eq_binder bs1 bs2 && + eq_term t1 t2 + | _ -> false + +and eq_decl (d1 d2:decl) : bool = + eq_decl' d1.d d2.d && + eq_list eq_qualifier d1.quals d2.quals && + eq_list eq_term d1.attrs d2.attrs diff --git a/src/parser/FStar.Parser.AST.Comparison.fsti b/src/parser/FStar.Parser.AST.Comparison.fsti new file mode 100755 index 00000000000..e87dee1a846 --- /dev/null +++ b/src/parser/FStar.Parser.AST.Comparison.fsti @@ -0,0 +1,24 @@ +(* + Copyright 2023 Microsoft Research + + Licensed under the Apache License, Version 2.0 (the "License"); + you may not use this file except in compliance with the License. + You may obtain a copy of the License at + + http://www.apache.org/licenses/LICENSE-2.0 + + Unless required by applicable law or agreed to in writing, software + distributed under the License is distributed on an "AS IS" BASIS, + WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or implied. + See the License for the specific language governing permissions and + limitations under the License. + + Authors: N. Swamy and Copilot +*) +module FStar.Parser.AST.Comparison +open FStar.Pervasives +open FStar.Compiler.Effect +open FStar.Compiler.List +open FStar.Parser.AST + +val eq_decl (d1 d2:decl) : bool From 23d4113c6ba75ce95e67bdab7be80f9288b0cc75 Mon Sep 17 00:00:00 2001 From: Nikhil Swamy Date: Fri, 10 Mar 2023 13:16:00 -0800 Subject: [PATCH 04/48] Generated files for incremental parser tests --- .../fstar-tests/generated/FStar_Tests_Pars.ml | 30 +++++++++++++++++-- .../fstar-tests/generated/FStar_Tests_Test.ml | 5 ++-- 2 files changed, 31 insertions(+), 4 deletions(-) diff --git a/ocaml/fstar-tests/generated/FStar_Tests_Pars.ml b/ocaml/fstar-tests/generated/FStar_Tests_Pars.ml index 5d794e5141d..334b8a7b2a8 100644 --- a/ocaml/fstar-tests/generated/FStar_Tests_Pars.ml +++ b/ocaml/fstar-tests/generated/FStar_Tests_Pars.ml @@ -564,7 +564,8 @@ let (pars_and_tc_fragment : Prims.string -> unit) = let uu___3 = let uu___4 = FStar_Compiler_Effect.op_Bang test_mod_ref in - FStar_Universal.tc_one_fragment uu___4 tcenv frag in + FStar_Universal.tc_one_fragment uu___4 tcenv + (frag, FStar_Pervasives_Native.None) in (match uu___3 with | (test_mod', tcenv') -> (FStar_Compiler_Effect.op_Colon_Equals @@ -621,4 +622,29 @@ let (test_hashes : unit -> unit) = if n = Prims.int_zero then () else (test_one_hash n; aux (n - Prims.int_one)) in - aux (Prims.of_int (100)); FStar_Options.init ()) \ No newline at end of file + aux (Prims.of_int (100)); FStar_Options.init ()) +let (parse_incremental_decls : unit -> unit) = + fun uu___ -> + let source = + "module Demo\nlet f x = match x with | Some x -> true | None -> false\nlet test y = if Some? y then f y else true\nlet some junk )(" in + let input = + FStar_Parser_ParseIt.Incremental + { + FStar_Parser_ParseIt.frag_fname = "Demo.fst"; + FStar_Parser_ParseIt.frag_text = source; + FStar_Parser_ParseIt.frag_line = Prims.int_zero; + FStar_Parser_ParseIt.frag_col = Prims.int_zero + } in + let uu___1 = FStar_Parser_ParseIt.parse input in + match uu___1 with + | FStar_Parser_ParseIt.IncrementalFragment + (d0::d1::d2::[], uu___2, FStar_Pervasives_Native.Some + (uu___3, uu___4, rng)) + when + let p = FStar_Compiler_Range.start_of_range rng in + (let uu___5 = FStar_Compiler_Range.line_of_pos p in + uu___5 = (Prims.of_int (3))) && + (let uu___5 = FStar_Compiler_Range.col_of_pos p in + uu___5 = (Prims.of_int (15))) + -> () + | uu___2 -> failwith "Incremental parsing failed" \ No newline at end of file diff --git a/ocaml/fstar-tests/generated/FStar_Tests_Test.ml b/ocaml/fstar-tests/generated/FStar_Tests_Test.ml index 2b2dcf9f2c9..08c1ce2e28e 100644 --- a/ocaml/fstar-tests/generated/FStar_Tests_Test.ml +++ b/ocaml/fstar-tests/generated/FStar_Tests_Test.ml @@ -10,9 +10,10 @@ let main : 'uuuuu 'uuuuu1 . 'uuuuu -> 'uuuuu1 = (let uu___4 = FStar_Tests_Pars.init () in FStar_Compiler_Effect.op_Bar_Greater uu___4 (fun uu___5 -> ())); + FStar_Tests_Pars.parse_incremental_decls (); FStar_Tests_Norm.run_all (); - (let uu___6 = FStar_Tests_Unif.run_all () in - if uu___6 + (let uu___7 = FStar_Tests_Unif.run_all () in + if uu___7 then () else FStar_Compiler_Effect.exit Prims.int_one); FStar_Compiler_Effect.exit Prims.int_zero)) () From 0c5212600be63915d8970d227b199a78259844c0 Mon Sep 17 00:00:00 2001 From: Nikhil Swamy Date: Fri, 10 Mar 2023 13:21:59 -0800 Subject: [PATCH 05/48] propagating AST syntax change --- src/syntax/FStar.Syntax.Resugar.fst | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/syntax/FStar.Syntax.Resugar.fst b/src/syntax/FStar.Syntax.Resugar.fst index d3a2b365231..ed5cf3fa66e 100644 --- a/src/syntax/FStar.Syntax.Resugar.fst +++ b/src/syntax/FStar.Syntax.Resugar.fst @@ -1425,7 +1425,7 @@ let resugar_sigelt' env se : option A.decl = | None, Some t -> A.LiftForFree t | _ -> failwith "Should not happen hopefully" in - Some (decl'_to_decl se (A.SubEffect({msource=src; mdest=dst; lift_op=op}))) + Some (decl'_to_decl se (A.SubEffect({msource=src; mdest=dst; lift_op=op; braced=false}))) | Sig_effect_abbrev (lid, vs, bs, c, flags) -> let bs, c = SS.open_comp bs c in From 539f8c91f298f2df4ea883c4fe9e665598d9976b Mon Sep 17 00:00:00 2001 From: Nikhil Swamy Date: Fri, 10 Mar 2023 14:00:28 -0800 Subject: [PATCH 06/48] Supporting a new full-buffer message in the IDE protocol, used by fstar-vscode-assistant --- .../generated/FStar_Interactive_Ide.ml | 1495 +- .../generated/FStar_Interactive_Ide_Types.ml | 379 + .../FStar_Interactive_Incremental.ml | 256 + .../generated/FStar_Interactive_JsonHelper.ml | 113 - .../generated/FStar_Interactive_Legacy.ml | 4 +- .../generated/FStar_Interactive_Lsp.ml | 52 +- .../generated/FStar_Interactive_PushHelper.ml | 367 +- .../FStar_Interactive_QueryHelper.ml | 8 +- .../generated/FStar_Interactive_ReplState.ml | 184 + ocaml/fstar-lib/generated/FStar_Parser_AST.ml | 27 +- ocaml/fstar-lib/generated/FStar_Parser_Dep.ml | 9 +- .../generated/FStar_Parser_Driver.ml | 2 + .../fstar-lib/generated/FStar_Parser_Parse.ml | 12167 ++++++++-------- .../generated/FStar_Syntax_Resugar.ml | 3 +- ocaml/fstar-lib/generated/FStar_Universal.ml | 193 +- .../fstar-tests/generated/FStar_Tests_Pars.ml | 2 +- src/Makefile.boot | 4 +- src/fstar/FStar.Interactive.Ide.Types.fst | 217 + src/fstar/FStar.Interactive.Ide.fst | 308 +- src/fstar/FStar.Interactive.Incremental.fst | 217 + src/fstar/FStar.Interactive.JsonHelper.fst | 1 + src/fstar/FStar.Interactive.JsonHelper.fsti | 38 +- src/fstar/FStar.Interactive.Legacy.fst | 2 +- src/fstar/FStar.Interactive.Lsp.fst | 1 + src/fstar/FStar.Interactive.PushHelper.fst | 3 +- src/fstar/FStar.Interactive.PushHelper.fsti | 1 + src/fstar/FStar.Interactive.QueryHelper.fst | 1 + src/fstar/FStar.Interactive.QueryHelper.fsti | 4 +- src/fstar/FStar.Interactive.ReplState.fst | 121 + src/fstar/FStar.Universal.fst | 79 +- src/fstar/FStar.Universal.fsti | 2 +- src/tests/FStar.Tests.Pars.fst | 2 +- 32 files changed, 8697 insertions(+), 7565 deletions(-) create mode 100755 ocaml/fstar-lib/generated/FStar_Interactive_Ide_Types.ml create mode 100755 ocaml/fstar-lib/generated/FStar_Interactive_Incremental.ml create mode 100755 ocaml/fstar-lib/generated/FStar_Interactive_ReplState.ml create mode 100755 src/fstar/FStar.Interactive.Ide.Types.fst create mode 100755 src/fstar/FStar.Interactive.Incremental.fst create mode 100755 src/fstar/FStar.Interactive.ReplState.fst diff --git a/ocaml/fstar-lib/generated/FStar_Interactive_Ide.ml b/ocaml/fstar-lib/generated/FStar_Interactive_Ide.ml index 8ec462707a2..43159941034 100644 --- a/ocaml/fstar-lib/generated/FStar_Interactive_Ide.ml +++ b/ocaml/fstar-lib/generated/FStar_Interactive_Ide.ml @@ -56,64 +56,12 @@ let with_captured_errors : fun f -> let uu___ = FStar_Options.trace_error () in if uu___ then f env else with_captured_errors' env sigint_handler f -let (t0 : FStar_Compiler_Util.time) = FStar_Compiler_Util.now () -let (dummy_tf_of_fname : - Prims.string -> FStar_Interactive_JsonHelper.timed_fname) = - fun fname -> - { - FStar_Interactive_JsonHelper.tf_fname = fname; - FStar_Interactive_JsonHelper.tf_modtime = t0 - } -let (string_of_timed_fname : - FStar_Interactive_JsonHelper.timed_fname -> Prims.string) = - fun uu___ -> - match uu___ with - | { FStar_Interactive_JsonHelper.tf_fname = fname; - FStar_Interactive_JsonHelper.tf_modtime = modtime;_} -> - if modtime = t0 - then FStar_Compiler_Util.format1 "{ %s }" fname - else - (let uu___2 = FStar_Compiler_Util.string_of_time modtime in - FStar_Compiler_Util.format2 "{ %s; %s }" fname uu___2) -type push_query = - { - push_kind: FStar_Interactive_PushHelper.push_kind ; - push_code: Prims.string ; - push_line: Prims.int ; - push_column: Prims.int ; - push_peek_only: Prims.bool } -let (__proj__Mkpush_query__item__push_kind : - push_query -> FStar_Interactive_PushHelper.push_kind) = - fun projectee -> - match projectee with - | { push_kind; push_code; push_line; push_column; push_peek_only;_} -> - push_kind -let (__proj__Mkpush_query__item__push_code : push_query -> Prims.string) = - fun projectee -> - match projectee with - | { push_kind; push_code; push_line; push_column; push_peek_only;_} -> - push_code -let (__proj__Mkpush_query__item__push_line : push_query -> Prims.int) = - fun projectee -> - match projectee with - | { push_kind; push_code; push_line; push_column; push_peek_only;_} -> - push_line -let (__proj__Mkpush_query__item__push_column : push_query -> Prims.int) = - fun projectee -> - match projectee with - | { push_kind; push_code; push_line; push_column; push_peek_only;_} -> - push_column -let (__proj__Mkpush_query__item__push_peek_only : push_query -> Prims.bool) = - fun projectee -> - match projectee with - | { push_kind; push_code; push_line; push_column; push_peek_only;_} -> - push_peek_only type env_t = FStar_TypeChecker_Env.env let (repl_current_qid : Prims.string FStar_Pervasives_Native.option FStar_Compiler_Effect.ref) = FStar_Compiler_Util.mk_ref FStar_Pervasives_Native.None let (nothing_left_to_pop : - FStar_Interactive_JsonHelper.repl_state -> Prims.bool) = + FStar_Interactive_ReplState.repl_state -> Prims.bool) = fun st -> let uu___ = let uu___1 = @@ -121,31 +69,13 @@ let (nothing_left_to_pop : FStar_Compiler_List.length uu___1 in uu___ = (FStar_Compiler_List.length - st.FStar_Interactive_JsonHelper.repl_deps_stack) -let (string_of_repl_task : - FStar_Interactive_JsonHelper.repl_task -> Prims.string) = - fun uu___ -> - match uu___ with - | FStar_Interactive_JsonHelper.LDInterleaved (intf, impl) -> - let uu___1 = string_of_timed_fname intf in - let uu___2 = string_of_timed_fname impl in - FStar_Compiler_Util.format2 "LDInterleaved (%s, %s)" uu___1 uu___2 - | FStar_Interactive_JsonHelper.LDSingle intf_or_impl -> - let uu___1 = string_of_timed_fname intf_or_impl in - FStar_Compiler_Util.format1 "LDSingle %s" uu___1 - | FStar_Interactive_JsonHelper.LDInterfaceOfCurrentFile intf -> - let uu___1 = string_of_timed_fname intf in - FStar_Compiler_Util.format1 "LDInterfaceOfCurrentFile %s" uu___1 - | FStar_Interactive_JsonHelper.PushFragment frag -> - FStar_Compiler_Util.format1 "PushFragment { code = %s }" - frag.FStar_Parser_ParseIt.frag_text - | FStar_Interactive_JsonHelper.Noop -> "Noop {}" + st.FStar_Interactive_ReplState.repl_deps_stack) let (run_repl_transaction : - FStar_Interactive_JsonHelper.repl_state -> + FStar_Interactive_ReplState.repl_state -> FStar_Interactive_PushHelper.push_kind -> Prims.bool -> - FStar_Interactive_JsonHelper.repl_task -> - (Prims.bool * FStar_Interactive_JsonHelper.repl_state)) + FStar_Interactive_ReplState.repl_task -> + (Prims.bool * FStar_Interactive_ReplState.repl_state)) = fun st -> fun push_kind -> @@ -156,7 +86,7 @@ let (run_repl_transaction : push_kind task st in let uu___ = FStar_Interactive_PushHelper.track_name_changes - st1.FStar_Interactive_JsonHelper.repl_env in + st1.FStar_Interactive_ReplState.repl_env in match uu___ with | (env, finish_name_tracking) -> let check_success uu___1 = @@ -169,7 +99,7 @@ let (run_repl_transaction : (fun env1 -> let uu___3 = FStar_Interactive_PushHelper.run_repl_task - st1.FStar_Interactive_JsonHelper.repl_curmod env1 + st1.FStar_Interactive_ReplState.repl_curmod env1 task in FStar_Compiler_Effect.op_Less_Bar (fun uu___4 -> FStar_Pervasives_Native.Some uu___4) @@ -178,7 +108,7 @@ let (run_repl_transaction : | FStar_Pervasives_Native.Some (curmod, env1) when check_success () -> (curmod, env1, true) | uu___3 -> - ((st1.FStar_Interactive_JsonHelper.repl_curmod), env, + ((st1.FStar_Interactive_ReplState.repl_curmod), env, false) in (match uu___1 with | (curmod, env1, success) -> @@ -190,22 +120,21 @@ let (run_repl_transaction : then let st3 = { - FStar_Interactive_JsonHelper.repl_line = - (st1.FStar_Interactive_JsonHelper.repl_line); - FStar_Interactive_JsonHelper.repl_column = - (st1.FStar_Interactive_JsonHelper.repl_column); - FStar_Interactive_JsonHelper.repl_fname = - (st1.FStar_Interactive_JsonHelper.repl_fname); - FStar_Interactive_JsonHelper.repl_deps_stack - = - (st1.FStar_Interactive_JsonHelper.repl_deps_stack); - FStar_Interactive_JsonHelper.repl_curmod = + FStar_Interactive_ReplState.repl_line = + (st1.FStar_Interactive_ReplState.repl_line); + FStar_Interactive_ReplState.repl_column = + (st1.FStar_Interactive_ReplState.repl_column); + FStar_Interactive_ReplState.repl_fname = + (st1.FStar_Interactive_ReplState.repl_fname); + FStar_Interactive_ReplState.repl_deps_stack = + (st1.FStar_Interactive_ReplState.repl_deps_stack); + FStar_Interactive_ReplState.repl_curmod = curmod; - FStar_Interactive_JsonHelper.repl_env = env2; - FStar_Interactive_JsonHelper.repl_stdin = - (st1.FStar_Interactive_JsonHelper.repl_stdin); - FStar_Interactive_JsonHelper.repl_names = - (st1.FStar_Interactive_JsonHelper.repl_names) + FStar_Interactive_ReplState.repl_env = env2; + FStar_Interactive_ReplState.repl_stdin = + (st1.FStar_Interactive_ReplState.repl_stdin); + FStar_Interactive_ReplState.repl_names = + (st1.FStar_Interactive_ReplState.repl_names) } in FStar_Interactive_PushHelper.commit_name_tracking st3 name_events @@ -214,11 +143,11 @@ let (run_repl_transaction : "run_repl_transaction" st1 in (success, st2))) let (run_repl_ld_transactions : - FStar_Interactive_JsonHelper.repl_state -> - FStar_Interactive_JsonHelper.repl_task Prims.list -> - (FStar_Interactive_JsonHelper.repl_task -> unit) -> - (FStar_Interactive_JsonHelper.repl_state, - FStar_Interactive_JsonHelper.repl_state) FStar_Pervasives.either) + FStar_Interactive_ReplState.repl_state -> + FStar_Interactive_ReplState.repl_task Prims.list -> + (FStar_Interactive_ReplState.repl_task -> unit) -> + (FStar_Interactive_ReplState.repl_state, + FStar_Interactive_ReplState.repl_state) FStar_Pervasives.either) = fun st -> fun tasks -> @@ -227,7 +156,7 @@ let (run_repl_ld_transactions : let uu___ = FStar_Options.debug_any () in if uu___ then - let uu___1 = string_of_repl_task task in + let uu___1 = FStar_Interactive_ReplState.string_of_repl_task task in FStar_Compiler_Util.print2 "%s %s" verb uu___1 else () in let rec revert_many st1 uu___ = @@ -252,27 +181,27 @@ let (run_repl_ld_transactions : "run_repl_ls_transactions" st1 in let dep_graph = FStar_TypeChecker_Env.dep_graph - st1.FStar_Interactive_JsonHelper.repl_env in + st1.FStar_Interactive_ReplState.repl_env in let st'1 = let uu___3 = FStar_TypeChecker_Env.set_dep_graph - st'.FStar_Interactive_JsonHelper.repl_env dep_graph in + st'.FStar_Interactive_ReplState.repl_env dep_graph in { - FStar_Interactive_JsonHelper.repl_line = - (st'.FStar_Interactive_JsonHelper.repl_line); - FStar_Interactive_JsonHelper.repl_column = - (st'.FStar_Interactive_JsonHelper.repl_column); - FStar_Interactive_JsonHelper.repl_fname = - (st'.FStar_Interactive_JsonHelper.repl_fname); - FStar_Interactive_JsonHelper.repl_deps_stack = - (st'.FStar_Interactive_JsonHelper.repl_deps_stack); - FStar_Interactive_JsonHelper.repl_curmod = - (st'.FStar_Interactive_JsonHelper.repl_curmod); - FStar_Interactive_JsonHelper.repl_env = uu___3; - FStar_Interactive_JsonHelper.repl_stdin = - (st'.FStar_Interactive_JsonHelper.repl_stdin); - FStar_Interactive_JsonHelper.repl_names = - (st'.FStar_Interactive_JsonHelper.repl_names) + FStar_Interactive_ReplState.repl_line = + (st'.FStar_Interactive_ReplState.repl_line); + FStar_Interactive_ReplState.repl_column = + (st'.FStar_Interactive_ReplState.repl_column); + FStar_Interactive_ReplState.repl_fname = + (st'.FStar_Interactive_ReplState.repl_fname); + FStar_Interactive_ReplState.repl_deps_stack = + (st'.FStar_Interactive_ReplState.repl_deps_stack); + FStar_Interactive_ReplState.repl_curmod = + (st'.FStar_Interactive_ReplState.repl_curmod); + FStar_Interactive_ReplState.repl_env = uu___3; + FStar_Interactive_ReplState.repl_stdin = + (st'.FStar_Interactive_ReplState.repl_stdin); + FStar_Interactive_ReplState.repl_names = + (st'.FStar_Interactive_ReplState.repl_names) } in revert_many st'1 entries)) in let rec aux st1 tasks1 previous = @@ -302,22 +231,22 @@ let (run_repl_ld_transactions : FStar_Compiler_Effect.op_Bang FStar_Interactive_PushHelper.repl_stack in { - FStar_Interactive_JsonHelper.repl_line = - (st2.FStar_Interactive_JsonHelper.repl_line); - FStar_Interactive_JsonHelper.repl_column = - (st2.FStar_Interactive_JsonHelper.repl_column); - FStar_Interactive_JsonHelper.repl_fname = - (st2.FStar_Interactive_JsonHelper.repl_fname); - FStar_Interactive_JsonHelper.repl_deps_stack = + FStar_Interactive_ReplState.repl_line = + (st2.FStar_Interactive_ReplState.repl_line); + FStar_Interactive_ReplState.repl_column = + (st2.FStar_Interactive_ReplState.repl_column); + FStar_Interactive_ReplState.repl_fname = + (st2.FStar_Interactive_ReplState.repl_fname); + FStar_Interactive_ReplState.repl_deps_stack = uu___5; - FStar_Interactive_JsonHelper.repl_curmod = - (st2.FStar_Interactive_JsonHelper.repl_curmod); - FStar_Interactive_JsonHelper.repl_env = - (st2.FStar_Interactive_JsonHelper.repl_env); - FStar_Interactive_JsonHelper.repl_stdin = - (st2.FStar_Interactive_JsonHelper.repl_stdin); - FStar_Interactive_JsonHelper.repl_names = - (st2.FStar_Interactive_JsonHelper.repl_names) + FStar_Interactive_ReplState.repl_curmod = + (st2.FStar_Interactive_ReplState.repl_curmod); + FStar_Interactive_ReplState.repl_env = + (st2.FStar_Interactive_ReplState.repl_env); + FStar_Interactive_ReplState.repl_stdin = + (st2.FStar_Interactive_ReplState.repl_stdin); + FStar_Interactive_ReplState.repl_names = + (st2.FStar_Interactive_ReplState.repl_names) } in aux uu___4 tasks2 [] else FStar_Pervasives.Inr st2)) @@ -331,240 +260,12 @@ let (run_repl_ld_transactions : let uu___ = revert_many st1 previous1 in aux uu___ tasks2 [] in aux st tasks (FStar_Compiler_List.rev - st.FStar_Interactive_JsonHelper.repl_deps_stack) -let (js_pushkind : - FStar_Compiler_Util.json -> FStar_Interactive_PushHelper.push_kind) = - fun s -> - let uu___ = FStar_Interactive_JsonHelper.js_str s in - match uu___ with - | "syntax" -> FStar_Interactive_PushHelper.SyntaxCheck - | "lax" -> FStar_Interactive_PushHelper.LaxCheck - | "full" -> FStar_Interactive_PushHelper.FullCheck - | uu___1 -> FStar_Interactive_JsonHelper.js_fail "push_kind" s -let (js_reductionrule : - FStar_Compiler_Util.json -> FStar_TypeChecker_Env.step) = - fun s -> - let uu___ = FStar_Interactive_JsonHelper.js_str s in - match uu___ with - | "beta" -> FStar_TypeChecker_Env.Beta - | "delta" -> - FStar_TypeChecker_Env.UnfoldUntil FStar_Syntax_Syntax.delta_constant - | "iota" -> FStar_TypeChecker_Env.Iota - | "zeta" -> FStar_TypeChecker_Env.Zeta - | "reify" -> FStar_TypeChecker_Env.Reify - | "pure-subterms" -> FStar_TypeChecker_Env.PureSubtermsWithinComputations - | uu___1 -> FStar_Interactive_JsonHelper.js_fail "reduction rule" s -type completion_context = - | CKCode - | CKOption of Prims.bool - | CKModuleOrNamespace of (Prims.bool * Prims.bool) -let (uu___is_CKCode : completion_context -> Prims.bool) = - fun projectee -> match projectee with | CKCode -> true | uu___ -> false -let (uu___is_CKOption : completion_context -> Prims.bool) = - fun projectee -> - match projectee with | CKOption _0 -> true | uu___ -> false -let (__proj__CKOption__item___0 : completion_context -> Prims.bool) = - fun projectee -> match projectee with | CKOption _0 -> _0 -let (uu___is_CKModuleOrNamespace : completion_context -> Prims.bool) = - fun projectee -> - match projectee with | CKModuleOrNamespace _0 -> true | uu___ -> false -let (__proj__CKModuleOrNamespace__item___0 : - completion_context -> (Prims.bool * Prims.bool)) = - fun projectee -> match projectee with | CKModuleOrNamespace _0 -> _0 -let (js_optional_completion_context : - FStar_Compiler_Util.json FStar_Pervasives_Native.option -> - completion_context) - = - fun k -> - match k with - | FStar_Pervasives_Native.None -> CKCode - | FStar_Pervasives_Native.Some k1 -> - let uu___ = FStar_Interactive_JsonHelper.js_str k1 in - (match uu___ with - | "symbol" -> CKCode - | "code" -> CKCode - | "set-options" -> CKOption false - | "reset-options" -> CKOption true - | "open" -> CKModuleOrNamespace (true, true) - | "let-open" -> CKModuleOrNamespace (true, true) - | "include" -> CKModuleOrNamespace (true, false) - | "module-alias" -> CKModuleOrNamespace (true, false) - | uu___1 -> - FStar_Interactive_JsonHelper.js_fail - "completion context (code, set-options, reset-options, open, let-open, include, module-alias)" - k1) -type lookup_context = - | LKSymbolOnly - | LKModule - | LKOption - | LKCode -let (uu___is_LKSymbolOnly : lookup_context -> Prims.bool) = - fun projectee -> - match projectee with | LKSymbolOnly -> true | uu___ -> false -let (uu___is_LKModule : lookup_context -> Prims.bool) = - fun projectee -> match projectee with | LKModule -> true | uu___ -> false -let (uu___is_LKOption : lookup_context -> Prims.bool) = - fun projectee -> match projectee with | LKOption -> true | uu___ -> false -let (uu___is_LKCode : lookup_context -> Prims.bool) = - fun projectee -> match projectee with | LKCode -> true | uu___ -> false -let (js_optional_lookup_context : - FStar_Compiler_Util.json FStar_Pervasives_Native.option -> lookup_context) - = - fun k -> - match k with - | FStar_Pervasives_Native.None -> LKSymbolOnly - | FStar_Pervasives_Native.Some k1 -> - let uu___ = FStar_Interactive_JsonHelper.js_str k1 in - (match uu___ with - | "symbol-only" -> LKSymbolOnly - | "code" -> LKCode - | "set-options" -> LKOption - | "reset-options" -> LKOption - | "open" -> LKModule - | "let-open" -> LKModule - | "include" -> LKModule - | "module-alias" -> LKModule - | uu___1 -> - FStar_Interactive_JsonHelper.js_fail - "lookup context (symbol-only, code, set-options, reset-options, open, let-open, include, module-alias)" - k1) -type position = (Prims.string * Prims.int * Prims.int) -type query' = - | Exit - | DescribeProtocol - | DescribeRepl - | Segment of Prims.string - | Pop - | Push of push_query - | VfsAdd of (Prims.string FStar_Pervasives_Native.option * Prims.string) - | AutoComplete of (Prims.string * completion_context) - | Lookup of (Prims.string * lookup_context * position - FStar_Pervasives_Native.option * Prims.string Prims.list) - | Compute of (Prims.string * FStar_TypeChecker_Env.step Prims.list - FStar_Pervasives_Native.option) - | Search of Prims.string - | GenericError of Prims.string - | ProtocolViolation of Prims.string -and query = { - qq: query' ; - qid: Prims.string } -let (uu___is_Exit : query' -> Prims.bool) = - fun projectee -> match projectee with | Exit -> true | uu___ -> false -let (uu___is_DescribeProtocol : query' -> Prims.bool) = - fun projectee -> - match projectee with | DescribeProtocol -> true | uu___ -> false -let (uu___is_DescribeRepl : query' -> Prims.bool) = - fun projectee -> - match projectee with | DescribeRepl -> true | uu___ -> false -let (uu___is_Segment : query' -> Prims.bool) = - fun projectee -> match projectee with | Segment _0 -> true | uu___ -> false -let (__proj__Segment__item___0 : query' -> Prims.string) = - fun projectee -> match projectee with | Segment _0 -> _0 -let (uu___is_Pop : query' -> Prims.bool) = - fun projectee -> match projectee with | Pop -> true | uu___ -> false -let (uu___is_Push : query' -> Prims.bool) = - fun projectee -> match projectee with | Push _0 -> true | uu___ -> false -let (__proj__Push__item___0 : query' -> push_query) = - fun projectee -> match projectee with | Push _0 -> _0 -let (uu___is_VfsAdd : query' -> Prims.bool) = - fun projectee -> match projectee with | VfsAdd _0 -> true | uu___ -> false -let (__proj__VfsAdd__item___0 : - query' -> (Prims.string FStar_Pervasives_Native.option * Prims.string)) = - fun projectee -> match projectee with | VfsAdd _0 -> _0 -let (uu___is_AutoComplete : query' -> Prims.bool) = - fun projectee -> - match projectee with | AutoComplete _0 -> true | uu___ -> false -let (__proj__AutoComplete__item___0 : - query' -> (Prims.string * completion_context)) = - fun projectee -> match projectee with | AutoComplete _0 -> _0 -let (uu___is_Lookup : query' -> Prims.bool) = - fun projectee -> match projectee with | Lookup _0 -> true | uu___ -> false -let (__proj__Lookup__item___0 : - query' -> - (Prims.string * lookup_context * position FStar_Pervasives_Native.option - * Prims.string Prims.list)) - = fun projectee -> match projectee with | Lookup _0 -> _0 -let (uu___is_Compute : query' -> Prims.bool) = - fun projectee -> match projectee with | Compute _0 -> true | uu___ -> false -let (__proj__Compute__item___0 : - query' -> - (Prims.string * FStar_TypeChecker_Env.step Prims.list - FStar_Pervasives_Native.option)) - = fun projectee -> match projectee with | Compute _0 -> _0 -let (uu___is_Search : query' -> Prims.bool) = - fun projectee -> match projectee with | Search _0 -> true | uu___ -> false -let (__proj__Search__item___0 : query' -> Prims.string) = - fun projectee -> match projectee with | Search _0 -> _0 -let (uu___is_GenericError : query' -> Prims.bool) = - fun projectee -> - match projectee with | GenericError _0 -> true | uu___ -> false -let (__proj__GenericError__item___0 : query' -> Prims.string) = - fun projectee -> match projectee with | GenericError _0 -> _0 -let (uu___is_ProtocolViolation : query' -> Prims.bool) = - fun projectee -> - match projectee with | ProtocolViolation _0 -> true | uu___ -> false -let (__proj__ProtocolViolation__item___0 : query' -> Prims.string) = - fun projectee -> match projectee with | ProtocolViolation _0 -> _0 -let (__proj__Mkquery__item__qq : query -> query') = - fun projectee -> match projectee with | { qq; qid;_} -> qq -let (__proj__Mkquery__item__qid : query -> Prims.string) = - fun projectee -> match projectee with | { qq; qid;_} -> qid -let (query_needs_current_module : query' -> Prims.bool) = - fun uu___ -> - match uu___ with - | Exit -> false - | DescribeProtocol -> false - | DescribeRepl -> false - | Segment uu___1 -> false - | Pop -> false - | Push - { push_kind = uu___1; push_code = uu___2; push_line = uu___3; - push_column = uu___4; push_peek_only = false;_} - -> false - | VfsAdd uu___1 -> false - | GenericError uu___1 -> false - | ProtocolViolation uu___1 -> false - | Push uu___1 -> true - | AutoComplete uu___1 -> true - | Lookup uu___1 -> true - | Compute uu___1 -> true - | Search uu___1 -> true -let (interactive_protocol_vernum : Prims.int) = (Prims.of_int (2)) -let (interactive_protocol_features : Prims.string Prims.list) = - ["autocomplete"; - "autocomplete/context"; - "compute"; - "compute/reify"; - "compute/pure-subterms"; - "describe-protocol"; - "describe-repl"; - "exit"; - "lookup"; - "lookup/context"; - "lookup/documentation"; - "lookup/definition"; - "peek"; - "pop"; - "push"; - "search"; - "segment"; - "vfs-add"; - "tactic-ranges"; - "interrupt"; - "progress"] -type query_status = - | QueryOK - | QueryNOK - | QueryViolatesProtocol -let (uu___is_QueryOK : query_status -> Prims.bool) = - fun projectee -> match projectee with | QueryOK -> true | uu___ -> false -let (uu___is_QueryNOK : query_status -> Prims.bool) = - fun projectee -> match projectee with | QueryNOK -> true | uu___ -> false -let (uu___is_QueryViolatesProtocol : query_status -> Prims.bool) = - fun projectee -> - match projectee with | QueryViolatesProtocol -> true | uu___ -> false + st.FStar_Interactive_ReplState.repl_deps_stack) let (wrap_js_failure : - Prims.string -> Prims.string -> FStar_Compiler_Util.json -> query) = + Prims.string -> + Prims.string -> + FStar_Compiler_Util.json -> FStar_Interactive_Ide_Types.query) + = fun qid -> fun expected -> fun got -> @@ -573,9 +274,13 @@ let (wrap_js_failure : let uu___2 = FStar_Interactive_JsonHelper.json_debug got in FStar_Compiler_Util.format2 "JSON decoding failed: expected %s, got %s" expected uu___2 in - ProtocolViolation uu___1 in - { qq = uu___; qid } -let (unpack_interactive_query : FStar_Compiler_Util.json -> query) = + FStar_Interactive_Ide_Types.ProtocolViolation uu___1 in + { + FStar_Interactive_Ide_Types.qq = uu___; + FStar_Interactive_Ide_Types.qid = qid + } +let (unpack_interactive_query : + FStar_Compiler_Util.json -> FStar_Interactive_Ide_Types.query) = fun json -> let assoc errloc key a = let uu___ = FStar_Interactive_JsonHelper.try_assoc key a in @@ -599,7 +304,7 @@ let (unpack_interactive_query : FStar_Compiler_Util.json -> query) = (fun uu___ -> match () with | () -> - let query1 = + let query = let uu___1 = assoc "query" "query" request in FStar_Compiler_Effect.op_Bar_Greater uu___1 FStar_Interactive_JsonHelper.js_str in @@ -615,69 +320,82 @@ let (unpack_interactive_query : FStar_Compiler_Util.json -> query) = -> FStar_Pervasives_Native.None | other -> other in let uu___1 = - match query1 with - | "exit" -> Exit - | "pop" -> Pop - | "describe-protocol" -> DescribeProtocol - | "describe-repl" -> DescribeRepl + match query with + | "exit" -> FStar_Interactive_Ide_Types.Exit + | "pop" -> FStar_Interactive_Ide_Types.Pop + | "describe-protocol" -> + FStar_Interactive_Ide_Types.DescribeProtocol + | "describe-repl" -> FStar_Interactive_Ide_Types.DescribeRepl | "segment" -> let uu___2 = let uu___3 = arg "code" in FStar_Compiler_Effect.op_Bar_Greater uu___3 FStar_Interactive_JsonHelper.js_str in - Segment uu___2 + FStar_Interactive_Ide_Types.Segment uu___2 | "peek" -> let uu___2 = let uu___3 = let uu___4 = arg "kind" in FStar_Compiler_Effect.op_Bar_Greater uu___4 - js_pushkind in + FStar_Interactive_Ide_Types.js_pushkind in let uu___4 = - let uu___5 = arg "code" in + let uu___5 = arg "line" in FStar_Compiler_Effect.op_Bar_Greater uu___5 - FStar_Interactive_JsonHelper.js_str in + FStar_Interactive_JsonHelper.js_int in let uu___5 = - let uu___6 = arg "line" in + let uu___6 = arg "column" in FStar_Compiler_Effect.op_Bar_Greater uu___6 FStar_Interactive_JsonHelper.js_int in let uu___6 = - let uu___7 = arg "column" in - FStar_Compiler_Effect.op_Bar_Greater uu___7 - FStar_Interactive_JsonHelper.js_int in + let uu___7 = + let uu___8 = arg "code" in + FStar_Compiler_Effect.op_Bar_Greater uu___8 + FStar_Interactive_JsonHelper.js_str in + FStar_Pervasives.Inl uu___7 in { - push_kind = uu___3; - push_code = uu___4; - push_line = uu___5; - push_column = uu___6; - push_peek_only = (query1 = "peek") + FStar_Interactive_Ide_Types.push_kind = uu___3; + FStar_Interactive_Ide_Types.push_line = uu___4; + FStar_Interactive_Ide_Types.push_column = uu___5; + FStar_Interactive_Ide_Types.push_peek_only = + (query = "peek"); + FStar_Interactive_Ide_Types.push_code_or_decl = uu___6 } in - Push uu___2 + FStar_Interactive_Ide_Types.Push uu___2 | "push" -> let uu___2 = let uu___3 = let uu___4 = arg "kind" in FStar_Compiler_Effect.op_Bar_Greater uu___4 - js_pushkind in + FStar_Interactive_Ide_Types.js_pushkind in let uu___4 = - let uu___5 = arg "code" in + let uu___5 = arg "line" in FStar_Compiler_Effect.op_Bar_Greater uu___5 - FStar_Interactive_JsonHelper.js_str in + FStar_Interactive_JsonHelper.js_int in let uu___5 = - let uu___6 = arg "line" in + let uu___6 = arg "column" in FStar_Compiler_Effect.op_Bar_Greater uu___6 FStar_Interactive_JsonHelper.js_int in let uu___6 = - let uu___7 = arg "column" in - FStar_Compiler_Effect.op_Bar_Greater uu___7 - FStar_Interactive_JsonHelper.js_int in + let uu___7 = + let uu___8 = arg "code" in + FStar_Compiler_Effect.op_Bar_Greater uu___8 + FStar_Interactive_JsonHelper.js_str in + FStar_Pervasives.Inl uu___7 in { - push_kind = uu___3; - push_code = uu___4; - push_line = uu___5; - push_column = uu___6; - push_peek_only = (query1 = "peek") + FStar_Interactive_Ide_Types.push_kind = uu___3; + FStar_Interactive_Ide_Types.push_line = uu___4; + FStar_Interactive_Ide_Types.push_column = uu___5; + FStar_Interactive_Ide_Types.push_peek_only = + (query = "peek"); + FStar_Interactive_Ide_Types.push_code_or_decl = uu___6 } in - Push uu___2 + FStar_Interactive_Ide_Types.Push uu___2 + | "full-buffer" -> + let uu___2 = + let uu___3 = arg "code" in + FStar_Compiler_Effect.op_Bar_Greater uu___3 + FStar_Interactive_JsonHelper.js_str in + FStar_Interactive_Ide_Types.FullBuffer uu___2 | "autocomplete" -> let uu___2 = let uu___3 = @@ -687,9 +405,9 @@ let (unpack_interactive_query : FStar_Compiler_Util.json -> query) = let uu___4 = let uu___5 = try_arg "context" in FStar_Compiler_Effect.op_Bar_Greater uu___5 - js_optional_completion_context in + FStar_Interactive_Ide_Types.js_optional_completion_context in (uu___3, uu___4) in - AutoComplete uu___2 + FStar_Interactive_Ide_Types.AutoComplete uu___2 | "lookup" -> let uu___2 = let uu___3 = @@ -699,7 +417,7 @@ let (unpack_interactive_query : FStar_Compiler_Util.json -> query) = let uu___4 = let uu___5 = try_arg "context" in FStar_Compiler_Effect.op_Bar_Greater uu___5 - js_optional_lookup_context in + FStar_Interactive_Ide_Types.js_optional_lookup_context in let uu___5 = let uu___6 = let uu___7 = try_arg "location" in @@ -730,7 +448,7 @@ let (unpack_interactive_query : FStar_Compiler_Util.json -> query) = (FStar_Interactive_JsonHelper.js_list FStar_Interactive_JsonHelper.js_str) in (uu___3, uu___4, uu___5, uu___6) in - Lookup uu___2 + FStar_Interactive_Ide_Types.Lookup uu___2 | "compute" -> let uu___2 = let uu___3 = @@ -742,15 +460,15 @@ let (unpack_interactive_query : FStar_Compiler_Util.json -> query) = FStar_Compiler_Effect.op_Bar_Greater uu___5 (FStar_Compiler_Util.map_option (FStar_Interactive_JsonHelper.js_list - js_reductionrule)) in + FStar_Interactive_Ide_Types.js_reductionrule)) in (uu___3, uu___4) in - Compute uu___2 + FStar_Interactive_Ide_Types.Compute uu___2 | "search" -> let uu___2 = let uu___3 = arg "terms" in FStar_Compiler_Effect.op_Bar_Greater uu___3 FStar_Interactive_JsonHelper.js_str in - Search uu___2 + FStar_Interactive_Ide_Types.Search uu___2 | "vfs-add" -> let uu___2 = let uu___3 = @@ -763,36 +481,55 @@ let (unpack_interactive_query : FStar_Compiler_Util.json -> query) = FStar_Compiler_Effect.op_Bar_Greater uu___5 FStar_Interactive_JsonHelper.js_str in (uu___3, uu___4) in - VfsAdd uu___2 + FStar_Interactive_Ide_Types.VfsAdd uu___2 | uu___2 -> let uu___3 = - FStar_Compiler_Util.format1 "Unknown query '%s'" query1 in - ProtocolViolation uu___3 in - { qq = uu___1; qid }) () + FStar_Compiler_Util.format1 "Unknown query '%s'" query in + FStar_Interactive_Ide_Types.ProtocolViolation uu___3 in + { + FStar_Interactive_Ide_Types.qq = uu___1; + FStar_Interactive_Ide_Types.qid = qid + }) () with | FStar_Interactive_JsonHelper.InvalidQuery msg -> - { qq = (ProtocolViolation msg); qid } + { + FStar_Interactive_Ide_Types.qq = + (FStar_Interactive_Ide_Types.ProtocolViolation msg); + FStar_Interactive_Ide_Types.qid = qid + } | FStar_Interactive_JsonHelper.UnexpectedJsonType (expected, got) -> wrap_js_failure qid expected got -let (deserialize_interactive_query : FStar_Compiler_Util.json -> query) = +let (deserialize_interactive_query : + FStar_Compiler_Util.json -> FStar_Interactive_Ide_Types.query) = fun js_query -> try (fun uu___ -> match () with | () -> unpack_interactive_query js_query) () with | FStar_Interactive_JsonHelper.InvalidQuery msg -> - { qq = (ProtocolViolation msg); qid = "?" } + { + FStar_Interactive_Ide_Types.qq = + (FStar_Interactive_Ide_Types.ProtocolViolation msg); + FStar_Interactive_Ide_Types.qid = "?" + } | FStar_Interactive_JsonHelper.UnexpectedJsonType (expected, got) -> wrap_js_failure "?" expected got -let (parse_interactive_query : Prims.string -> query) = +let (parse_interactive_query : + Prims.string -> FStar_Interactive_Ide_Types.query) = fun query_str -> let uu___ = FStar_Compiler_Util.json_of_string query_str in match uu___ with | FStar_Pervasives_Native.None -> - { qq = (ProtocolViolation "Json parsing failed."); qid = "?" } + { + FStar_Interactive_Ide_Types.qq = + (FStar_Interactive_Ide_Types.ProtocolViolation + "Json parsing failed."); + FStar_Interactive_Ide_Types.qid = "?" + } | FStar_Pervasives_Native.Some request -> deserialize_interactive_query request -let (read_interactive_query : FStar_Compiler_Util.stream_reader -> query) = +let (read_interactive_query : + FStar_Compiler_Util.stream_reader -> FStar_Interactive_Ide_Types.query) = fun stream -> let uu___ = FStar_Compiler_Util.read_line stream in match uu___ with @@ -808,60 +545,6 @@ let json_of_opt : fun opt_a -> let uu___ = FStar_Compiler_Util.map_option json_of_a opt_a in FStar_Compiler_Util.dflt FStar_Compiler_Util.JsonNull uu___ -let (json_of_issue_level : - FStar_Errors.issue_level -> FStar_Compiler_Util.json) = - fun i -> - FStar_Compiler_Util.JsonStr - (match i with - | FStar_Errors.ENotImplemented -> "not-implemented" - | FStar_Errors.EInfo -> "info" - | FStar_Errors.EWarning -> "warning" - | FStar_Errors.EError -> "error") -let (json_of_issue : FStar_Errors.issue -> FStar_Compiler_Util.json) = - fun issue -> - let uu___ = - let uu___1 = - let uu___2 = - let uu___3 = - let uu___4 = - let uu___5 = FStar_Errors.issue_message issue in - FStar_Compiler_Util.JsonStr uu___5 in - ("message", uu___4) in - let uu___4 = - let uu___5 = - let uu___6 = - let uu___7 = - let uu___8 = - match issue.FStar_Errors.issue_range with - | FStar_Pervasives_Native.None -> [] - | FStar_Pervasives_Native.Some r -> - let uu___9 = FStar_Compiler_Range.json_of_use_range r in - [uu___9] in - let uu___9 = - match issue.FStar_Errors.issue_range with - | FStar_Pervasives_Native.Some r when - let uu___10 = FStar_Compiler_Range.def_range r in - let uu___11 = FStar_Compiler_Range.use_range r in - uu___10 <> uu___11 -> - let uu___10 = - FStar_Compiler_Range.json_of_def_range r in - [uu___10] - | uu___10 -> [] in - FStar_Compiler_List.op_At uu___8 uu___9 in - FStar_Compiler_Util.JsonList uu___7 in - ("ranges", uu___6) in - [uu___5] in - uu___3 :: uu___4 in - FStar_Compiler_List.op_At - (match issue.FStar_Errors.issue_number with - | FStar_Pervasives_Native.None -> [] - | FStar_Pervasives_Native.Some n -> - [("number", (FStar_Compiler_Util.JsonInt n))]) uu___2 in - FStar_Compiler_List.op_At - [("level", (json_of_issue_level issue.FStar_Errors.issue_level))] - uu___1 in - FStar_Compiler_Effect.op_Less_Bar - (fun uu___1 -> FStar_Compiler_Util.JsonAssoc uu___1) uu___ let (alist_of_symbol_lookup_result : FStar_Interactive_QueryHelper.sl_reponse -> (Prims.string * FStar_Compiler_Util.json) Prims.list) @@ -902,12 +585,14 @@ let (alist_of_symbol_lookup_result : :: uu___ let (alist_of_protocol_info : (Prims.string * FStar_Compiler_Util.json) Prims.list) = - let js_version = FStar_Compiler_Util.JsonInt interactive_protocol_vernum in + let js_version = + FStar_Compiler_Util.JsonInt + FStar_Interactive_Ide_Types.interactive_protocol_vernum in let js_features = let uu___ = FStar_Compiler_List.map (fun uu___1 -> FStar_Compiler_Util.JsonStr uu___1) - interactive_protocol_features in + FStar_Interactive_Ide_Types.interactive_protocol_features in FStar_Compiler_Effect.op_Less_Bar (fun uu___1 -> FStar_Compiler_Util.JsonList uu___1) uu___ in [("version", js_version); ("features", js_features)] @@ -1078,7 +763,8 @@ let (json_of_fstar_option : fstar_option -> FStar_Compiler_Util.json) = FStar_Compiler_Util.JsonAssoc uu___ let (json_of_response : Prims.string -> - query_status -> FStar_Compiler_Util.json -> FStar_Compiler_Util.json) + FStar_Interactive_Ide_Types.query_status -> + FStar_Compiler_Util.json -> FStar_Compiler_Util.json) = fun qid -> fun status -> @@ -1086,9 +772,11 @@ let (json_of_response : let qid1 = FStar_Compiler_Util.JsonStr qid in let status1 = match status with - | QueryOK -> FStar_Compiler_Util.JsonStr "success" - | QueryNOK -> FStar_Compiler_Util.JsonStr "failure" - | QueryViolatesProtocol -> + | FStar_Interactive_Ide_Types.QueryOK -> + FStar_Compiler_Util.JsonStr "success" + | FStar_Interactive_Ide_Types.QueryNOK -> + FStar_Compiler_Util.JsonStr "failure" + | FStar_Interactive_Ide_Types.QueryViolatesProtocol -> FStar_Compiler_Util.JsonStr "protocol-violation" in FStar_Compiler_Util.JsonAssoc [("kind", (FStar_Compiler_Util.JsonStr "response")); @@ -1096,7 +784,10 @@ let (json_of_response : ("status", status1); ("response", response)] let (write_response : - Prims.string -> query_status -> FStar_Compiler_Util.json -> unit) = + Prims.string -> + FStar_Interactive_Ide_Types.query_status -> + FStar_Compiler_Util.json -> unit) + = fun qid -> fun status -> fun response -> @@ -1129,12 +820,14 @@ let forward_message : fun contents -> let uu___ = json_of_message level contents in callback uu___ let (json_of_hello : FStar_Compiler_Util.json) = - let js_version = FStar_Compiler_Util.JsonInt interactive_protocol_vernum in + let js_version = + FStar_Compiler_Util.JsonInt + FStar_Interactive_Ide_Types.interactive_protocol_vernum in let js_features = let uu___ = FStar_Compiler_List.map (fun uu___1 -> FStar_Compiler_Util.JsonStr uu___1) - interactive_protocol_features in + FStar_Interactive_Ide_Types.interactive_protocol_features in FStar_Compiler_Util.JsonList uu___ in FStar_Compiler_Util.JsonAssoc (("kind", (FStar_Compiler_Util.JsonStr "protocol-info")) :: @@ -1222,19 +915,19 @@ let (trim_option_name : Prims.string -> (Prims.string * Prims.string)) = (opt_prefix, uu___) else ("", opt_name) let (json_of_repl_state : - FStar_Interactive_JsonHelper.repl_state -> FStar_Compiler_Util.json) = + FStar_Interactive_ReplState.repl_state -> FStar_Compiler_Util.json) = fun st -> let filenames uu___ = match uu___ with | (uu___1, (task, uu___2)) -> (match task with - | FStar_Interactive_JsonHelper.LDInterleaved (intf, impl) -> - [intf.FStar_Interactive_JsonHelper.tf_fname; - impl.FStar_Interactive_JsonHelper.tf_fname] - | FStar_Interactive_JsonHelper.LDSingle intf_or_impl -> - [intf_or_impl.FStar_Interactive_JsonHelper.tf_fname] - | FStar_Interactive_JsonHelper.LDInterfaceOfCurrentFile intf -> - [intf.FStar_Interactive_JsonHelper.tf_fname] + | FStar_Interactive_ReplState.LDInterleaved (intf, impl) -> + [intf.FStar_Interactive_ReplState.tf_fname; + impl.FStar_Interactive_ReplState.tf_fname] + | FStar_Interactive_ReplState.LDSingle intf_or_impl -> + [intf_or_impl.FStar_Interactive_ReplState.tf_fname] + | FStar_Interactive_ReplState.LDInterfaceOfCurrentFile intf -> + [intf.FStar_Interactive_ReplState.tf_fname] | uu___3 -> []) in let uu___ = let uu___1 = @@ -1242,7 +935,7 @@ let (json_of_repl_state : let uu___3 = let uu___4 = FStar_Compiler_List.concatMap filenames - st.FStar_Interactive_JsonHelper.repl_deps_stack in + st.FStar_Interactive_ReplState.repl_deps_stack in FStar_Compiler_List.map (fun uu___5 -> FStar_Compiler_Util.JsonStr uu___5) uu___4 in FStar_Compiler_Util.JsonList uu___3 in @@ -1261,62 +954,68 @@ let (json_of_repl_state : let run_exit : 'uuuuu 'uuuuu1 . 'uuuuu -> - ((query_status * FStar_Compiler_Util.json) * ('uuuuu1, Prims.int) - FStar_Pervasives.either) + ((FStar_Interactive_Ide_Types.query_status * FStar_Compiler_Util.json) + * ('uuuuu1, Prims.int) FStar_Pervasives.either) = fun st -> - ((QueryOK, FStar_Compiler_Util.JsonNull), + ((FStar_Interactive_Ide_Types.QueryOK, FStar_Compiler_Util.JsonNull), (FStar_Pervasives.Inr Prims.int_zero)) let run_describe_protocol : 'uuuuu 'uuuuu1 . 'uuuuu -> - ((query_status * FStar_Compiler_Util.json) * ('uuuuu, 'uuuuu1) - FStar_Pervasives.either) + ((FStar_Interactive_Ide_Types.query_status * FStar_Compiler_Util.json) + * ('uuuuu, 'uuuuu1) FStar_Pervasives.either) = fun st -> - ((QueryOK, (FStar_Compiler_Util.JsonAssoc alist_of_protocol_info)), + ((FStar_Interactive_Ide_Types.QueryOK, + (FStar_Compiler_Util.JsonAssoc alist_of_protocol_info)), (FStar_Pervasives.Inl st)) let run_describe_repl : 'uuuuu . - FStar_Interactive_JsonHelper.repl_state -> - ((query_status * FStar_Compiler_Util.json) * - (FStar_Interactive_JsonHelper.repl_state, 'uuuuu) + FStar_Interactive_ReplState.repl_state -> + ((FStar_Interactive_Ide_Types.query_status * FStar_Compiler_Util.json) + * (FStar_Interactive_ReplState.repl_state, 'uuuuu) FStar_Pervasives.either) = fun st -> - let uu___ = let uu___1 = json_of_repl_state st in (QueryOK, uu___1) in + let uu___ = + let uu___1 = json_of_repl_state st in + (FStar_Interactive_Ide_Types.QueryOK, uu___1) in (uu___, (FStar_Pervasives.Inl st)) let run_protocol_violation : 'uuuuu 'uuuuu1 . 'uuuuu -> Prims.string -> - ((query_status * FStar_Compiler_Util.json) * ('uuuuu, 'uuuuu1) + ((FStar_Interactive_Ide_Types.query_status * + FStar_Compiler_Util.json) * ('uuuuu, 'uuuuu1) FStar_Pervasives.either) = fun st -> fun message -> - ((QueryViolatesProtocol, (FStar_Compiler_Util.JsonStr message)), - (FStar_Pervasives.Inl st)) + ((FStar_Interactive_Ide_Types.QueryViolatesProtocol, + (FStar_Compiler_Util.JsonStr message)), (FStar_Pervasives.Inl st)) let run_generic_error : 'uuuuu 'uuuuu1 . 'uuuuu -> Prims.string -> - ((query_status * FStar_Compiler_Util.json) * ('uuuuu, 'uuuuu1) + ((FStar_Interactive_Ide_Types.query_status * + FStar_Compiler_Util.json) * ('uuuuu, 'uuuuu1) FStar_Pervasives.either) = fun st -> fun message -> - ((QueryNOK, (FStar_Compiler_Util.JsonStr message)), - (FStar_Pervasives.Inl st)) + ((FStar_Interactive_Ide_Types.QueryNOK, + (FStar_Compiler_Util.JsonStr message)), (FStar_Pervasives.Inl st)) let (collect_errors : unit -> FStar_Errors.issue Prims.list) = fun uu___ -> let errors = FStar_Errors.report_all () in FStar_Errors.clear (); errors let run_segment : 'uuuuu . - FStar_Interactive_JsonHelper.repl_state -> + FStar_Interactive_ReplState.repl_state -> Prims.string -> - ((query_status * FStar_Compiler_Util.json) * - (FStar_Interactive_JsonHelper.repl_state, 'uuuuu) + ((FStar_Interactive_Ide_Types.query_status * + FStar_Compiler_Util.json) * + (FStar_Interactive_ReplState.repl_state, 'uuuuu) FStar_Pervasives.either) = fun st -> @@ -1338,7 +1037,7 @@ let run_segment : | FStar_Parser_Driver.Modul (FStar_Parser_AST.Interface (uu___2, decls, uu___3)) -> decls in let uu___ = - with_captured_errors st.FStar_Interactive_JsonHelper.repl_env + with_captured_errors st.FStar_Interactive_ReplState.repl_env FStar_Compiler_Util.sigint_ignore (fun uu___1 -> let uu___2 = collect_decls () in @@ -1349,8 +1048,10 @@ let run_segment : let errors = let uu___1 = collect_errors () in FStar_Compiler_Effect.op_Bar_Greater uu___1 - (FStar_Compiler_List.map json_of_issue) in - ((QueryNOK, (FStar_Compiler_Util.JsonList errors)), + (FStar_Compiler_List.map + FStar_Interactive_Ide_Types.json_of_issue) in + ((FStar_Interactive_Ide_Types.QueryNOK, + (FStar_Compiler_Util.JsonList errors)), (FStar_Pervasives.Inl st)) | FStar_Pervasives_Native.Some decls -> let json_of_decl decl = @@ -1366,41 +1067,46 @@ let run_segment : let uu___1 = FStar_Compiler_List.map json_of_decl decls in FStar_Compiler_Effect.op_Less_Bar (fun uu___2 -> FStar_Compiler_Util.JsonList uu___2) uu___1 in - ((QueryOK, (FStar_Compiler_Util.JsonAssoc [("decls", js_decls)])), + ((FStar_Interactive_Ide_Types.QueryOK, + (FStar_Compiler_Util.JsonAssoc [("decls", js_decls)])), (FStar_Pervasives.Inl st)) let run_vfs_add : 'uuuuu . - FStar_Interactive_JsonHelper.repl_state -> + FStar_Interactive_ReplState.repl_state -> Prims.string FStar_Pervasives_Native.option -> Prims.string -> - ((query_status * FStar_Compiler_Util.json) * - (FStar_Interactive_JsonHelper.repl_state, 'uuuuu) + ((FStar_Interactive_Ide_Types.query_status * + FStar_Compiler_Util.json) * + (FStar_Interactive_ReplState.repl_state, 'uuuuu) FStar_Pervasives.either) = fun st -> fun opt_fname -> fun contents -> let fname = - FStar_Compiler_Util.dflt st.FStar_Interactive_JsonHelper.repl_fname + FStar_Compiler_Util.dflt st.FStar_Interactive_ReplState.repl_fname opt_fname in FStar_Parser_ParseIt.add_vfs_entry fname contents; - ((QueryOK, FStar_Compiler_Util.JsonNull), (FStar_Pervasives.Inl st)) + ((FStar_Interactive_Ide_Types.QueryOK, FStar_Compiler_Util.JsonNull), + (FStar_Pervasives.Inl st)) let run_pop : 'uuuuu . - FStar_Interactive_JsonHelper.repl_state -> - ((query_status * FStar_Compiler_Util.json) * - (FStar_Interactive_JsonHelper.repl_state, 'uuuuu) + FStar_Interactive_ReplState.repl_state -> + ((FStar_Interactive_Ide_Types.query_status * FStar_Compiler_Util.json) + * (FStar_Interactive_ReplState.repl_state, 'uuuuu) FStar_Pervasives.either) = fun st -> let uu___ = nothing_left_to_pop st in if uu___ then - ((QueryNOK, (FStar_Compiler_Util.JsonStr "Too many pops")), + ((FStar_Interactive_Ide_Types.QueryNOK, + (FStar_Compiler_Util.JsonStr "Too many pops")), (FStar_Pervasives.Inl st)) else (let st' = FStar_Interactive_PushHelper.pop_repl "pop_query" st in - ((QueryOK, FStar_Compiler_Util.JsonNull), (FStar_Pervasives.Inl st'))) + ((FStar_Interactive_Ide_Types.QueryOK, FStar_Compiler_Util.JsonNull), + (FStar_Pervasives.Inl st'))) let (write_progress : Prims.string FStar_Pervasives_Native.option -> (Prims.string * FStar_Compiler_Util.json) Prims.list -> unit) @@ -1416,42 +1122,48 @@ let (write_progress : json_of_message "progress" (FStar_Compiler_Util.JsonAssoc js_contents) in FStar_Interactive_JsonHelper.write_json uu___ +let (write_error : + (Prims.string * FStar_Compiler_Util.json) Prims.list -> unit) = + fun contents -> + let uu___ = + json_of_message "error" (FStar_Compiler_Util.JsonAssoc contents) in + FStar_Interactive_JsonHelper.write_json uu___ let (write_repl_ld_task_progress : - FStar_Interactive_JsonHelper.repl_task -> unit) = + FStar_Interactive_ReplState.repl_task -> unit) = fun task -> match task with - | FStar_Interactive_JsonHelper.LDInterleaved (uu___, tf) -> + | FStar_Interactive_ReplState.LDInterleaved (uu___, tf) -> let modname = FStar_Parser_Dep.module_name_of_file - tf.FStar_Interactive_JsonHelper.tf_fname in + tf.FStar_Interactive_ReplState.tf_fname in write_progress (FStar_Pervasives_Native.Some "loading-dependency") [("modname", (FStar_Compiler_Util.JsonStr modname))] - | FStar_Interactive_JsonHelper.LDSingle tf -> + | FStar_Interactive_ReplState.LDSingle tf -> let modname = FStar_Parser_Dep.module_name_of_file - tf.FStar_Interactive_JsonHelper.tf_fname in + tf.FStar_Interactive_ReplState.tf_fname in write_progress (FStar_Pervasives_Native.Some "loading-dependency") [("modname", (FStar_Compiler_Util.JsonStr modname))] - | FStar_Interactive_JsonHelper.LDInterfaceOfCurrentFile tf -> + | FStar_Interactive_ReplState.LDInterfaceOfCurrentFile tf -> let modname = FStar_Parser_Dep.module_name_of_file - tf.FStar_Interactive_JsonHelper.tf_fname in + tf.FStar_Interactive_ReplState.tf_fname in write_progress (FStar_Pervasives_Native.Some "loading-dependency") [("modname", (FStar_Compiler_Util.JsonStr modname))] | uu___ -> () let (load_deps : - FStar_Interactive_JsonHelper.repl_state -> - ((FStar_Interactive_JsonHelper.repl_state * Prims.string Prims.list), - FStar_Interactive_JsonHelper.repl_state) FStar_Pervasives.either) + FStar_Interactive_ReplState.repl_state -> + ((FStar_Interactive_ReplState.repl_state * Prims.string Prims.list), + FStar_Interactive_ReplState.repl_state) FStar_Pervasives.either) = fun st -> let uu___ = - with_captured_errors st.FStar_Interactive_JsonHelper.repl_env + with_captured_errors st.FStar_Interactive_ReplState.repl_env FStar_Compiler_Util.sigint_ignore (fun _env -> let uu___1 = FStar_Interactive_PushHelper.deps_and_repl_ld_tasks_of_our_file - st.FStar_Interactive_JsonHelper.repl_fname in + st.FStar_Interactive_ReplState.repl_fname in FStar_Compiler_Effect.op_Less_Bar (fun uu___2 -> FStar_Pervasives_Native.Some uu___2) uu___1) in match uu___ with @@ -1460,23 +1172,23 @@ let (load_deps : let st1 = let uu___1 = FStar_TypeChecker_Env.set_dep_graph - st.FStar_Interactive_JsonHelper.repl_env dep_graph in + st.FStar_Interactive_ReplState.repl_env dep_graph in { - FStar_Interactive_JsonHelper.repl_line = - (st.FStar_Interactive_JsonHelper.repl_line); - FStar_Interactive_JsonHelper.repl_column = - (st.FStar_Interactive_JsonHelper.repl_column); - FStar_Interactive_JsonHelper.repl_fname = - (st.FStar_Interactive_JsonHelper.repl_fname); - FStar_Interactive_JsonHelper.repl_deps_stack = - (st.FStar_Interactive_JsonHelper.repl_deps_stack); - FStar_Interactive_JsonHelper.repl_curmod = - (st.FStar_Interactive_JsonHelper.repl_curmod); - FStar_Interactive_JsonHelper.repl_env = uu___1; - FStar_Interactive_JsonHelper.repl_stdin = - (st.FStar_Interactive_JsonHelper.repl_stdin); - FStar_Interactive_JsonHelper.repl_names = - (st.FStar_Interactive_JsonHelper.repl_names) + FStar_Interactive_ReplState.repl_line = + (st.FStar_Interactive_ReplState.repl_line); + FStar_Interactive_ReplState.repl_column = + (st.FStar_Interactive_ReplState.repl_column); + FStar_Interactive_ReplState.repl_fname = + (st.FStar_Interactive_ReplState.repl_fname); + FStar_Interactive_ReplState.repl_deps_stack = + (st.FStar_Interactive_ReplState.repl_deps_stack); + FStar_Interactive_ReplState.repl_curmod = + (st.FStar_Interactive_ReplState.repl_curmod); + FStar_Interactive_ReplState.repl_env = uu___1; + FStar_Interactive_ReplState.repl_stdin = + (st.FStar_Interactive_ReplState.repl_stdin); + FStar_Interactive_ReplState.repl_names = + (st.FStar_Interactive_ReplState.repl_names) } in let uu___1 = run_repl_ld_transactions st1 tasks write_repl_ld_task_progress in @@ -1500,30 +1212,59 @@ let (rephrase_dependency_error : FStar_Errors.issue -> FStar_Errors.issue) = FStar_Errors.issue_number = (issue.FStar_Errors.issue_number); FStar_Errors.issue_ctx = (issue.FStar_Errors.issue_ctx) } -let run_push_without_deps : - 'uuuuu . - FStar_Interactive_JsonHelper.repl_state -> - push_query -> - ((query_status * FStar_Compiler_Util.json) * - (FStar_Interactive_JsonHelper.repl_state, 'uuuuu) - FStar_Pervasives.either) +let (write_full_buffer_fragment_progress : + (FStar_Parser_AST.decl, FStar_Errors.issue Prims.list) + FStar_Pervasives.either -> unit) + = + fun di -> + match di with + | FStar_Pervasives.Inl d -> + let uu___ = + let uu___1 = + let uu___2 = + FStar_Compiler_Range.json_of_def_range + d.FStar_Parser_AST.drange in + ("ranges", uu___2) in + [uu___1] in + write_progress + (FStar_Pervasives_Native.Some "full-buffer-fragment-ok") uu___ + | FStar_Pervasives.Inr issues -> + let qid = + let uu___ = FStar_Compiler_Effect.op_Bang repl_current_qid in + match uu___ with + | FStar_Pervasives_Native.None -> "unknown" + | FStar_Pervasives_Native.Some q -> q in + let uu___ = + let uu___1 = + let uu___2 = + FStar_Compiler_List.map + FStar_Interactive_Ide_Types.json_of_issue issues in + FStar_Compiler_Util.JsonList uu___2 in + json_of_response qid FStar_Interactive_Ide_Types.QueryNOK uu___1 in + FStar_Interactive_JsonHelper.write_json uu___ +let (run_push_without_deps : + FStar_Interactive_ReplState.repl_state -> + FStar_Interactive_Ide_Types.push_query -> + ((FStar_Interactive_Ide_Types.query_status * FStar_Compiler_Util.json) + * (FStar_Interactive_ReplState.repl_state, Prims.int) + FStar_Pervasives.either)) = fun st -> - fun query1 -> + fun query -> let set_nosynth_flag st1 flag = { - FStar_Interactive_JsonHelper.repl_line = - (st1.FStar_Interactive_JsonHelper.repl_line); - FStar_Interactive_JsonHelper.repl_column = - (st1.FStar_Interactive_JsonHelper.repl_column); - FStar_Interactive_JsonHelper.repl_fname = - (st1.FStar_Interactive_JsonHelper.repl_fname); - FStar_Interactive_JsonHelper.repl_deps_stack = - (st1.FStar_Interactive_JsonHelper.repl_deps_stack); - FStar_Interactive_JsonHelper.repl_curmod = - (st1.FStar_Interactive_JsonHelper.repl_curmod); - FStar_Interactive_JsonHelper.repl_env = - (let uu___ = st1.FStar_Interactive_JsonHelper.repl_env in + FStar_Interactive_ReplState.repl_line = + (st1.FStar_Interactive_ReplState.repl_line); + FStar_Interactive_ReplState.repl_column = + (st1.FStar_Interactive_ReplState.repl_column); + FStar_Interactive_ReplState.repl_fname = + (st1.FStar_Interactive_ReplState.repl_fname); + FStar_Interactive_ReplState.repl_deps_stack = + (st1.FStar_Interactive_ReplState.repl_deps_stack); + FStar_Interactive_ReplState.repl_curmod = + (st1.FStar_Interactive_ReplState.repl_curmod); + FStar_Interactive_ReplState.repl_env = + (let uu___ = st1.FStar_Interactive_ReplState.repl_env in { FStar_TypeChecker_Env.solver = (uu___.FStar_TypeChecker_Env.solver); @@ -1623,82 +1364,104 @@ let run_push_without_deps : FStar_TypeChecker_Env.core_check = (uu___.FStar_TypeChecker_Env.core_check) }); - FStar_Interactive_JsonHelper.repl_stdin = - (st1.FStar_Interactive_JsonHelper.repl_stdin); - FStar_Interactive_JsonHelper.repl_names = - (st1.FStar_Interactive_JsonHelper.repl_names) + FStar_Interactive_ReplState.repl_stdin = + (st1.FStar_Interactive_ReplState.repl_stdin); + FStar_Interactive_ReplState.repl_names = + (st1.FStar_Interactive_ReplState.repl_names) } in - let uu___ = query1 in + let uu___ = query in match uu___ with - | { push_kind; push_code = text; push_line = line; - push_column = column; push_peek_only = peek_only;_} -> - let frag = - { - FStar_Parser_ParseIt.frag_fname = ""; - FStar_Parser_ParseIt.frag_text = text; - FStar_Parser_ParseIt.frag_line = line; - FStar_Parser_ParseIt.frag_col = column - } in + | { FStar_Interactive_Ide_Types.push_kind = push_kind; + FStar_Interactive_Ide_Types.push_line = line; + FStar_Interactive_Ide_Types.push_column = column; + FStar_Interactive_Ide_Types.push_peek_only = peek_only; + FStar_Interactive_Ide_Types.push_code_or_decl = code_or_decl;_} -> ((let uu___2 = FStar_Options.ide_id_info_off () in if uu___2 then FStar_TypeChecker_Env.toggle_id_info - st.FStar_Interactive_JsonHelper.repl_env false + st.FStar_Interactive_ReplState.repl_env false else FStar_TypeChecker_Env.toggle_id_info - st.FStar_Interactive_JsonHelper.repl_env true); - (let st1 = set_nosynth_flag st peek_only in + st.FStar_Interactive_ReplState.repl_env true); + (let frag = + match code_or_decl with + | FStar_Pervasives.Inl text -> + FStar_Pervasives.Inl + { + FStar_Parser_ParseIt.frag_fname = ""; + FStar_Parser_ParseIt.frag_text = text; + FStar_Parser_ParseIt.frag_line = line; + FStar_Parser_ParseIt.frag_col = column + } + | FStar_Pervasives.Inr decl -> FStar_Pervasives.Inr decl in + let st1 = set_nosynth_flag st peek_only in let uu___2 = run_repl_transaction st1 push_kind peek_only - (FStar_Interactive_JsonHelper.PushFragment frag) in + (FStar_Interactive_ReplState.PushFragment frag) in match uu___2 with | (success, st2) -> let st3 = set_nosynth_flag st2 false in let status = - if success || peek_only then QueryOK else QueryNOK in - let json_errors = - let uu___3 = - let uu___4 = collect_errors () in - FStar_Compiler_Effect.op_Bar_Greater uu___4 - (FStar_Compiler_List.map json_of_issue) in - FStar_Compiler_Util.JsonList uu___3 in - let st4 = - if success - then - { - FStar_Interactive_JsonHelper.repl_line = line; - FStar_Interactive_JsonHelper.repl_column = column; - FStar_Interactive_JsonHelper.repl_fname = - (st3.FStar_Interactive_JsonHelper.repl_fname); - FStar_Interactive_JsonHelper.repl_deps_stack = - (st3.FStar_Interactive_JsonHelper.repl_deps_stack); - FStar_Interactive_JsonHelper.repl_curmod = - (st3.FStar_Interactive_JsonHelper.repl_curmod); - FStar_Interactive_JsonHelper.repl_env = - (st3.FStar_Interactive_JsonHelper.repl_env); - FStar_Interactive_JsonHelper.repl_stdin = - (st3.FStar_Interactive_JsonHelper.repl_stdin); - FStar_Interactive_JsonHelper.repl_names = - (st3.FStar_Interactive_JsonHelper.repl_names) - } - else st3 in - ((status, json_errors), (FStar_Pervasives.Inl st4)))) -let run_push_with_deps : - 'uuuuu . - FStar_Interactive_JsonHelper.repl_state -> - push_query -> - ((query_status * FStar_Compiler_Util.json) * - (FStar_Interactive_JsonHelper.repl_state, 'uuuuu) - FStar_Pervasives.either) + if success || peek_only + then FStar_Interactive_Ide_Types.QueryOK + else FStar_Interactive_Ide_Types.QueryNOK in + let errs = collect_errors () in + let has_error = + FStar_Compiler_List.existsb + (fun i -> + match i.FStar_Errors.issue_level with + | FStar_Errors.EError -> true + | FStar_Errors.ENotImplemented -> true + | uu___3 -> false) errs in + ((match code_or_decl with + | FStar_Pervasives.Inr d when Prims.op_Negation has_error + -> + write_full_buffer_fragment_progress + (FStar_Pervasives.Inl d) + | uu___4 -> ()); + (let json_errors = + let uu___4 = + FStar_Compiler_Effect.op_Bar_Greater errs + (FStar_Compiler_List.map + FStar_Interactive_Ide_Types.json_of_issue) in + FStar_Compiler_Util.JsonList uu___4 in + let st4 = + if success + then + { + FStar_Interactive_ReplState.repl_line = line; + FStar_Interactive_ReplState.repl_column = column; + FStar_Interactive_ReplState.repl_fname = + (st3.FStar_Interactive_ReplState.repl_fname); + FStar_Interactive_ReplState.repl_deps_stack = + (st3.FStar_Interactive_ReplState.repl_deps_stack); + FStar_Interactive_ReplState.repl_curmod = + (st3.FStar_Interactive_ReplState.repl_curmod); + FStar_Interactive_ReplState.repl_env = + (st3.FStar_Interactive_ReplState.repl_env); + FStar_Interactive_ReplState.repl_stdin = + (st3.FStar_Interactive_ReplState.repl_stdin); + FStar_Interactive_ReplState.repl_names = + (st3.FStar_Interactive_ReplState.repl_names) + } + else st3 in + ((status, json_errors), (FStar_Pervasives.Inl st4)))))) +let (run_push_with_deps : + FStar_Interactive_ReplState.repl_state -> + FStar_Interactive_Ide_Types.push_query -> + ((FStar_Interactive_Ide_Types.query_status * FStar_Compiler_Util.json) + * (FStar_Interactive_ReplState.repl_state, Prims.int) + FStar_Pervasives.either)) = fun st -> - fun query1 -> + fun query -> (let uu___1 = FStar_Options.debug_any () in if uu___1 then FStar_Compiler_Util.print_string "Reloading dependencies" else ()); FStar_TypeChecker_Env.toggle_id_info - st.FStar_Interactive_JsonHelper.repl_env false; + st.FStar_Interactive_ReplState.repl_env false; (let uu___2 = load_deps st in match uu___2 with | FStar_Pervasives.Inr st1 -> @@ -1707,50 +1470,51 @@ let run_push_with_deps : FStar_Compiler_List.map rephrase_dependency_error uu___3 in let js_errors = FStar_Compiler_Effect.op_Bar_Greater errors - (FStar_Compiler_List.map json_of_issue) in - ((QueryNOK, (FStar_Compiler_Util.JsonList js_errors)), + (FStar_Compiler_List.map + FStar_Interactive_Ide_Types.json_of_issue) in + ((FStar_Interactive_Ide_Types.QueryNOK, + (FStar_Compiler_Util.JsonList js_errors)), (FStar_Pervasives.Inl st1)) | FStar_Pervasives.Inl (st1, deps) -> ((let uu___4 = FStar_Options.restore_cmd_line_options false in FStar_Compiler_Effect.op_Bar_Greater uu___4 (fun uu___5 -> ())); (let names = FStar_Interactive_PushHelper.add_module_completions - st1.FStar_Interactive_JsonHelper.repl_fname deps - st1.FStar_Interactive_JsonHelper.repl_names in + st1.FStar_Interactive_ReplState.repl_fname deps + st1.FStar_Interactive_ReplState.repl_names in run_push_without_deps { - FStar_Interactive_JsonHelper.repl_line = - (st1.FStar_Interactive_JsonHelper.repl_line); - FStar_Interactive_JsonHelper.repl_column = - (st1.FStar_Interactive_JsonHelper.repl_column); - FStar_Interactive_JsonHelper.repl_fname = - (st1.FStar_Interactive_JsonHelper.repl_fname); - FStar_Interactive_JsonHelper.repl_deps_stack = - (st1.FStar_Interactive_JsonHelper.repl_deps_stack); - FStar_Interactive_JsonHelper.repl_curmod = - (st1.FStar_Interactive_JsonHelper.repl_curmod); - FStar_Interactive_JsonHelper.repl_env = - (st1.FStar_Interactive_JsonHelper.repl_env); - FStar_Interactive_JsonHelper.repl_stdin = - (st1.FStar_Interactive_JsonHelper.repl_stdin); - FStar_Interactive_JsonHelper.repl_names = names - } query1))) -let run_push : - 'uuuuu . - FStar_Interactive_JsonHelper.repl_state -> - push_query -> - ((query_status * FStar_Compiler_Util.json) * - (FStar_Interactive_JsonHelper.repl_state, 'uuuuu) - FStar_Pervasives.either) + FStar_Interactive_ReplState.repl_line = + (st1.FStar_Interactive_ReplState.repl_line); + FStar_Interactive_ReplState.repl_column = + (st1.FStar_Interactive_ReplState.repl_column); + FStar_Interactive_ReplState.repl_fname = + (st1.FStar_Interactive_ReplState.repl_fname); + FStar_Interactive_ReplState.repl_deps_stack = + (st1.FStar_Interactive_ReplState.repl_deps_stack); + FStar_Interactive_ReplState.repl_curmod = + (st1.FStar_Interactive_ReplState.repl_curmod); + FStar_Interactive_ReplState.repl_env = + (st1.FStar_Interactive_ReplState.repl_env); + FStar_Interactive_ReplState.repl_stdin = + (st1.FStar_Interactive_ReplState.repl_stdin); + FStar_Interactive_ReplState.repl_names = names + } query))) +let (run_push : + FStar_Interactive_ReplState.repl_state -> + FStar_Interactive_Ide_Types.push_query -> + ((FStar_Interactive_Ide_Types.query_status * FStar_Compiler_Util.json) + * (FStar_Interactive_ReplState.repl_state, Prims.int) + FStar_Pervasives.either)) = fun st -> - fun query1 -> + fun query -> let uu___ = nothing_left_to_pop st in if uu___ - then run_push_with_deps st query1 - else run_push_without_deps st query1 + then run_push_with_deps st query + else run_push_without_deps st query let (run_symbol_lookup : - FStar_Interactive_JsonHelper.repl_state -> + FStar_Interactive_ReplState.repl_state -> Prims.string -> FStar_Interactive_QueryHelper.position FStar_Pervasives_Native.option -> @@ -1766,7 +1530,7 @@ let (run_symbol_lookup : fun requested_info -> let uu___ = FStar_Interactive_QueryHelper.symlookup - st.FStar_Interactive_JsonHelper.repl_env symbol pos_opt + st.FStar_Interactive_ReplState.repl_env symbol pos_opt requested_info in match uu___ with | FStar_Pervasives_Native.None -> @@ -1800,7 +1564,7 @@ let (run_option_lookup : ("option", uu___4) in FStar_Pervasives.Inr uu___3) let (run_module_lookup : - FStar_Interactive_JsonHelper.repl_state -> + FStar_Interactive_ReplState.repl_state -> Prims.string -> (Prims.string, (Prims.string * (Prims.string * FStar_Compiler_Util.json) Prims.list)) @@ -1808,10 +1572,10 @@ let (run_module_lookup : = fun st -> fun symbol -> - let query1 = FStar_Compiler_Util.split symbol "." in + let query = FStar_Compiler_Util.split symbol "." in let uu___ = FStar_Interactive_CompletionTable.find_module_or_ns - st.FStar_Interactive_JsonHelper.repl_names query1 in + st.FStar_Interactive_ReplState.repl_names query in match uu___ with | FStar_Pervasives_Native.None -> FStar_Pervasives.Inl "No such module or namespace" @@ -1830,7 +1594,7 @@ let (run_module_lookup : ("namespace", uu___2) in FStar_Pervasives.Inr uu___1 let (run_code_lookup : - FStar_Interactive_JsonHelper.repl_state -> + FStar_Interactive_ReplState.repl_state -> Prims.string -> FStar_Interactive_QueryHelper.position FStar_Pervasives_Native.option -> @@ -1855,9 +1619,9 @@ let (run_code_lookup : FStar_Pervasives.Inl "No such symbol, module, or namespace.") let (run_lookup' : - FStar_Interactive_JsonHelper.repl_state -> + FStar_Interactive_ReplState.repl_state -> Prims.string -> - lookup_context -> + FStar_Interactive_Ide_Types.lookup_context -> FStar_Interactive_QueryHelper.position FStar_Pervasives_Native.option -> Prims.string Prims.list -> @@ -1872,21 +1636,25 @@ let (run_lookup' : fun pos_opt -> fun requested_info -> match context with - | LKSymbolOnly -> + | FStar_Interactive_Ide_Types.LKSymbolOnly -> run_symbol_lookup st symbol pos_opt requested_info - | LKModule -> run_module_lookup st symbol - | LKOption -> run_option_lookup symbol - | LKCode -> run_code_lookup st symbol pos_opt requested_info + | FStar_Interactive_Ide_Types.LKModule -> + run_module_lookup st symbol + | FStar_Interactive_Ide_Types.LKOption -> + run_option_lookup symbol + | FStar_Interactive_Ide_Types.LKCode -> + run_code_lookup st symbol pos_opt requested_info let run_lookup : 'uuuuu . - FStar_Interactive_JsonHelper.repl_state -> + FStar_Interactive_ReplState.repl_state -> Prims.string -> - lookup_context -> + FStar_Interactive_Ide_Types.lookup_context -> FStar_Interactive_QueryHelper.position FStar_Pervasives_Native.option -> Prims.string Prims.list -> - ((query_status * FStar_Compiler_Util.json) * - (FStar_Interactive_JsonHelper.repl_state, 'uuuuu) + ((FStar_Interactive_Ide_Types.query_status * + FStar_Compiler_Util.json) * + (FStar_Interactive_ReplState.repl_state, 'uuuuu) FStar_Pervasives.either) = fun st -> @@ -1897,19 +1665,21 @@ let run_lookup : let uu___ = run_lookup' st symbol context pos_opt requested_info in match uu___ with | FStar_Pervasives.Inl err_msg -> - ((QueryNOK, (FStar_Compiler_Util.JsonStr err_msg)), + ((FStar_Interactive_Ide_Types.QueryNOK, + (FStar_Compiler_Util.JsonStr err_msg)), (FStar_Pervasives.Inl st)) | FStar_Pervasives.Inr (kind, info) -> - ((QueryOK, + ((FStar_Interactive_Ide_Types.QueryOK, (FStar_Compiler_Util.JsonAssoc (("kind", (FStar_Compiler_Util.JsonStr kind)) :: info))), (FStar_Pervasives.Inl st)) let run_code_autocomplete : 'uuuuu . - FStar_Interactive_JsonHelper.repl_state -> + FStar_Interactive_ReplState.repl_state -> Prims.string -> - ((query_status * FStar_Compiler_Util.json) * - (FStar_Interactive_JsonHelper.repl_state, 'uuuuu) + ((FStar_Interactive_Ide_Types.query_status * + FStar_Compiler_Util.json) * + (FStar_Interactive_ReplState.repl_state, 'uuuuu) FStar_Pervasives.either) = fun st -> @@ -1918,16 +1688,17 @@ let run_code_autocomplete : let js = FStar_Compiler_List.map FStar_Interactive_CompletionTable.json_of_completion_result result in - ((QueryOK, (FStar_Compiler_Util.JsonList js)), - (FStar_Pervasives.Inl st)) + ((FStar_Interactive_Ide_Types.QueryOK, + (FStar_Compiler_Util.JsonList js)), (FStar_Pervasives.Inl st)) let run_module_autocomplete : 'uuuuu 'uuuuu1 'uuuuu2 . - FStar_Interactive_JsonHelper.repl_state -> + FStar_Interactive_ReplState.repl_state -> Prims.string -> 'uuuuu -> 'uuuuu1 -> - ((query_status * FStar_Compiler_Util.json) * - (FStar_Interactive_JsonHelper.repl_state, 'uuuuu2) + ((FStar_Interactive_Ide_Types.query_status * + FStar_Compiler_Util.json) * + (FStar_Interactive_ReplState.repl_state, 'uuuuu2) FStar_Pervasives.either) = fun st -> @@ -1937,14 +1708,14 @@ let run_module_autocomplete : let needle = FStar_Compiler_Util.split search_term "." in let mods_and_nss = FStar_Interactive_CompletionTable.autocomplete_mod_or_ns - st.FStar_Interactive_JsonHelper.repl_names needle + st.FStar_Interactive_ReplState.repl_names needle (fun uu___ -> FStar_Pervasives_Native.Some uu___) in let json = FStar_Compiler_List.map FStar_Interactive_CompletionTable.json_of_completion_result mods_and_nss in - ((QueryOK, (FStar_Compiler_Util.JsonList json)), - (FStar_Pervasives.Inl st)) + ((FStar_Interactive_Ide_Types.QueryOK, + (FStar_Compiler_Util.JsonList json)), (FStar_Pervasives.Inl st)) let candidates_of_fstar_option : 'uuuuu . Prims.int -> @@ -1985,7 +1756,8 @@ let run_option_autocomplete : 'uuuuu -> Prims.string -> 'uuuuu1 -> - ((query_status * FStar_Compiler_Util.json) * ('uuuuu, 'uuuuu2) + ((FStar_Interactive_Ide_Types.query_status * + FStar_Compiler_Util.json) * ('uuuuu, 'uuuuu2) FStar_Pervasives.either) = fun st -> @@ -2006,36 +1778,40 @@ let run_option_autocomplete : FStar_Compiler_List.map FStar_Interactive_CompletionTable.json_of_completion_result results in - ((QueryOK, (FStar_Compiler_Util.JsonList json)), + ((FStar_Interactive_Ide_Types.QueryOK, + (FStar_Compiler_Util.JsonList json)), (FStar_Pervasives.Inl st)) | (uu___1, uu___2) -> - ((QueryNOK, + ((FStar_Interactive_Ide_Types.QueryNOK, (FStar_Compiler_Util.JsonStr "Options should start with '--'")), (FStar_Pervasives.Inl st)) let run_autocomplete : 'uuuuu . - FStar_Interactive_JsonHelper.repl_state -> + FStar_Interactive_ReplState.repl_state -> Prims.string -> - completion_context -> - ((query_status * FStar_Compiler_Util.json) * - (FStar_Interactive_JsonHelper.repl_state, 'uuuuu) + FStar_Interactive_Ide_Types.completion_context -> + ((FStar_Interactive_Ide_Types.query_status * + FStar_Compiler_Util.json) * + (FStar_Interactive_ReplState.repl_state, 'uuuuu) FStar_Pervasives.either) = fun st -> fun search_term -> fun context -> match context with - | CKCode -> run_code_autocomplete st search_term - | CKOption is_reset -> + | FStar_Interactive_Ide_Types.CKCode -> + run_code_autocomplete st search_term + | FStar_Interactive_Ide_Types.CKOption is_reset -> run_option_autocomplete st search_term is_reset - | CKModuleOrNamespace (modules, namespaces) -> + | FStar_Interactive_Ide_Types.CKModuleOrNamespace + (modules, namespaces) -> run_module_autocomplete st search_term modules namespaces let run_and_rewind : 'uuuuu 'uuuuu1 . - FStar_Interactive_JsonHelper.repl_state -> + FStar_Interactive_ReplState.repl_state -> 'uuuuu -> - (FStar_Interactive_JsonHelper.repl_state -> 'uuuuu) -> - ('uuuuu * (FStar_Interactive_JsonHelper.repl_state, 'uuuuu1) + (FStar_Interactive_ReplState.repl_state -> 'uuuuu) -> + ('uuuuu * (FStar_Interactive_ReplState.repl_state, 'uuuuu1) FStar_Pervasives.either) = fun st -> @@ -2044,7 +1820,7 @@ let run_and_rewind : let st1 = FStar_Interactive_PushHelper.push_repl "run_and_rewind" FStar_Interactive_PushHelper.FullCheck - FStar_Interactive_JsonHelper.Noop st in + FStar_Interactive_ReplState.Noop st in let results = try (fun uu___ -> @@ -2067,16 +1843,18 @@ let run_and_rewind : | FStar_Pervasives.Inr e -> FStar_Compiler_Effect.raise e let run_with_parsed_and_tc_term : 'uuuuu 'uuuuu1 'uuuuu2 . - FStar_Interactive_JsonHelper.repl_state -> + FStar_Interactive_ReplState.repl_state -> Prims.string -> 'uuuuu -> 'uuuuu1 -> (FStar_TypeChecker_Env.env -> FStar_Syntax_Syntax.term -> - (query_status * FStar_Compiler_Util.json)) + (FStar_Interactive_Ide_Types.query_status * + FStar_Compiler_Util.json)) -> - ((query_status * FStar_Compiler_Util.json) * - (FStar_Interactive_JsonHelper.repl_state, 'uuuuu2) + ((FStar_Interactive_Ide_Types.query_status * + FStar_Compiler_Util.json) * + (FStar_Interactive_ReplState.repl_state, 'uuuuu2) FStar_Pervasives.either) = fun st -> @@ -2117,11 +1895,10 @@ let run_with_parsed_and_tc_term : let parse frag = let uu___ = FStar_Parser_ParseIt.parse - (FStar_Parser_ParseIt.Toplevel frag) in + (FStar_Parser_ParseIt.Incremental frag) in match uu___ with - | FStar_Parser_ParseIt.ASTFragment - (FStar_Pervasives.Inr decls, uu___1) -> - FStar_Pervasives_Native.Some decls + | FStar_Parser_ParseIt.IncrementalFragment + (decls, uu___1, _err) -> FStar_Pervasives_Native.Some decls | uu___1 -> FStar_Pervasives_Native.None in let desugar env decls = let uu___ = @@ -2132,15 +1909,15 @@ let run_with_parsed_and_tc_term : let uu___ = FStar_TypeChecker_Tc.tc_decls tcenv decls in match uu___ with | (ses, uu___1) -> ses in run_and_rewind st - (QueryNOK, + (FStar_Interactive_Ide_Types.QueryNOK, (FStar_Compiler_Util.JsonStr "Computation interrupted")) (fun st1 -> - let tcenv = st1.FStar_Interactive_JsonHelper.repl_env in + let tcenv = st1.FStar_Interactive_ReplState.repl_env in let frag = dummy_let_fragment term in let uu___ = parse frag in match uu___ with | FStar_Pervasives_Native.None -> - (QueryNOK, + (FStar_Interactive_Ide_Types.QueryNOK, (FStar_Compiler_Util.JsonStr "Could not parse this term")) | FStar_Pervasives_Native.Some decls -> @@ -2149,7 +1926,7 @@ let run_with_parsed_and_tc_term : let ses = typecheck tcenv decls1 in match find_let_body ses with | FStar_Pervasives_Native.None -> - (QueryNOK, + (FStar_Interactive_Ide_Types.QueryNOK, (FStar_Compiler_Util.JsonStr "Typechecking yielded an unexpected term")) | FStar_Pervasives_Native.Some (univs, def) -> @@ -2175,17 +1952,19 @@ let run_with_parsed_and_tc_term : let uu___6 = FStar_Errors.format_issue issue in FStar_Compiler_Util.JsonStr uu___6 in - (QueryNOK, uu___5) + (FStar_Interactive_Ide_Types.QueryNOK, + uu___5) | FStar_Pervasives_Native.None -> FStar_Compiler_Effect.raise uu___3))) let run_compute : 'uuuuu . - FStar_Interactive_JsonHelper.repl_state -> + FStar_Interactive_ReplState.repl_state -> Prims.string -> FStar_TypeChecker_Env.step Prims.list FStar_Pervasives_Native.option -> - ((query_status * FStar_Compiler_Util.json) * - (FStar_Interactive_JsonHelper.repl_state, 'uuuuu) + ((FStar_Interactive_Ide_Types.query_status * + FStar_Compiler_Util.json) * + (FStar_Interactive_ReplState.repl_state, 'uuuuu) FStar_Pervasives.either) = fun st -> @@ -2216,7 +1995,7 @@ let run_compute : FStar_Interactive_QueryHelper.term_to_string tcenv normalized in FStar_Compiler_Util.JsonStr uu___1 in - (QueryOK, uu___)) + (FStar_Interactive_Ide_Types.QueryOK, uu___)) type search_term' = | NameContainsStr of Prims.string | TypeContainsLid of FStar_Ident.lid @@ -2338,15 +2117,16 @@ let (__proj__InvalidSearch__item__uu___ : Prims.exn -> Prims.string) = fun projectee -> match projectee with | InvalidSearch uu___ -> uu___ let run_search : 'uuuuu . - FStar_Interactive_JsonHelper.repl_state -> + FStar_Interactive_ReplState.repl_state -> Prims.string -> - ((query_status * FStar_Compiler_Util.json) * - (FStar_Interactive_JsonHelper.repl_state, 'uuuuu) + ((FStar_Interactive_Ide_Types.query_status * + FStar_Compiler_Util.json) * + (FStar_Interactive_ReplState.repl_state, 'uuuuu) FStar_Pervasives.either) = fun st -> fun search_str -> - let tcenv = st.FStar_Interactive_JsonHelper.repl_env in + let tcenv = st.FStar_Interactive_ReplState.repl_env in let empty_fv_set = FStar_Syntax_Syntax.new_fv_set () in let st_matches candidate term = let found = @@ -2447,98 +2227,185 @@ let run_search : "No results found for query [%s]" kwds in InvalidSearch uu___2 in FStar_Compiler_Effect.raise uu___1 - | uu___1 -> (QueryOK, (FStar_Compiler_Util.JsonList js)))) - () - with | InvalidSearch s -> (QueryNOK, (FStar_Compiler_Util.JsonStr s)) in + | uu___1 -> + (FStar_Interactive_Ide_Types.QueryOK, + (FStar_Compiler_Util.JsonList js)))) () + with + | InvalidSearch s -> + (FStar_Interactive_Ide_Types.QueryNOK, + (FStar_Compiler_Util.JsonStr s)) in (results, (FStar_Pervasives.Inl st)) -let (run_query : - FStar_Interactive_JsonHelper.repl_state -> - query' -> - ((query_status * FStar_Compiler_Util.json) * - (FStar_Interactive_JsonHelper.repl_state, Prims.int) - FStar_Pervasives.either)) +let (as_json_list : + ((FStar_Interactive_Ide_Types.query_status * FStar_Compiler_Util.json) * + (FStar_Interactive_ReplState.repl_state, Prims.int) + FStar_Pervasives.either) -> + ((FStar_Interactive_Ide_Types.query_status * FStar_Compiler_Util.json + Prims.list) * (FStar_Interactive_ReplState.repl_state, Prims.int) + FStar_Pervasives.either)) = - fun st -> - fun q -> - match q with - | Exit -> run_exit st - | DescribeProtocol -> run_describe_protocol st - | DescribeRepl -> run_describe_repl st - | GenericError message -> run_generic_error st message - | ProtocolViolation query1 -> run_protocol_violation st query1 - | Segment c -> run_segment st c - | VfsAdd (fname, contents) -> run_vfs_add st fname contents - | Push pquery -> run_push st pquery - | Pop -> run_pop st - | AutoComplete (search_term1, context) -> - run_autocomplete st search_term1 context - | Lookup (symbol, context, pos_opt, rq_info) -> - run_lookup st symbol context pos_opt rq_info - | Compute (term, rules) -> run_compute st term rules - | Search term -> run_search st term + fun q -> let uu___ = q in match uu___ with | ((q1, j), s) -> ((q1, [j]), s) +type run_query_result = + ((FStar_Interactive_Ide_Types.query_status * FStar_Compiler_Util.json + Prims.list) * (FStar_Interactive_ReplState.repl_state, Prims.int) + FStar_Pervasives.either) +let rec (fold_query : + (FStar_Interactive_ReplState.repl_state -> + FStar_Interactive_Ide_Types.query -> run_query_result) + -> + FStar_Interactive_Ide_Types.query Prims.list -> + FStar_Interactive_ReplState.repl_state -> + FStar_Compiler_Util.json Prims.list -> run_query_result) + = + fun f -> + fun l -> + fun st -> + fun responses -> + match l with + | [] -> + ((FStar_Interactive_Ide_Types.QueryOK, responses), + (FStar_Pervasives.Inl st)) + | q::l1 -> + let uu___ = f st q in + (match uu___ with + | ((status, resp), st') -> + let responses1 = FStar_Compiler_List.op_At responses resp in + (match (status, st') with + | (FStar_Interactive_Ide_Types.QueryOK, + FStar_Pervasives.Inl st1) -> + fold_query f l1 st1 responses1 + | uu___1 -> ((status, responses1), st'))) let (validate_query : - FStar_Interactive_JsonHelper.repl_state -> query -> query) = + FStar_Interactive_ReplState.repl_state -> + FStar_Interactive_Ide_Types.query -> FStar_Interactive_Ide_Types.query) + = fun st -> fun q -> - match q.qq with - | Push - { push_kind = FStar_Interactive_PushHelper.SyntaxCheck; - push_code = uu___; push_line = uu___1; push_column = uu___2; - push_peek_only = false;_} + match q.FStar_Interactive_Ide_Types.qq with + | FStar_Interactive_Ide_Types.Push + { + FStar_Interactive_Ide_Types.push_kind = + FStar_Interactive_PushHelper.SyntaxCheck; + FStar_Interactive_Ide_Types.push_line = uu___; + FStar_Interactive_Ide_Types.push_column = uu___1; + FStar_Interactive_Ide_Types.push_peek_only = false; + FStar_Interactive_Ide_Types.push_code_or_decl = uu___2;_} -> { - qq = - (ProtocolViolation + FStar_Interactive_Ide_Types.qq = + (FStar_Interactive_Ide_Types.ProtocolViolation "Cannot use 'kind': 'syntax' with 'query': 'push'"); - qid = (q.qid) + FStar_Interactive_Ide_Types.qid = + (q.FStar_Interactive_Ide_Types.qid) } | uu___ -> - (match st.FStar_Interactive_JsonHelper.repl_curmod with + (match st.FStar_Interactive_ReplState.repl_curmod with | FStar_Pervasives_Native.None when - query_needs_current_module q.qq -> - { qq = (GenericError "Current module unset"); qid = (q.qid) } + FStar_Interactive_Ide_Types.query_needs_current_module + q.FStar_Interactive_Ide_Types.qq + -> + { + FStar_Interactive_Ide_Types.qq = + (FStar_Interactive_Ide_Types.GenericError + "Current module unset"); + FStar_Interactive_Ide_Types.qid = + (q.FStar_Interactive_Ide_Types.qid) + } | uu___1 -> q) -let (validate_and_run_query : - FStar_Interactive_JsonHelper.repl_state -> - query -> - ((query_status * FStar_Compiler_Util.json) * - (FStar_Interactive_JsonHelper.repl_state, Prims.int) +let rec (run_query : + FStar_Interactive_ReplState.repl_state -> + FStar_Interactive_Ide_Types.query -> + ((FStar_Interactive_Ide_Types.query_status * FStar_Compiler_Util.json + Prims.list) * (FStar_Interactive_ReplState.repl_state, Prims.int) FStar_Pervasives.either)) = fun st -> - fun query1 -> - let query2 = validate_query st query1 in + fun q -> + match q.FStar_Interactive_Ide_Types.qq with + | FStar_Interactive_Ide_Types.Exit -> as_json_list (run_exit st) + | FStar_Interactive_Ide_Types.DescribeProtocol -> + as_json_list (run_describe_protocol st) + | FStar_Interactive_Ide_Types.DescribeRepl -> + let uu___ = run_describe_repl st in as_json_list uu___ + | FStar_Interactive_Ide_Types.GenericError message -> + as_json_list (run_generic_error st message) + | FStar_Interactive_Ide_Types.ProtocolViolation query -> + as_json_list (run_protocol_violation st query) + | FStar_Interactive_Ide_Types.Segment c -> + let uu___ = run_segment st c in as_json_list uu___ + | FStar_Interactive_Ide_Types.VfsAdd (fname, contents) -> + let uu___ = run_vfs_add st fname contents in as_json_list uu___ + | FStar_Interactive_Ide_Types.Push pquery -> + let uu___ = run_push st pquery in as_json_list uu___ + | FStar_Interactive_Ide_Types.Pop -> + let uu___ = run_pop st in as_json_list uu___ + | FStar_Interactive_Ide_Types.FullBuffer code -> + let queries = + FStar_Interactive_Incremental.run_full_buffer st + q.FStar_Interactive_Ide_Types.qid code + write_full_buffer_fragment_progress in + fold_query validate_and_run_query queries st [] + | FStar_Interactive_Ide_Types.AutoComplete (search_term1, context) -> + let uu___ = run_autocomplete st search_term1 context in + as_json_list uu___ + | FStar_Interactive_Ide_Types.Lookup + (symbol, context, pos_opt, rq_info) -> + let uu___ = run_lookup st symbol context pos_opt rq_info in + as_json_list uu___ + | FStar_Interactive_Ide_Types.Compute (term, rules) -> + let uu___ = run_compute st term rules in as_json_list uu___ + | FStar_Interactive_Ide_Types.Search term -> + let uu___ = run_search st term in as_json_list uu___ +and (validate_and_run_query : + FStar_Interactive_ReplState.repl_state -> + FStar_Interactive_Ide_Types.query -> run_query_result) + = + fun st -> + fun query -> + let query1 = validate_query st query in FStar_Compiler_Effect.op_Colon_Equals repl_current_qid - (FStar_Pervasives_Native.Some (query2.qid)); - run_query st query2.qq + (FStar_Pervasives_Native.Some + (query1.FStar_Interactive_Ide_Types.qid)); + (let uu___2 = FStar_Options.debug_any () in + if uu___2 + then + let uu___3 = FStar_Interactive_Ide_Types.query_to_string query1 in + FStar_Compiler_Util.print2 "Running query %s: %s\n" + query1.FStar_Interactive_Ide_Types.qid uu___3 + else ()); + run_query st query1 let (js_repl_eval : - FStar_Interactive_JsonHelper.repl_state -> - query -> - (FStar_Compiler_Util.json * (FStar_Interactive_JsonHelper.repl_state, - Prims.int) FStar_Pervasives.either)) + FStar_Interactive_ReplState.repl_state -> + FStar_Interactive_Ide_Types.query -> + (FStar_Compiler_Util.json Prims.list * + (FStar_Interactive_ReplState.repl_state, Prims.int) + FStar_Pervasives.either)) = fun st -> - fun query1 -> - let uu___ = validate_and_run_query st query1 in + fun query -> + let uu___ = validate_and_run_query st query in match uu___ with - | ((status, response), st_opt) -> - let js_response = json_of_response query1.qid status response in - (js_response, st_opt) + | ((status, responses), st_opt) -> + let js_responses = + FStar_Compiler_List.map + (json_of_response query.FStar_Interactive_Ide_Types.qid status) + responses in + (js_responses, st_opt) let (js_repl_eval_js : - FStar_Interactive_JsonHelper.repl_state -> + FStar_Interactive_ReplState.repl_state -> FStar_Compiler_Util.json -> - (FStar_Compiler_Util.json * (FStar_Interactive_JsonHelper.repl_state, - Prims.int) FStar_Pervasives.either)) + (FStar_Compiler_Util.json Prims.list * + (FStar_Interactive_ReplState.repl_state, Prims.int) + FStar_Pervasives.either)) = fun st -> fun query_js -> let uu___ = deserialize_interactive_query query_js in js_repl_eval st uu___ let (js_repl_eval_str : - FStar_Interactive_JsonHelper.repl_state -> + FStar_Interactive_ReplState.repl_state -> Prims.string -> - (Prims.string * (FStar_Interactive_JsonHelper.repl_state, Prims.int) - FStar_Pervasives.either)) + (Prims.string Prims.list * (FStar_Interactive_ReplState.repl_state, + Prims.int) FStar_Pervasives.either)) = fun st -> fun query_str -> @@ -2547,7 +2414,9 @@ let (js_repl_eval_str : js_repl_eval st uu___1 in match uu___ with | (js_response, st_opt) -> - let uu___1 = FStar_Compiler_Util.string_of_json js_response in + let uu___1 = + FStar_Compiler_List.map FStar_Compiler_Util.string_of_json + js_response in (uu___1, st_opt) let (js_repl_init_opts : unit -> unit) = fun uu___ -> @@ -2567,17 +2436,29 @@ let (js_repl_init_opts : unit -> unit) = failwith "repl_init: Too many file names given in --ide invocation" | uu___2 -> ())) -let rec (go : FStar_Interactive_JsonHelper.repl_state -> Prims.int) = +let rec (go : FStar_Interactive_ReplState.repl_state -> Prims.int) = fun st -> - let query1 = - read_interactive_query st.FStar_Interactive_JsonHelper.repl_stdin in - let uu___ = validate_and_run_query st query1 in - match uu___ with - | ((status, response), state_opt) -> - (write_response query1.qid status response; - (match state_opt with - | FStar_Pervasives.Inl st' -> go st' - | FStar_Pervasives.Inr exitcode -> exitcode)) + (let uu___1 = FStar_Options.debug_any () in + if uu___1 + then + let uu___2 = + let uu___3 = + FStar_Compiler_Effect.op_Bang + FStar_Interactive_PushHelper.repl_stack in + FStar_Interactive_ReplState.string_of_repl_stack uu___3 in + FStar_Compiler_Util.print1 "Repl stack is:\n%s\n" uu___2 + else ()); + (let query = + read_interactive_query st.FStar_Interactive_ReplState.repl_stdin in + let uu___1 = validate_and_run_query st query in + match uu___1 with + | ((status, responses), state_opt) -> + (FStar_Compiler_List.iter + (write_response query.FStar_Interactive_Ide_Types.qid status) + responses; + (match state_opt with + | FStar_Pervasives.Inl st' -> go st' + | FStar_Pervasives.Inr exitcode -> exitcode))) let (interactive_error_handler : FStar_Errors.error_handler) = let issues = FStar_Compiler_Util.mk_ref [] in let add_one e = @@ -2627,29 +2508,27 @@ let (install_ide_mode_hooks : (FStar_Compiler_Util.json -> unit) -> unit) = fun printer -> FStar_Compiler_Util.set_printer (interactive_printer printer); FStar_Errors.set_handler interactive_error_handler -let (initial_range : FStar_Compiler_Range.range) = - let uu___ = FStar_Compiler_Range.mk_pos Prims.int_one Prims.int_zero in - let uu___1 = FStar_Compiler_Range.mk_pos Prims.int_one Prims.int_zero in - FStar_Compiler_Range.mk_range "" uu___ uu___1 let (build_initial_repl_state : - Prims.string -> FStar_Interactive_JsonHelper.repl_state) = + Prims.string -> FStar_Interactive_ReplState.repl_state) = fun filename -> let env = FStar_Universal.init_env FStar_Parser_Dep.empty_deps in - let env1 = FStar_TypeChecker_Env.set_range env initial_range in + let env1 = + FStar_TypeChecker_Env.set_range env + FStar_Interactive_Ide_Types.initial_range in let uu___ = FStar_Compiler_Util.open_stdin () in { - FStar_Interactive_JsonHelper.repl_line = Prims.int_one; - FStar_Interactive_JsonHelper.repl_column = Prims.int_zero; - FStar_Interactive_JsonHelper.repl_fname = filename; - FStar_Interactive_JsonHelper.repl_deps_stack = []; - FStar_Interactive_JsonHelper.repl_curmod = FStar_Pervasives_Native.None; - FStar_Interactive_JsonHelper.repl_env = env1; - FStar_Interactive_JsonHelper.repl_stdin = uu___; - FStar_Interactive_JsonHelper.repl_names = + FStar_Interactive_ReplState.repl_line = Prims.int_one; + FStar_Interactive_ReplState.repl_column = Prims.int_zero; + FStar_Interactive_ReplState.repl_fname = filename; + FStar_Interactive_ReplState.repl_deps_stack = []; + FStar_Interactive_ReplState.repl_curmod = FStar_Pervasives_Native.None; + FStar_Interactive_ReplState.repl_env = env1; + FStar_Interactive_ReplState.repl_stdin = uu___; + FStar_Interactive_ReplState.repl_names = FStar_Interactive_CompletionTable.empty } let interactive_mode' : - 'uuuuu . FStar_Interactive_JsonHelper.repl_state -> 'uuuuu = + 'uuuuu . FStar_Interactive_ReplState.repl_state -> 'uuuuu = fun init_st -> write_hello (); (let exit_code = diff --git a/ocaml/fstar-lib/generated/FStar_Interactive_Ide_Types.ml b/ocaml/fstar-lib/generated/FStar_Interactive_Ide_Types.ml new file mode 100755 index 00000000000..c28b680d470 --- /dev/null +++ b/ocaml/fstar-lib/generated/FStar_Interactive_Ide_Types.ml @@ -0,0 +1,379 @@ +open Prims +let (initial_range : FStar_Compiler_Range.range) = + let uu___ = FStar_Compiler_Range.mk_pos Prims.int_one Prims.int_zero in + let uu___1 = FStar_Compiler_Range.mk_pos Prims.int_one Prims.int_zero in + FStar_Compiler_Range.mk_range "" uu___ uu___1 +let (js_pushkind : + FStar_Compiler_Util.json -> FStar_Interactive_PushHelper.push_kind) = + fun s -> + let uu___ = FStar_Interactive_JsonHelper.js_str s in + match uu___ with + | "syntax" -> FStar_Interactive_PushHelper.SyntaxCheck + | "lax" -> FStar_Interactive_PushHelper.LaxCheck + | "full" -> FStar_Interactive_PushHelper.FullCheck + | uu___1 -> FStar_Interactive_JsonHelper.js_fail "push_kind" s +let (js_reductionrule : + FStar_Compiler_Util.json -> FStar_TypeChecker_Env.step) = + fun s -> + let uu___ = FStar_Interactive_JsonHelper.js_str s in + match uu___ with + | "beta" -> FStar_TypeChecker_Env.Beta + | "delta" -> + FStar_TypeChecker_Env.UnfoldUntil FStar_Syntax_Syntax.delta_constant + | "iota" -> FStar_TypeChecker_Env.Iota + | "zeta" -> FStar_TypeChecker_Env.Zeta + | "reify" -> FStar_TypeChecker_Env.Reify + | "pure-subterms" -> FStar_TypeChecker_Env.PureSubtermsWithinComputations + | uu___1 -> FStar_Interactive_JsonHelper.js_fail "reduction rule" s +type completion_context = + | CKCode + | CKOption of Prims.bool + | CKModuleOrNamespace of (Prims.bool * Prims.bool) +let (uu___is_CKCode : completion_context -> Prims.bool) = + fun projectee -> match projectee with | CKCode -> true | uu___ -> false +let (uu___is_CKOption : completion_context -> Prims.bool) = + fun projectee -> + match projectee with | CKOption _0 -> true | uu___ -> false +let (__proj__CKOption__item___0 : completion_context -> Prims.bool) = + fun projectee -> match projectee with | CKOption _0 -> _0 +let (uu___is_CKModuleOrNamespace : completion_context -> Prims.bool) = + fun projectee -> + match projectee with | CKModuleOrNamespace _0 -> true | uu___ -> false +let (__proj__CKModuleOrNamespace__item___0 : + completion_context -> (Prims.bool * Prims.bool)) = + fun projectee -> match projectee with | CKModuleOrNamespace _0 -> _0 +let (js_optional_completion_context : + FStar_Compiler_Util.json FStar_Pervasives_Native.option -> + completion_context) + = + fun k -> + match k with + | FStar_Pervasives_Native.None -> CKCode + | FStar_Pervasives_Native.Some k1 -> + let uu___ = FStar_Interactive_JsonHelper.js_str k1 in + (match uu___ with + | "symbol" -> CKCode + | "code" -> CKCode + | "set-options" -> CKOption false + | "reset-options" -> CKOption true + | "open" -> CKModuleOrNamespace (true, true) + | "let-open" -> CKModuleOrNamespace (true, true) + | "include" -> CKModuleOrNamespace (true, false) + | "module-alias" -> CKModuleOrNamespace (true, false) + | uu___1 -> + FStar_Interactive_JsonHelper.js_fail + "completion context (code, set-options, reset-options, open, let-open, include, module-alias)" + k1) +type lookup_context = + | LKSymbolOnly + | LKModule + | LKOption + | LKCode +let (uu___is_LKSymbolOnly : lookup_context -> Prims.bool) = + fun projectee -> + match projectee with | LKSymbolOnly -> true | uu___ -> false +let (uu___is_LKModule : lookup_context -> Prims.bool) = + fun projectee -> match projectee with | LKModule -> true | uu___ -> false +let (uu___is_LKOption : lookup_context -> Prims.bool) = + fun projectee -> match projectee with | LKOption -> true | uu___ -> false +let (uu___is_LKCode : lookup_context -> Prims.bool) = + fun projectee -> match projectee with | LKCode -> true | uu___ -> false +let (js_optional_lookup_context : + FStar_Compiler_Util.json FStar_Pervasives_Native.option -> lookup_context) + = + fun k -> + match k with + | FStar_Pervasives_Native.None -> LKSymbolOnly + | FStar_Pervasives_Native.Some k1 -> + let uu___ = FStar_Interactive_JsonHelper.js_str k1 in + (match uu___ with + | "symbol-only" -> LKSymbolOnly + | "code" -> LKCode + | "set-options" -> LKOption + | "reset-options" -> LKOption + | "open" -> LKModule + | "let-open" -> LKModule + | "include" -> LKModule + | "module-alias" -> LKModule + | uu___1 -> + FStar_Interactive_JsonHelper.js_fail + "lookup context (symbol-only, code, set-options, reset-options, open, let-open, include, module-alias)" + k1) +type position = (Prims.string * Prims.int * Prims.int) +type push_query = + { + push_kind: FStar_Interactive_PushHelper.push_kind ; + push_line: Prims.int ; + push_column: Prims.int ; + push_peek_only: Prims.bool ; + push_code_or_decl: + (Prims.string, FStar_Parser_AST.decl) FStar_Pervasives.either } +let (__proj__Mkpush_query__item__push_kind : + push_query -> FStar_Interactive_PushHelper.push_kind) = + fun projectee -> + match projectee with + | { push_kind; push_line; push_column; push_peek_only; + push_code_or_decl;_} -> push_kind +let (__proj__Mkpush_query__item__push_line : push_query -> Prims.int) = + fun projectee -> + match projectee with + | { push_kind; push_line; push_column; push_peek_only; + push_code_or_decl;_} -> push_line +let (__proj__Mkpush_query__item__push_column : push_query -> Prims.int) = + fun projectee -> + match projectee with + | { push_kind; push_line; push_column; push_peek_only; + push_code_or_decl;_} -> push_column +let (__proj__Mkpush_query__item__push_peek_only : push_query -> Prims.bool) = + fun projectee -> + match projectee with + | { push_kind; push_line; push_column; push_peek_only; + push_code_or_decl;_} -> push_peek_only +let (__proj__Mkpush_query__item__push_code_or_decl : + push_query -> (Prims.string, FStar_Parser_AST.decl) FStar_Pervasives.either) + = + fun projectee -> + match projectee with + | { push_kind; push_line; push_column; push_peek_only; + push_code_or_decl;_} -> push_code_or_decl +type query' = + | Exit + | DescribeProtocol + | DescribeRepl + | Segment of Prims.string + | Pop + | Push of push_query + | VfsAdd of (Prims.string FStar_Pervasives_Native.option * Prims.string) + | AutoComplete of (Prims.string * completion_context) + | Lookup of (Prims.string * lookup_context * position + FStar_Pervasives_Native.option * Prims.string Prims.list) + | Compute of (Prims.string * FStar_TypeChecker_Env.step Prims.list + FStar_Pervasives_Native.option) + | Search of Prims.string + | GenericError of Prims.string + | ProtocolViolation of Prims.string + | FullBuffer of Prims.string +and query = { + qq: query' ; + qid: Prims.string } +let (uu___is_Exit : query' -> Prims.bool) = + fun projectee -> match projectee with | Exit -> true | uu___ -> false +let (uu___is_DescribeProtocol : query' -> Prims.bool) = + fun projectee -> + match projectee with | DescribeProtocol -> true | uu___ -> false +let (uu___is_DescribeRepl : query' -> Prims.bool) = + fun projectee -> + match projectee with | DescribeRepl -> true | uu___ -> false +let (uu___is_Segment : query' -> Prims.bool) = + fun projectee -> match projectee with | Segment _0 -> true | uu___ -> false +let (__proj__Segment__item___0 : query' -> Prims.string) = + fun projectee -> match projectee with | Segment _0 -> _0 +let (uu___is_Pop : query' -> Prims.bool) = + fun projectee -> match projectee with | Pop -> true | uu___ -> false +let (uu___is_Push : query' -> Prims.bool) = + fun projectee -> match projectee with | Push _0 -> true | uu___ -> false +let (__proj__Push__item___0 : query' -> push_query) = + fun projectee -> match projectee with | Push _0 -> _0 +let (uu___is_VfsAdd : query' -> Prims.bool) = + fun projectee -> match projectee with | VfsAdd _0 -> true | uu___ -> false +let (__proj__VfsAdd__item___0 : + query' -> (Prims.string FStar_Pervasives_Native.option * Prims.string)) = + fun projectee -> match projectee with | VfsAdd _0 -> _0 +let (uu___is_AutoComplete : query' -> Prims.bool) = + fun projectee -> + match projectee with | AutoComplete _0 -> true | uu___ -> false +let (__proj__AutoComplete__item___0 : + query' -> (Prims.string * completion_context)) = + fun projectee -> match projectee with | AutoComplete _0 -> _0 +let (uu___is_Lookup : query' -> Prims.bool) = + fun projectee -> match projectee with | Lookup _0 -> true | uu___ -> false +let (__proj__Lookup__item___0 : + query' -> + (Prims.string * lookup_context * position FStar_Pervasives_Native.option + * Prims.string Prims.list)) + = fun projectee -> match projectee with | Lookup _0 -> _0 +let (uu___is_Compute : query' -> Prims.bool) = + fun projectee -> match projectee with | Compute _0 -> true | uu___ -> false +let (__proj__Compute__item___0 : + query' -> + (Prims.string * FStar_TypeChecker_Env.step Prims.list + FStar_Pervasives_Native.option)) + = fun projectee -> match projectee with | Compute _0 -> _0 +let (uu___is_Search : query' -> Prims.bool) = + fun projectee -> match projectee with | Search _0 -> true | uu___ -> false +let (__proj__Search__item___0 : query' -> Prims.string) = + fun projectee -> match projectee with | Search _0 -> _0 +let (uu___is_GenericError : query' -> Prims.bool) = + fun projectee -> + match projectee with | GenericError _0 -> true | uu___ -> false +let (__proj__GenericError__item___0 : query' -> Prims.string) = + fun projectee -> match projectee with | GenericError _0 -> _0 +let (uu___is_ProtocolViolation : query' -> Prims.bool) = + fun projectee -> + match projectee with | ProtocolViolation _0 -> true | uu___ -> false +let (__proj__ProtocolViolation__item___0 : query' -> Prims.string) = + fun projectee -> match projectee with | ProtocolViolation _0 -> _0 +let (uu___is_FullBuffer : query' -> Prims.bool) = + fun projectee -> + match projectee with | FullBuffer _0 -> true | uu___ -> false +let (__proj__FullBuffer__item___0 : query' -> Prims.string) = + fun projectee -> match projectee with | FullBuffer _0 -> _0 +let (__proj__Mkquery__item__qq : query -> query') = + fun projectee -> match projectee with | { qq; qid;_} -> qq +let (__proj__Mkquery__item__qid : query -> Prims.string) = + fun projectee -> match projectee with | { qq; qid;_} -> qid +let (push_query_to_string : push_query -> Prims.string) = + fun pq -> + let pk = + match pq.push_kind with + | FStar_Interactive_PushHelper.SyntaxCheck -> "SyntaxCheck" + | FStar_Interactive_PushHelper.LaxCheck -> "LaxCheck" + | FStar_Interactive_PushHelper.FullCheck -> "FullCheck" in + let code_or_decl = + match pq.push_code_or_decl with + | FStar_Pervasives.Inl code -> code + | FStar_Pervasives.Inr decl -> FStar_Parser_AST.decl_to_string decl in + let uu___ = + let uu___1 = + let uu___2 = FStar_Compiler_Util.string_of_int pq.push_line in + let uu___3 = + let uu___4 = FStar_Compiler_Util.string_of_int pq.push_column in + let uu___5 = + let uu___6 = FStar_Compiler_Util.string_of_bool pq.push_peek_only in + [uu___6; code_or_decl] in + uu___4 :: uu___5 in + uu___2 :: uu___3 in + pk :: uu___1 in + FStar_Compiler_Util.format + "{ push_kind = %s; push_line = %s; push_column = %s; push_peek_only = %s; push_code_or_decl = %s }" + uu___ +let (query_to_string : query -> Prims.string) = + fun q -> + match q.qq with + | Exit -> "Exit" + | DescribeProtocol -> "DescribeProtocol" + | DescribeRepl -> "DescribeRepl" + | Segment uu___ -> "Segment" + | Pop -> "Pop" + | Push pq -> + let uu___ = + let uu___1 = push_query_to_string pq in Prims.op_Hat uu___1 ")" in + Prims.op_Hat "(Push " uu___ + | VfsAdd uu___ -> "VfsAdd" + | AutoComplete uu___ -> "AutoComplete" + | Lookup uu___ -> "Lookup" + | Compute uu___ -> "Compute" + | Search uu___ -> "Search" + | GenericError uu___ -> "GenericError" + | ProtocolViolation uu___ -> "ProtocolViolation" + | FullBuffer uu___ -> "FullBuffer" +let (query_needs_current_module : query' -> Prims.bool) = + fun uu___ -> + match uu___ with + | Exit -> false + | DescribeProtocol -> false + | DescribeRepl -> false + | Segment uu___1 -> false + | Pop -> false + | Push + { push_kind = uu___1; push_line = uu___2; push_column = uu___3; + push_peek_only = false; push_code_or_decl = uu___4;_} + -> false + | VfsAdd uu___1 -> false + | GenericError uu___1 -> false + | ProtocolViolation uu___1 -> false + | FullBuffer uu___1 -> false + | Push uu___1 -> true + | AutoComplete uu___1 -> true + | Lookup uu___1 -> true + | Compute uu___1 -> true + | Search uu___1 -> true +let (interactive_protocol_vernum : Prims.int) = (Prims.of_int (2)) +let (interactive_protocol_features : Prims.string Prims.list) = + ["autocomplete"; + "autocomplete/context"; + "compute"; + "compute/reify"; + "compute/pure-subterms"; + "describe-protocol"; + "describe-repl"; + "exit"; + "lookup"; + "lookup/context"; + "lookup/documentation"; + "lookup/definition"; + "peek"; + "pop"; + "push"; + "search"; + "segment"; + "vfs-add"; + "tactic-ranges"; + "interrupt"; + "progress"; + "full-buffer"] +type query_status = + | QueryOK + | QueryNOK + | QueryViolatesProtocol +let (uu___is_QueryOK : query_status -> Prims.bool) = + fun projectee -> match projectee with | QueryOK -> true | uu___ -> false +let (uu___is_QueryNOK : query_status -> Prims.bool) = + fun projectee -> match projectee with | QueryNOK -> true | uu___ -> false +let (uu___is_QueryViolatesProtocol : query_status -> Prims.bool) = + fun projectee -> + match projectee with | QueryViolatesProtocol -> true | uu___ -> false +let (json_of_issue_level : + FStar_Errors.issue_level -> FStar_Compiler_Util.json) = + fun i -> + FStar_Compiler_Util.JsonStr + (match i with + | FStar_Errors.ENotImplemented -> "not-implemented" + | FStar_Errors.EInfo -> "info" + | FStar_Errors.EWarning -> "warning" + | FStar_Errors.EError -> "error") +let (json_of_issue : FStar_Errors.issue -> FStar_Compiler_Util.json) = + fun issue -> + let uu___ = + let uu___1 = + let uu___2 = + let uu___3 = + let uu___4 = + let uu___5 = FStar_Errors.issue_message issue in + FStar_Compiler_Util.JsonStr uu___5 in + ("message", uu___4) in + let uu___4 = + let uu___5 = + let uu___6 = + let uu___7 = + let uu___8 = + match issue.FStar_Errors.issue_range with + | FStar_Pervasives_Native.None -> [] + | FStar_Pervasives_Native.Some r -> + let uu___9 = FStar_Compiler_Range.json_of_use_range r in + [uu___9] in + let uu___9 = + match issue.FStar_Errors.issue_range with + | FStar_Pervasives_Native.Some r when + let uu___10 = FStar_Compiler_Range.def_range r in + let uu___11 = FStar_Compiler_Range.use_range r in + uu___10 <> uu___11 -> + let uu___10 = + FStar_Compiler_Range.json_of_def_range r in + [uu___10] + | uu___10 -> [] in + FStar_Compiler_List.op_At uu___8 uu___9 in + FStar_Compiler_Util.JsonList uu___7 in + ("ranges", uu___6) in + [uu___5] in + uu___3 :: uu___4 in + FStar_Compiler_List.op_At + (match issue.FStar_Errors.issue_number with + | FStar_Pervasives_Native.None -> [] + | FStar_Pervasives_Native.Some n -> + [("number", (FStar_Compiler_Util.JsonInt n))]) uu___2 in + FStar_Compiler_List.op_At + [("level", (json_of_issue_level issue.FStar_Errors.issue_level))] + uu___1 in + FStar_Compiler_Effect.op_Less_Bar + (fun uu___1 -> FStar_Compiler_Util.JsonAssoc uu___1) uu___ \ No newline at end of file diff --git a/ocaml/fstar-lib/generated/FStar_Interactive_Incremental.ml b/ocaml/fstar-lib/generated/FStar_Interactive_Incremental.ml new file mode 100755 index 00000000000..171a63d9b64 --- /dev/null +++ b/ocaml/fstar-lib/generated/FStar_Interactive_Incremental.ml @@ -0,0 +1,256 @@ +open Prims +type qid = (Prims.string * Prims.int) +type 'a qst = qid -> ('a * qid) +let return : 'a . 'a -> 'a qst = fun x -> fun q -> (x, q) +let op_let_Bang : 'a 'b . 'a qst -> ('a -> 'b qst) -> 'b qst = + fun f -> + fun g -> + fun q -> + let uu___ = f q in + match uu___ with | (x, q') -> let uu___1 = g x in uu___1 q' +let run_qst : 'a . 'a qst -> Prims.string -> 'a = + fun f -> + fun q -> + let uu___ = f (q, Prims.int_zero) in FStar_Pervasives_Native.fst uu___ +let rec map : 'a 'b . ('a -> 'b qst) -> 'a Prims.list -> 'b Prims.list qst = + fun f -> + fun l -> + match l with + | [] -> return [] + | hd::tl -> + let uu___ = f hd in + op_let_Bang uu___ + (fun hd1 -> + let uu___1 = map f tl in + op_let_Bang uu___1 (fun tl1 -> return (hd1 :: tl1))) +let (shift_qid : qid -> Prims.int -> (Prims.string * Prims.int)) = + fun q -> + fun i -> + ((FStar_Pervasives_Native.fst q), + ((FStar_Pervasives_Native.snd q) + i)) +let (next_qid : qid qst) = + fun q -> let q1 = shift_qid q Prims.int_one in (q1, q1) +let (get_qid : qid qst) = fun q -> (q, q) +let (as_query : + FStar_Interactive_Ide_Types.query' -> FStar_Interactive_Ide_Types.query qst) + = + fun q -> + op_let_Bang next_qid + (fun uu___ -> + match uu___ with + | (qid_prefix, i) -> + let uu___1 = + let uu___2 = + let uu___3 = + let uu___4 = FStar_Compiler_Util.string_of_int i in + Prims.op_Hat "." uu___4 in + Prims.op_Hat qid_prefix uu___3 in + { + FStar_Interactive_Ide_Types.qq = q; + FStar_Interactive_Ide_Types.qid = uu___2 + } in + return uu___1) +let (push_decl : + FStar_Parser_AST.decl -> FStar_Interactive_Ide_Types.query qst) = + fun d -> + let pq = + let uu___ = + let uu___1 = + FStar_Compiler_Range.start_of_range d.FStar_Parser_AST.drange in + FStar_Compiler_Range.line_of_pos uu___1 in + let uu___1 = + let uu___2 = + FStar_Compiler_Range.start_of_range d.FStar_Parser_AST.drange in + FStar_Compiler_Range.col_of_pos uu___2 in + { + FStar_Interactive_Ide_Types.push_kind = + FStar_Interactive_PushHelper.FullCheck; + FStar_Interactive_Ide_Types.push_line = uu___; + FStar_Interactive_Ide_Types.push_column = uu___1; + FStar_Interactive_Ide_Types.push_peek_only = false; + FStar_Interactive_Ide_Types.push_code_or_decl = + (FStar_Pervasives.Inr d) + } in + as_query (FStar_Interactive_Ide_Types.Push pq) +let (push_decls : + FStar_Parser_AST.decl Prims.list -> + FStar_Interactive_Ide_Types.query Prims.list qst) + = fun ds -> map push_decl ds +let (pop_entries : + FStar_Interactive_ReplState.repl_stack_entry_t Prims.list -> + FStar_Interactive_Ide_Types.query Prims.list qst) + = fun e -> map (fun uu___ -> as_query FStar_Interactive_Ide_Types.Pop) e +let (response_success : + FStar_Parser_AST.decl -> FStar_Compiler_Util.json qst) = + fun d -> + op_let_Bang get_qid + (fun uu___ -> + match uu___ with + | (q, uu___1) -> + let contents = + let uu___2 = + let uu___3 = + let uu___4 = + FStar_Compiler_Range.json_of_def_range + d.FStar_Parser_AST.drange in + ("ranges", uu___4) in + [uu___3] in + FStar_Compiler_Util.JsonAssoc uu___2 in + return + (FStar_Compiler_Util.JsonAssoc + [("kind", (FStar_Compiler_Util.JsonStr "response")); + ("query-id", (FStar_Compiler_Util.JsonStr q)); + ("status", (FStar_Compiler_Util.JsonStr "success")); + ("contents", contents)])) +let (inspect_repl_stack : + FStar_Interactive_ReplState.repl_stack_t -> + FStar_Parser_AST.decl Prims.list -> + ((FStar_Parser_AST.decl, FStar_Errors.issue Prims.list) + FStar_Pervasives.either -> unit) + -> FStar_Interactive_Ide_Types.query Prims.list qst) + = + fun s -> + fun ds -> + fun write_full_buffer_fragment_progress -> + let entries = FStar_Compiler_List.rev s in + let uu___ = + FStar_Compiler_Util.prefix_until + (fun uu___1 -> + match uu___1 with + | (uu___2, + (FStar_Interactive_ReplState.PushFragment uu___3, uu___4)) + -> true + | uu___2 -> false) entries in + match uu___ with + | FStar_Pervasives_Native.None -> + let uu___1 = push_decls ds in + op_let_Bang uu___1 (fun ds1 -> return ds1) + | FStar_Pervasives_Native.Some (prefix, first_push, rest) -> + let entries1 = first_push :: rest in + let repl_task uu___1 = + match uu___1 with | (uu___2, (p, uu___3)) -> p in + let rec matching_prefix entries2 ds1 = + match (entries2, ds1) with + | ([], []) -> return [] + | (e::entries3, d::ds2) -> + (match repl_task e with + | FStar_Interactive_ReplState.Noop -> + matching_prefix entries3 (d :: ds2) + | FStar_Interactive_ReplState.PushFragment + (FStar_Pervasives.Inl frag) -> + let uu___1 = pop_entries (e :: entries3) in + op_let_Bang uu___1 + (fun pops -> + let uu___2 = push_decls (d :: ds2) in + op_let_Bang uu___2 + (fun pushes -> + return + (FStar_Compiler_List.op_At pops pushes))) + | FStar_Interactive_ReplState.PushFragment + (FStar_Pervasives.Inr d') -> + let uu___1 = FStar_Parser_AST_Comparison.eq_decl d d' in + if uu___1 + then + (write_full_buffer_fragment_progress + (FStar_Pervasives.Inl d); + matching_prefix entries3 ds2) + else + (let uu___3 = pop_entries (e :: entries3) in + op_let_Bang uu___3 + (fun pops -> + let uu___4 = push_decls (d :: ds2) in + op_let_Bang uu___4 + (fun pushes -> + return + (FStar_Compiler_List.op_At pops pushes))))) + | ([], ds2) -> + let uu___1 = push_decls ds2 in + op_let_Bang uu___1 (fun pushes -> return pushes) + | (es, []) -> + let uu___1 = pop_entries es in + op_let_Bang uu___1 (fun pops -> return pops) in + matching_prefix entries1 ds +let (run_full_buffer : + FStar_Interactive_ReplState.repl_state -> + Prims.string -> + Prims.string -> + ((FStar_Parser_AST.decl, FStar_Errors.issue Prims.list) + FStar_Pervasives.either -> unit) + -> FStar_Interactive_Ide_Types.query Prims.list) + = + fun st -> + fun qid1 -> + fun code -> + fun write_full_buffer_fragment_progress -> + let parse_result = + let uu___ = + let uu___1 = + let uu___2 = + FStar_Compiler_Range.file_of_range + FStar_Interactive_Ide_Types.initial_range in + let uu___3 = + let uu___4 = + FStar_Compiler_Range.start_of_range + FStar_Interactive_Ide_Types.initial_range in + FStar_Compiler_Range.line_of_pos uu___4 in + let uu___4 = + let uu___5 = + FStar_Compiler_Range.start_of_range + FStar_Interactive_Ide_Types.initial_range in + FStar_Compiler_Range.col_of_pos uu___5 in + { + FStar_Parser_ParseIt.frag_fname = uu___2; + FStar_Parser_ParseIt.frag_text = code; + FStar_Parser_ParseIt.frag_line = uu___3; + FStar_Parser_ParseIt.frag_col = uu___4 + } in + FStar_Parser_ParseIt.Incremental uu___1 in + FStar_Parser_ParseIt.parse uu___ in + let log_syntax_issues err = + match err with + | FStar_Pervasives_Native.None -> () + | FStar_Pervasives_Native.Some (raw_error, msg, range) -> + let uu___ = FStar_Errors.lookup raw_error in + (match uu___ with + | (uu___1, uu___2, num) -> + let issue = + { + FStar_Errors.issue_msg = msg; + FStar_Errors.issue_level = FStar_Errors.EError; + FStar_Errors.issue_range = + (FStar_Pervasives_Native.Some range); + FStar_Errors.issue_number = + (FStar_Pervasives_Native.Some num); + FStar_Errors.issue_ctx = [] + } in + write_full_buffer_fragment_progress + (FStar_Pervasives.Inr [issue])) in + let qs = + match parse_result with + | FStar_Parser_ParseIt.IncrementalFragment + (decls, uu___, err_opt) -> + let queries = + let uu___1 = + let uu___2 = + FStar_Compiler_Effect.op_Bang + FStar_Interactive_PushHelper.repl_stack in + inspect_repl_stack uu___2 decls + write_full_buffer_fragment_progress in + run_qst uu___1 qid1 in + (log_syntax_issues err_opt; + (let uu___3 = FStar_Options.debug_any () in + if uu___3 + then + let uu___4 = + let uu___5 = + FStar_Compiler_List.map + FStar_Interactive_Ide_Types.query_to_string queries in + FStar_String.concat "\n" uu___5 in + FStar_Compiler_Util.print1 "Generating queries\n%s\n" + uu___4 + else ()); + queries) + | FStar_Parser_ParseIt.ParseError err -> + (log_syntax_issues (FStar_Pervasives_Native.Some err); []) + | uu___ -> failwith "Unexpected parse result" in + qs \ No newline at end of file diff --git a/ocaml/fstar-lib/generated/FStar_Interactive_JsonHelper.ml b/ocaml/fstar-lib/generated/FStar_Interactive_JsonHelper.ml index 4a81467f69f..6ccd25b8b1a 100644 --- a/ocaml/fstar-lib/generated/FStar_Interactive_JsonHelper.ml +++ b/ocaml/fstar-lib/generated/FStar_Interactive_JsonHelper.ml @@ -486,119 +486,6 @@ let (__proj__Mklsp_query__item__query_id : fun projectee -> match projectee with | { query_id; q;_} -> query_id let (__proj__Mklsp_query__item__q : lsp_query -> lquery) = fun projectee -> match projectee with | { query_id; q;_} -> q -type repl_depth_t = (FStar_TypeChecker_Env.tcenv_depth_t * Prims.int) -type optmod_t = FStar_Syntax_Syntax.modul FStar_Pervasives_Native.option -type timed_fname = - { - tf_fname: Prims.string ; - tf_modtime: FStar_Compiler_Util.time } -let (__proj__Mktimed_fname__item__tf_fname : timed_fname -> Prims.string) = - fun projectee -> - match projectee with | { tf_fname; tf_modtime;_} -> tf_fname -let (__proj__Mktimed_fname__item__tf_modtime : - timed_fname -> FStar_Compiler_Util.time) = - fun projectee -> - match projectee with | { tf_fname; tf_modtime;_} -> tf_modtime -type repl_task = - | LDInterleaved of (timed_fname * timed_fname) - | LDSingle of timed_fname - | LDInterfaceOfCurrentFile of timed_fname - | PushFragment of FStar_Parser_ParseIt.input_frag - | Noop -let (uu___is_LDInterleaved : repl_task -> Prims.bool) = - fun projectee -> - match projectee with | LDInterleaved _0 -> true | uu___ -> false -let (__proj__LDInterleaved__item___0 : - repl_task -> (timed_fname * timed_fname)) = - fun projectee -> match projectee with | LDInterleaved _0 -> _0 -let (uu___is_LDSingle : repl_task -> Prims.bool) = - fun projectee -> - match projectee with | LDSingle _0 -> true | uu___ -> false -let (__proj__LDSingle__item___0 : repl_task -> timed_fname) = - fun projectee -> match projectee with | LDSingle _0 -> _0 -let (uu___is_LDInterfaceOfCurrentFile : repl_task -> Prims.bool) = - fun projectee -> - match projectee with - | LDInterfaceOfCurrentFile _0 -> true - | uu___ -> false -let (__proj__LDInterfaceOfCurrentFile__item___0 : repl_task -> timed_fname) = - fun projectee -> match projectee with | LDInterfaceOfCurrentFile _0 -> _0 -let (uu___is_PushFragment : repl_task -> Prims.bool) = - fun projectee -> - match projectee with | PushFragment _0 -> true | uu___ -> false -let (__proj__PushFragment__item___0 : - repl_task -> FStar_Parser_ParseIt.input_frag) = - fun projectee -> match projectee with | PushFragment _0 -> _0 -let (uu___is_Noop : repl_task -> Prims.bool) = - fun projectee -> match projectee with | Noop -> true | uu___ -> false -type repl_state = - { - repl_line: Prims.int ; - repl_column: Prims.int ; - repl_fname: Prims.string ; - repl_deps_stack: (repl_depth_t * (repl_task * repl_state)) Prims.list ; - repl_curmod: optmod_t ; - repl_env: FStar_TypeChecker_Env.env ; - repl_stdin: FStar_Compiler_Util.stream_reader ; - repl_names: FStar_Interactive_CompletionTable.table } -let (__proj__Mkrepl_state__item__repl_line : repl_state -> Prims.int) = - fun projectee -> - match projectee with - | { repl_line; repl_column; repl_fname; repl_deps_stack; repl_curmod; - repl_env; repl_stdin; repl_names;_} -> repl_line -let (__proj__Mkrepl_state__item__repl_column : repl_state -> Prims.int) = - fun projectee -> - match projectee with - | { repl_line; repl_column; repl_fname; repl_deps_stack; repl_curmod; - repl_env; repl_stdin; repl_names;_} -> repl_column -let (__proj__Mkrepl_state__item__repl_fname : repl_state -> Prims.string) = - fun projectee -> - match projectee with - | { repl_line; repl_column; repl_fname; repl_deps_stack; repl_curmod; - repl_env; repl_stdin; repl_names;_} -> repl_fname -let (__proj__Mkrepl_state__item__repl_deps_stack : - repl_state -> (repl_depth_t * (repl_task * repl_state)) Prims.list) = - fun projectee -> - match projectee with - | { repl_line; repl_column; repl_fname; repl_deps_stack; repl_curmod; - repl_env; repl_stdin; repl_names;_} -> repl_deps_stack -let (__proj__Mkrepl_state__item__repl_curmod : repl_state -> optmod_t) = - fun projectee -> - match projectee with - | { repl_line; repl_column; repl_fname; repl_deps_stack; repl_curmod; - repl_env; repl_stdin; repl_names;_} -> repl_curmod -let (__proj__Mkrepl_state__item__repl_env : - repl_state -> FStar_TypeChecker_Env.env) = - fun projectee -> - match projectee with - | { repl_line; repl_column; repl_fname; repl_deps_stack; repl_curmod; - repl_env; repl_stdin; repl_names;_} -> repl_env -let (__proj__Mkrepl_state__item__repl_stdin : - repl_state -> FStar_Compiler_Util.stream_reader) = - fun projectee -> - match projectee with - | { repl_line; repl_column; repl_fname; repl_deps_stack; repl_curmod; - repl_env; repl_stdin; repl_names;_} -> repl_stdin -let (__proj__Mkrepl_state__item__repl_names : - repl_state -> FStar_Interactive_CompletionTable.table) = - fun projectee -> - match projectee with - | { repl_line; repl_column; repl_fname; repl_deps_stack; repl_curmod; - repl_env; repl_stdin; repl_names;_} -> repl_names -type repl_stack_entry_t = (repl_depth_t * (repl_task * repl_state)) -type repl_stack_t = (repl_depth_t * (repl_task * repl_state)) Prims.list -type grepl_state = - { - grepl_repls: repl_state FStar_Compiler_Util.psmap ; - grepl_stdin: FStar_Compiler_Util.stream_reader } -let (__proj__Mkgrepl_state__item__grepl_repls : - grepl_state -> repl_state FStar_Compiler_Util.psmap) = - fun projectee -> - match projectee with | { grepl_repls; grepl_stdin;_} -> grepl_repls -let (__proj__Mkgrepl_state__item__grepl_stdin : - grepl_state -> FStar_Compiler_Util.stream_reader) = - fun projectee -> - match projectee with | { grepl_repls; grepl_stdin;_} -> grepl_stdin type error_code = | ParseError | InvalidRequest diff --git a/ocaml/fstar-lib/generated/FStar_Interactive_Legacy.ml b/ocaml/fstar-lib/generated/FStar_Interactive_Legacy.ml index 60466c32c1d..b1db4759bf7 100644 --- a/ocaml/fstar-lib/generated/FStar_Interactive_Legacy.ml +++ b/ocaml/fstar-lib/generated/FStar_Interactive_Legacy.ml @@ -167,7 +167,9 @@ let (check_frag : (fun uu___ -> match () with | () -> - let uu___1 = FStar_Universal.tc_one_fragment curmod env frag in + let uu___1 = + FStar_Universal.tc_one_fragment curmod env + (FStar_Pervasives.Inl frag) in (match uu___1 with | (m, env1) -> let uu___2 = diff --git a/ocaml/fstar-lib/generated/FStar_Interactive_Lsp.ml b/ocaml/fstar-lib/generated/FStar_Interactive_Lsp.ml index 96bfa0a808f..ab383721160 100644 --- a/ocaml/fstar-lib/generated/FStar_Interactive_Lsp.ml +++ b/ocaml/fstar-lib/generated/FStar_Interactive_Lsp.ml @@ -210,7 +210,7 @@ let (parse_lsp_query : } | FStar_Pervasives_Native.Some request -> deserialize_lsp_query request let (repl_state_init : - Prims.string -> FStar_Interactive_JsonHelper.repl_state) = + Prims.string -> FStar_Interactive_ReplState.repl_state) = fun fname -> let intial_range = let uu___ = FStar_Compiler_Range.mk_pos Prims.int_one Prims.int_zero in @@ -220,23 +220,23 @@ let (repl_state_init : let env1 = FStar_TypeChecker_Env.set_range env intial_range in let uu___ = FStar_Compiler_Util.open_stdin () in { - FStar_Interactive_JsonHelper.repl_line = Prims.int_one; - FStar_Interactive_JsonHelper.repl_column = Prims.int_zero; - FStar_Interactive_JsonHelper.repl_fname = fname; - FStar_Interactive_JsonHelper.repl_deps_stack = []; - FStar_Interactive_JsonHelper.repl_curmod = FStar_Pervasives_Native.None; - FStar_Interactive_JsonHelper.repl_env = env1; - FStar_Interactive_JsonHelper.repl_stdin = uu___; - FStar_Interactive_JsonHelper.repl_names = + FStar_Interactive_ReplState.repl_line = Prims.int_one; + FStar_Interactive_ReplState.repl_column = Prims.int_zero; + FStar_Interactive_ReplState.repl_fname = fname; + FStar_Interactive_ReplState.repl_deps_stack = []; + FStar_Interactive_ReplState.repl_curmod = FStar_Pervasives_Native.None; + FStar_Interactive_ReplState.repl_env = env1; + FStar_Interactive_ReplState.repl_stdin = uu___; + FStar_Interactive_ReplState.repl_names = FStar_Interactive_CompletionTable.empty } type optresponse = FStar_Interactive_JsonHelper.assoct FStar_Pervasives_Native.option type either_gst_exit = - (FStar_Interactive_JsonHelper.grepl_state, Prims.int) + (FStar_Interactive_ReplState.grepl_state, Prims.int) FStar_Pervasives.either let (invoke_full_lax : - FStar_Interactive_JsonHelper.grepl_state -> + FStar_Interactive_ReplState.grepl_state -> Prims.string -> Prims.string -> Prims.bool -> (optresponse * either_gst_exit)) = @@ -253,7 +253,7 @@ let (invoke_full_lax : | (diag, st') -> let repls = FStar_Compiler_Util.psmap_add - gst.FStar_Interactive_JsonHelper.grepl_repls fname st' in + gst.FStar_Interactive_ReplState.grepl_repls fname st' in let diag1 = if FStar_Compiler_Util.is_some diag then diag @@ -264,13 +264,13 @@ let (invoke_full_lax : (diag1, (FStar_Pervasives.Inl { - FStar_Interactive_JsonHelper.grepl_repls = repls; - FStar_Interactive_JsonHelper.grepl_stdin = - (gst.FStar_Interactive_JsonHelper.grepl_stdin) + FStar_Interactive_ReplState.grepl_repls = repls; + FStar_Interactive_ReplState.grepl_stdin = + (gst.FStar_Interactive_ReplState.grepl_stdin) }))) in let uu___ = FStar_Compiler_Util.psmap_try_find - gst.FStar_Interactive_JsonHelper.grepl_repls fname in + gst.FStar_Interactive_ReplState.grepl_repls fname in match uu___ with | FStar_Pervasives_Native.Some uu___1 -> if force @@ -278,7 +278,7 @@ let (invoke_full_lax : else (FStar_Pervasives_Native.None, (FStar_Pervasives.Inl gst)) | FStar_Pervasives_Native.None -> aux () let (run_query : - FStar_Interactive_JsonHelper.grepl_state -> + FStar_Interactive_ReplState.grepl_state -> FStar_Interactive_JsonHelper.lquery -> (optresponse * either_gst_exit)) = fun gst -> @@ -334,7 +334,7 @@ let (run_query : | FStar_Interactive_JsonHelper.Completion (txpos, ctx) -> let uu___ = FStar_Compiler_Util.psmap_try_find - gst.FStar_Interactive_JsonHelper.grepl_repls + gst.FStar_Interactive_ReplState.grepl_repls txpos.FStar_Interactive_JsonHelper.path in (match uu___ with | FStar_Pervasives_Native.Some st -> @@ -349,13 +349,13 @@ let (run_query : | FStar_Interactive_JsonHelper.Hover txpos -> let uu___ = FStar_Compiler_Util.psmap_try_find - gst.FStar_Interactive_JsonHelper.grepl_repls + gst.FStar_Interactive_ReplState.grepl_repls txpos.FStar_Interactive_JsonHelper.path in (match uu___ with | FStar_Pervasives_Native.Some st -> let uu___1 = FStar_Interactive_QueryHelper.hoverlookup - st.FStar_Interactive_JsonHelper.repl_env txpos in + st.FStar_Interactive_ReplState.repl_env txpos in (uu___1, (FStar_Pervasives.Inl gst)) | FStar_Pervasives_Native.None -> (FStar_Interactive_JsonHelper.nullResponse, @@ -369,13 +369,13 @@ let (run_query : | FStar_Interactive_JsonHelper.Definition txpos -> let uu___ = FStar_Compiler_Util.psmap_try_find - gst.FStar_Interactive_JsonHelper.grepl_repls + gst.FStar_Interactive_ReplState.grepl_repls txpos.FStar_Interactive_JsonHelper.path in (match uu___ with | FStar_Pervasives_Native.Some st -> let uu___1 = FStar_Interactive_QueryHelper.deflookup - st.FStar_Interactive_JsonHelper.repl_env txpos in + st.FStar_Interactive_ReplState.repl_env txpos in (uu___1, (FStar_Pervasives.Inl gst)) | FStar_Pervasives_Native.None -> (FStar_Interactive_JsonHelper.nullResponse, @@ -495,9 +495,9 @@ let rec (read_lsp_query : (FStar_Compiler_Util.print_error "[E] Malformed Content Header\n"; read_lsp_query stream) | FStar_Interactive_JsonHelper.InputExhausted -> read_lsp_query stream -let rec (go : FStar_Interactive_JsonHelper.grepl_state -> Prims.int) = +let rec (go : FStar_Interactive_ReplState.grepl_state -> Prims.int) = fun gst -> - let query = read_lsp_query gst.FStar_Interactive_JsonHelper.grepl_stdin in + let query = read_lsp_query gst.FStar_Interactive_ReplState.grepl_stdin in let uu___ = run_query gst query.FStar_Interactive_JsonHelper.q in match uu___ with | (r, state_opt) -> @@ -523,8 +523,8 @@ let (start_server : unit -> unit) = let uu___3 = FStar_Compiler_Util.psmap_empty () in let uu___4 = FStar_Compiler_Util.open_stdin () in { - FStar_Interactive_JsonHelper.grepl_repls = uu___3; - FStar_Interactive_JsonHelper.grepl_stdin = uu___4 + FStar_Interactive_ReplState.grepl_repls = uu___3; + FStar_Interactive_ReplState.grepl_stdin = uu___4 } in go uu___2 in FStar_Compiler_Effect.exit uu___1 \ No newline at end of file diff --git a/ocaml/fstar-lib/generated/FStar_Interactive_PushHelper.ml b/ocaml/fstar-lib/generated/FStar_Interactive_PushHelper.ml index 88dba1c09dc..566d1fc0fc1 100644 --- a/ocaml/fstar-lib/generated/FStar_Interactive_PushHelper.ml +++ b/ocaml/fstar-lib/generated/FStar_Interactive_PushHelper.ml @@ -14,8 +14,8 @@ type ctx_depth_t = (Prims.int * Prims.int * FStar_TypeChecker_Env.solver_depth_t * Prims.int) type deps_t = FStar_Parser_Dep.deps type either_replst = - (FStar_Interactive_JsonHelper.repl_state, - FStar_Interactive_JsonHelper.repl_state) FStar_Pervasives.either + (FStar_Interactive_ReplState.repl_state, + FStar_Interactive_ReplState.repl_state) FStar_Pervasives.either type name_tracking_event = | NTAlias of (FStar_Ident.lid * FStar_Ident.ident * FStar_Ident.lid) | NTOpen of (FStar_Ident.lid * FStar_Syntax_DsEnv.open_module_or_namespace) @@ -50,7 +50,7 @@ let (__proj__NTBinding__item___0 : FStar_Pervasives.either) = fun projectee -> match projectee with | NTBinding _0 -> _0 let (repl_stack : - FStar_Interactive_JsonHelper.repl_stack_t FStar_Compiler_Effect.ref) = + FStar_Interactive_ReplState.repl_stack_t FStar_Compiler_Effect.ref) = FStar_Compiler_Util.mk_ref [] let (set_check_kind : FStar_TypeChecker_Env.env_t -> push_kind -> FStar_TypeChecker_Env.env_t) = @@ -144,16 +144,16 @@ let (set_check_kind : } let (repl_ld_tasks_of_deps : Prims.string Prims.list -> - FStar_Interactive_JsonHelper.repl_task Prims.list -> - FStar_Interactive_JsonHelper.repl_task Prims.list) + FStar_Interactive_ReplState.repl_task Prims.list -> + FStar_Interactive_ReplState.repl_task Prims.list) = fun deps -> fun final_tasks -> let wrap fname = let uu___ = FStar_Compiler_Util.now () in { - FStar_Interactive_JsonHelper.tf_fname = fname; - FStar_Interactive_JsonHelper.tf_modtime = uu___ + FStar_Interactive_ReplState.tf_fname = fname; + FStar_Interactive_ReplState.tf_modtime = uu___ } in let rec aux deps1 final_tasks1 = match deps1 with @@ -163,18 +163,18 @@ let (repl_ld_tasks_of_deps : let uu___1 = let uu___2 = wrap intf in let uu___3 = wrap impl in (uu___2, uu___3) in - FStar_Interactive_JsonHelper.LDInterleaved uu___1 in + FStar_Interactive_ReplState.LDInterleaved uu___1 in let uu___1 = aux deps' final_tasks1 in uu___ :: uu___1 | intf_or_impl::deps' -> let uu___ = let uu___1 = wrap intf_or_impl in - FStar_Interactive_JsonHelper.LDSingle uu___1 in + FStar_Interactive_ReplState.LDSingle uu___1 in let uu___1 = aux deps' final_tasks1 in uu___ :: uu___1 | [] -> final_tasks1 in aux deps final_tasks let (deps_and_repl_ld_tasks_of_our_file : Prims.string -> - (Prims.string Prims.list * FStar_Interactive_JsonHelper.repl_task + (Prims.string Prims.list * FStar_Interactive_ReplState.repl_task Prims.list * deps_t)) = fun filename -> @@ -221,10 +221,10 @@ let (deps_and_repl_ld_tasks_of_our_file : let uu___5 = let uu___6 = FStar_Compiler_Util.now () in { - FStar_Interactive_JsonHelper.tf_fname = intf; - FStar_Interactive_JsonHelper.tf_modtime = uu___6 + FStar_Interactive_ReplState.tf_fname = intf; + FStar_Interactive_ReplState.tf_modtime = uu___6 } in - FStar_Interactive_JsonHelper.LDInterfaceOfCurrentFile + FStar_Interactive_ReplState.LDInterfaceOfCurrentFile uu___5 in [uu___4])) | impl::[] -> [] @@ -243,7 +243,7 @@ let (deps_and_repl_ld_tasks_of_our_file : let (snapshot_env : FStar_TypeChecker_Env.env -> Prims.string -> - (FStar_Interactive_JsonHelper.repl_depth_t * + (FStar_Interactive_ReplState.repl_depth_t * FStar_TypeChecker_Env.env_t)) = fun env -> @@ -257,16 +257,16 @@ let (snapshot_env : let (push_repl : Prims.string -> push_kind -> - FStar_Interactive_JsonHelper.repl_task -> - FStar_Interactive_JsonHelper.repl_state -> - FStar_Interactive_JsonHelper.repl_state) + FStar_Interactive_ReplState.repl_task -> + FStar_Interactive_ReplState.repl_state -> + FStar_Interactive_ReplState.repl_state) = fun msg -> fun push_kind1 -> fun task -> fun st -> let uu___ = - snapshot_env st.FStar_Interactive_JsonHelper.repl_env msg in + snapshot_env st.FStar_Interactive_ReplState.repl_env msg in match uu___ with | (depth, env) -> ((let uu___2 = @@ -275,21 +275,21 @@ let (push_repl : FStar_Compiler_Effect.op_Colon_Equals repl_stack uu___2); (let uu___2 = set_check_kind env push_kind1 in { - FStar_Interactive_JsonHelper.repl_line = - (st.FStar_Interactive_JsonHelper.repl_line); - FStar_Interactive_JsonHelper.repl_column = - (st.FStar_Interactive_JsonHelper.repl_column); - FStar_Interactive_JsonHelper.repl_fname = - (st.FStar_Interactive_JsonHelper.repl_fname); - FStar_Interactive_JsonHelper.repl_deps_stack = - (st.FStar_Interactive_JsonHelper.repl_deps_stack); - FStar_Interactive_JsonHelper.repl_curmod = - (st.FStar_Interactive_JsonHelper.repl_curmod); - FStar_Interactive_JsonHelper.repl_env = uu___2; - FStar_Interactive_JsonHelper.repl_stdin = - (st.FStar_Interactive_JsonHelper.repl_stdin); - FStar_Interactive_JsonHelper.repl_names = - (st.FStar_Interactive_JsonHelper.repl_names) + FStar_Interactive_ReplState.repl_line = + (st.FStar_Interactive_ReplState.repl_line); + FStar_Interactive_ReplState.repl_column = + (st.FStar_Interactive_ReplState.repl_column); + FStar_Interactive_ReplState.repl_fname = + (st.FStar_Interactive_ReplState.repl_fname); + FStar_Interactive_ReplState.repl_deps_stack = + (st.FStar_Interactive_ReplState.repl_deps_stack); + FStar_Interactive_ReplState.repl_curmod = + (st.FStar_Interactive_ReplState.repl_curmod); + FStar_Interactive_ReplState.repl_env = uu___2; + FStar_Interactive_ReplState.repl_stdin = + (st.FStar_Interactive_ReplState.repl_stdin); + FStar_Interactive_ReplState.repl_names = + (st.FStar_Interactive_ReplState.repl_names) })) let (rollback_env : FStar_TypeChecker_Env.solver_t -> @@ -309,8 +309,8 @@ let (rollback_env : env) let (pop_repl : Prims.string -> - FStar_Interactive_JsonHelper.repl_state -> - FStar_Interactive_JsonHelper.repl_state) + FStar_Interactive_ReplState.repl_state -> + FStar_Interactive_ReplState.repl_state) = fun msg -> fun st -> @@ -320,12 +320,12 @@ let (pop_repl : | (depth, (uu___1, st'))::stack_tl -> let env = rollback_env - (st.FStar_Interactive_JsonHelper.repl_env).FStar_TypeChecker_Env.solver + (st.FStar_Interactive_ReplState.repl_env).FStar_TypeChecker_Env.solver msg depth in (FStar_Compiler_Effect.op_Colon_Equals repl_stack stack_tl; (let uu___4 = FStar_Compiler_Util.physical_equality env - st'.FStar_Interactive_JsonHelper.repl_env in + st'.FStar_Interactive_ReplState.repl_env in FStar_Common.runtime_assert uu___4 "Inconsistent stack state"); st') let (tc_one : @@ -345,35 +345,35 @@ let (tc_one : FStar_Universal.tc_one_file_for_ide env intf_opt modf parse_data in match uu___ with | (uu___1, env1) -> env1 let (run_repl_task : - FStar_Interactive_JsonHelper.optmod_t -> + FStar_Interactive_ReplState.optmod_t -> FStar_TypeChecker_Env.env_t -> - FStar_Interactive_JsonHelper.repl_task -> - (FStar_Interactive_JsonHelper.optmod_t * FStar_TypeChecker_Env.env_t)) + FStar_Interactive_ReplState.repl_task -> + (FStar_Interactive_ReplState.optmod_t * FStar_TypeChecker_Env.env_t)) = fun curmod -> fun env -> fun task -> match task with - | FStar_Interactive_JsonHelper.LDInterleaved (intf, impl) -> + | FStar_Interactive_ReplState.LDInterleaved (intf, impl) -> let uu___ = tc_one env (FStar_Pervasives_Native.Some - (intf.FStar_Interactive_JsonHelper.tf_fname)) - impl.FStar_Interactive_JsonHelper.tf_fname in + (intf.FStar_Interactive_ReplState.tf_fname)) + impl.FStar_Interactive_ReplState.tf_fname in (curmod, uu___) - | FStar_Interactive_JsonHelper.LDSingle intf_or_impl -> + | FStar_Interactive_ReplState.LDSingle intf_or_impl -> let uu___ = tc_one env FStar_Pervasives_Native.None - intf_or_impl.FStar_Interactive_JsonHelper.tf_fname in + intf_or_impl.FStar_Interactive_ReplState.tf_fname in (curmod, uu___) - | FStar_Interactive_JsonHelper.LDInterfaceOfCurrentFile intf -> + | FStar_Interactive_ReplState.LDInterfaceOfCurrentFile intf -> let uu___ = FStar_Universal.load_interface_decls env - intf.FStar_Interactive_JsonHelper.tf_fname in + intf.FStar_Interactive_ReplState.tf_fname in (curmod, uu___) - | FStar_Interactive_JsonHelper.PushFragment frag -> + | FStar_Interactive_ReplState.PushFragment frag -> FStar_Universal.tc_one_fragment curmod env frag - | FStar_Interactive_JsonHelper.Noop -> (curmod, env) + | FStar_Interactive_ReplState.Noop -> (curmod, env) let (query_of_ids : FStar_Ident.ident Prims.list -> FStar_Interactive_CompletionTable.query) = fun ids -> FStar_Compiler_List.map FStar_Ident.string_of_id ids @@ -462,30 +462,30 @@ let (commit_name_tracking' : let updater = update_names_from_event cur_mod_str in FStar_Compiler_List.fold_left updater names name_events let (commit_name_tracking : - FStar_Interactive_JsonHelper.repl_state -> - name_tracking_event Prims.list -> FStar_Interactive_JsonHelper.repl_state) + FStar_Interactive_ReplState.repl_state -> + name_tracking_event Prims.list -> FStar_Interactive_ReplState.repl_state) = fun st -> fun name_events -> let names = - commit_name_tracking' st.FStar_Interactive_JsonHelper.repl_curmod - st.FStar_Interactive_JsonHelper.repl_names name_events in + commit_name_tracking' st.FStar_Interactive_ReplState.repl_curmod + st.FStar_Interactive_ReplState.repl_names name_events in { - FStar_Interactive_JsonHelper.repl_line = - (st.FStar_Interactive_JsonHelper.repl_line); - FStar_Interactive_JsonHelper.repl_column = - (st.FStar_Interactive_JsonHelper.repl_column); - FStar_Interactive_JsonHelper.repl_fname = - (st.FStar_Interactive_JsonHelper.repl_fname); - FStar_Interactive_JsonHelper.repl_deps_stack = - (st.FStar_Interactive_JsonHelper.repl_deps_stack); - FStar_Interactive_JsonHelper.repl_curmod = - (st.FStar_Interactive_JsonHelper.repl_curmod); - FStar_Interactive_JsonHelper.repl_env = - (st.FStar_Interactive_JsonHelper.repl_env); - FStar_Interactive_JsonHelper.repl_stdin = - (st.FStar_Interactive_JsonHelper.repl_stdin); - FStar_Interactive_JsonHelper.repl_names = names + FStar_Interactive_ReplState.repl_line = + (st.FStar_Interactive_ReplState.repl_line); + FStar_Interactive_ReplState.repl_column = + (st.FStar_Interactive_ReplState.repl_column); + FStar_Interactive_ReplState.repl_fname = + (st.FStar_Interactive_ReplState.repl_fname); + FStar_Interactive_ReplState.repl_deps_stack = + (st.FStar_Interactive_ReplState.repl_deps_stack); + FStar_Interactive_ReplState.repl_curmod = + (st.FStar_Interactive_ReplState.repl_curmod); + FStar_Interactive_ReplState.repl_env = + (st.FStar_Interactive_ReplState.repl_env); + FStar_Interactive_ReplState.repl_stdin = + (st.FStar_Interactive_ReplState.repl_stdin); + FStar_Interactive_ReplState.repl_names = names } let (fresh_name_tracking_hooks : unit -> @@ -563,11 +563,11 @@ let (track_name_changes : FStar_Compiler_List.rev uu___5 in (uu___3, uu___4))))) let (repl_tx : - FStar_Interactive_JsonHelper.repl_state -> + FStar_Interactive_ReplState.repl_state -> push_kind -> - FStar_Interactive_JsonHelper.repl_task -> + FStar_Interactive_ReplState.repl_task -> (FStar_Interactive_JsonHelper.assoct FStar_Pervasives_Native.option * - FStar_Interactive_JsonHelper.repl_state)) + FStar_Interactive_ReplState.repl_state)) = fun st -> fun push_kind1 -> @@ -579,32 +579,32 @@ let (repl_tx : let st1 = push_repl "repl_tx" push_kind1 task st in let uu___1 = track_name_changes - st1.FStar_Interactive_JsonHelper.repl_env in + st1.FStar_Interactive_ReplState.repl_env in (match uu___1 with | (env, finish_name_tracking) -> let uu___2 = run_repl_task - st1.FStar_Interactive_JsonHelper.repl_curmod env + st1.FStar_Interactive_ReplState.repl_curmod env task in (match uu___2 with | (curmod, env1) -> let st2 = { - FStar_Interactive_JsonHelper.repl_line = - (st1.FStar_Interactive_JsonHelper.repl_line); - FStar_Interactive_JsonHelper.repl_column = - (st1.FStar_Interactive_JsonHelper.repl_column); - FStar_Interactive_JsonHelper.repl_fname = - (st1.FStar_Interactive_JsonHelper.repl_fname); - FStar_Interactive_JsonHelper.repl_deps_stack = - (st1.FStar_Interactive_JsonHelper.repl_deps_stack); - FStar_Interactive_JsonHelper.repl_curmod = + FStar_Interactive_ReplState.repl_line = + (st1.FStar_Interactive_ReplState.repl_line); + FStar_Interactive_ReplState.repl_column = + (st1.FStar_Interactive_ReplState.repl_column); + FStar_Interactive_ReplState.repl_fname = + (st1.FStar_Interactive_ReplState.repl_fname); + FStar_Interactive_ReplState.repl_deps_stack = + (st1.FStar_Interactive_ReplState.repl_deps_stack); + FStar_Interactive_ReplState.repl_curmod = curmod; - FStar_Interactive_JsonHelper.repl_env = env1; - FStar_Interactive_JsonHelper.repl_stdin = - (st1.FStar_Interactive_JsonHelper.repl_stdin); - FStar_Interactive_JsonHelper.repl_names = - (st1.FStar_Interactive_JsonHelper.repl_names) + FStar_Interactive_ReplState.repl_env = env1; + FStar_Interactive_ReplState.repl_stdin = + (st1.FStar_Interactive_ReplState.repl_stdin); + FStar_Interactive_ReplState.repl_names = + (st1.FStar_Interactive_ReplState.repl_names) } in let uu___3 = finish_name_tracking env1 in (match uu___3 with @@ -617,7 +617,7 @@ let (repl_tx : let uu___1 = let uu___2 = FStar_Interactive_JsonHelper.js_diag - st.FStar_Interactive_JsonHelper.repl_fname msg + st.FStar_Interactive_ReplState.repl_fname msg FStar_Pervasives_Native.None in FStar_Pervasives_Native.Some uu___2 in (uu___1, st) @@ -628,7 +628,7 @@ let (repl_tx : let uu___1 = let uu___2 = FStar_Interactive_JsonHelper.js_diag - st.FStar_Interactive_JsonHelper.repl_fname msg + st.FStar_Interactive_ReplState.repl_fname msg (FStar_Pervasives_Native.Some r) in FStar_Pervasives_Native.Some uu___2 in (uu___1, st) @@ -636,44 +636,43 @@ let (repl_tx : let uu___1 = let uu___2 = FStar_Interactive_JsonHelper.js_diag - st.FStar_Interactive_JsonHelper.repl_fname msg + st.FStar_Interactive_ReplState.repl_fname msg FStar_Pervasives_Native.None in FStar_Pervasives_Native.Some uu___2 in (uu___1, st) | FStar_Errors.Stop -> (FStar_Compiler_Util.print_error "[E] Stop"; (FStar_Pervasives_Native.None, st)) -let (tf_of_fname : Prims.string -> FStar_Interactive_JsonHelper.timed_fname) - = +let (tf_of_fname : Prims.string -> FStar_Interactive_ReplState.timed_fname) = fun fname -> let uu___ = FStar_Parser_ParseIt.get_file_last_modification_time fname in { - FStar_Interactive_JsonHelper.tf_fname = fname; - FStar_Interactive_JsonHelper.tf_modtime = uu___ + FStar_Interactive_ReplState.tf_fname = fname; + FStar_Interactive_ReplState.tf_modtime = uu___ } let (update_task_timestamps : - FStar_Interactive_JsonHelper.repl_task -> - FStar_Interactive_JsonHelper.repl_task) + FStar_Interactive_ReplState.repl_task -> + FStar_Interactive_ReplState.repl_task) = fun uu___ -> match uu___ with - | FStar_Interactive_JsonHelper.LDInterleaved (intf, impl) -> + | FStar_Interactive_ReplState.LDInterleaved (intf, impl) -> let uu___1 = - let uu___2 = tf_of_fname intf.FStar_Interactive_JsonHelper.tf_fname in - let uu___3 = tf_of_fname impl.FStar_Interactive_JsonHelper.tf_fname in + let uu___2 = tf_of_fname intf.FStar_Interactive_ReplState.tf_fname in + let uu___3 = tf_of_fname impl.FStar_Interactive_ReplState.tf_fname in (uu___2, uu___3) in - FStar_Interactive_JsonHelper.LDInterleaved uu___1 - | FStar_Interactive_JsonHelper.LDSingle intf_or_impl -> + FStar_Interactive_ReplState.LDInterleaved uu___1 + | FStar_Interactive_ReplState.LDSingle intf_or_impl -> let uu___1 = - tf_of_fname intf_or_impl.FStar_Interactive_JsonHelper.tf_fname in - FStar_Interactive_JsonHelper.LDSingle uu___1 - | FStar_Interactive_JsonHelper.LDInterfaceOfCurrentFile intf -> - let uu___1 = tf_of_fname intf.FStar_Interactive_JsonHelper.tf_fname in - FStar_Interactive_JsonHelper.LDInterfaceOfCurrentFile uu___1 + tf_of_fname intf_or_impl.FStar_Interactive_ReplState.tf_fname in + FStar_Interactive_ReplState.LDSingle uu___1 + | FStar_Interactive_ReplState.LDInterfaceOfCurrentFile intf -> + let uu___1 = tf_of_fname intf.FStar_Interactive_ReplState.tf_fname in + FStar_Interactive_ReplState.LDInterfaceOfCurrentFile uu___1 | other -> other let (repl_ldtx : - FStar_Interactive_JsonHelper.repl_state -> - FStar_Interactive_JsonHelper.repl_task Prims.list -> either_replst) + FStar_Interactive_ReplState.repl_state -> + FStar_Interactive_ReplState.repl_task Prims.list -> either_replst) = fun st -> fun tasks -> @@ -684,27 +683,27 @@ let (repl_ldtx : let st' = pop_repl "repl_ldtx" st1 in let dep_graph = FStar_TypeChecker_Env.dep_graph - st1.FStar_Interactive_JsonHelper.repl_env in + st1.FStar_Interactive_ReplState.repl_env in let st'1 = let uu___1 = FStar_TypeChecker_Env.set_dep_graph - st'.FStar_Interactive_JsonHelper.repl_env dep_graph in + st'.FStar_Interactive_ReplState.repl_env dep_graph in { - FStar_Interactive_JsonHelper.repl_line = - (st'.FStar_Interactive_JsonHelper.repl_line); - FStar_Interactive_JsonHelper.repl_column = - (st'.FStar_Interactive_JsonHelper.repl_column); - FStar_Interactive_JsonHelper.repl_fname = - (st'.FStar_Interactive_JsonHelper.repl_fname); - FStar_Interactive_JsonHelper.repl_deps_stack = - (st'.FStar_Interactive_JsonHelper.repl_deps_stack); - FStar_Interactive_JsonHelper.repl_curmod = - (st'.FStar_Interactive_JsonHelper.repl_curmod); - FStar_Interactive_JsonHelper.repl_env = uu___1; - FStar_Interactive_JsonHelper.repl_stdin = - (st'.FStar_Interactive_JsonHelper.repl_stdin); - FStar_Interactive_JsonHelper.repl_names = - (st'.FStar_Interactive_JsonHelper.repl_names) + FStar_Interactive_ReplState.repl_line = + (st'.FStar_Interactive_ReplState.repl_line); + FStar_Interactive_ReplState.repl_column = + (st'.FStar_Interactive_ReplState.repl_column); + FStar_Interactive_ReplState.repl_fname = + (st'.FStar_Interactive_ReplState.repl_fname); + FStar_Interactive_ReplState.repl_deps_stack = + (st'.FStar_Interactive_ReplState.repl_deps_stack); + FStar_Interactive_ReplState.repl_curmod = + (st'.FStar_Interactive_ReplState.repl_curmod); + FStar_Interactive_ReplState.repl_env = uu___1; + FStar_Interactive_ReplState.repl_stdin = + (st'.FStar_Interactive_ReplState.repl_stdin); + FStar_Interactive_ReplState.repl_names = + (st'.FStar_Interactive_ReplState.repl_names) } in revert_many st'1 entries in let rec aux st1 tasks1 previous = @@ -720,21 +719,21 @@ let (repl_ldtx : let uu___1 = let uu___2 = FStar_Compiler_Effect.op_Bang repl_stack in { - FStar_Interactive_JsonHelper.repl_line = - (st2.FStar_Interactive_JsonHelper.repl_line); - FStar_Interactive_JsonHelper.repl_column = - (st2.FStar_Interactive_JsonHelper.repl_column); - FStar_Interactive_JsonHelper.repl_fname = - (st2.FStar_Interactive_JsonHelper.repl_fname); - FStar_Interactive_JsonHelper.repl_deps_stack = uu___2; - FStar_Interactive_JsonHelper.repl_curmod = - (st2.FStar_Interactive_JsonHelper.repl_curmod); - FStar_Interactive_JsonHelper.repl_env = - (st2.FStar_Interactive_JsonHelper.repl_env); - FStar_Interactive_JsonHelper.repl_stdin = - (st2.FStar_Interactive_JsonHelper.repl_stdin); - FStar_Interactive_JsonHelper.repl_names = - (st2.FStar_Interactive_JsonHelper.repl_names) + FStar_Interactive_ReplState.repl_line = + (st2.FStar_Interactive_ReplState.repl_line); + FStar_Interactive_ReplState.repl_column = + (st2.FStar_Interactive_ReplState.repl_column); + FStar_Interactive_ReplState.repl_fname = + (st2.FStar_Interactive_ReplState.repl_fname); + FStar_Interactive_ReplState.repl_deps_stack = uu___2; + FStar_Interactive_ReplState.repl_curmod = + (st2.FStar_Interactive_ReplState.repl_curmod); + FStar_Interactive_ReplState.repl_env = + (st2.FStar_Interactive_ReplState.repl_env); + FStar_Interactive_ReplState.repl_stdin = + (st2.FStar_Interactive_ReplState.repl_stdin); + FStar_Interactive_ReplState.repl_names = + (st2.FStar_Interactive_ReplState.repl_names) } in aux uu___1 tasks2 [] else FStar_Pervasives.Inr st2) @@ -747,11 +746,11 @@ let (repl_ldtx : let uu___ = revert_many st1 previous1 in aux uu___ tasks2 [] in aux st tasks (FStar_Compiler_List.rev - st.FStar_Interactive_JsonHelper.repl_deps_stack) + st.FStar_Interactive_ReplState.repl_deps_stack) let (ld_deps : - FStar_Interactive_JsonHelper.repl_state -> - ((FStar_Interactive_JsonHelper.repl_state * Prims.string Prims.list), - FStar_Interactive_JsonHelper.repl_state) FStar_Pervasives.either) + FStar_Interactive_ReplState.repl_state -> + ((FStar_Interactive_ReplState.repl_state * Prims.string Prims.list), + FStar_Interactive_ReplState.repl_state) FStar_Pervasives.either) = fun st -> try @@ -760,29 +759,29 @@ let (ld_deps : | () -> let uu___1 = deps_and_repl_ld_tasks_of_our_file - st.FStar_Interactive_JsonHelper.repl_fname in + st.FStar_Interactive_ReplState.repl_fname in (match uu___1 with | (deps, tasks, dep_graph) -> let st1 = let uu___2 = FStar_TypeChecker_Env.set_dep_graph - st.FStar_Interactive_JsonHelper.repl_env dep_graph in + st.FStar_Interactive_ReplState.repl_env dep_graph in { - FStar_Interactive_JsonHelper.repl_line = - (st.FStar_Interactive_JsonHelper.repl_line); - FStar_Interactive_JsonHelper.repl_column = - (st.FStar_Interactive_JsonHelper.repl_column); - FStar_Interactive_JsonHelper.repl_fname = - (st.FStar_Interactive_JsonHelper.repl_fname); - FStar_Interactive_JsonHelper.repl_deps_stack = - (st.FStar_Interactive_JsonHelper.repl_deps_stack); - FStar_Interactive_JsonHelper.repl_curmod = - (st.FStar_Interactive_JsonHelper.repl_curmod); - FStar_Interactive_JsonHelper.repl_env = uu___2; - FStar_Interactive_JsonHelper.repl_stdin = - (st.FStar_Interactive_JsonHelper.repl_stdin); - FStar_Interactive_JsonHelper.repl_names = - (st.FStar_Interactive_JsonHelper.repl_names) + FStar_Interactive_ReplState.repl_line = + (st.FStar_Interactive_ReplState.repl_line); + FStar_Interactive_ReplState.repl_column = + (st.FStar_Interactive_ReplState.repl_column); + FStar_Interactive_ReplState.repl_fname = + (st.FStar_Interactive_ReplState.repl_fname); + FStar_Interactive_ReplState.repl_deps_stack = + (st.FStar_Interactive_ReplState.repl_deps_stack); + FStar_Interactive_ReplState.repl_curmod = + (st.FStar_Interactive_ReplState.repl_curmod); + FStar_Interactive_ReplState.repl_env = uu___2; + FStar_Interactive_ReplState.repl_stdin = + (st.FStar_Interactive_ReplState.repl_stdin); + FStar_Interactive_ReplState.repl_names = + (st.FStar_Interactive_ReplState.repl_names) } in let uu___2 = repl_ldtx st1 tasks in (match uu___2 with @@ -848,18 +847,18 @@ let (add_module_completions : (FStar_Compiler_List.rev mods) let (full_lax : Prims.string -> - FStar_Interactive_JsonHelper.repl_state -> + FStar_Interactive_ReplState.repl_state -> (FStar_Interactive_JsonHelper.assoct FStar_Pervasives_Native.option * - FStar_Interactive_JsonHelper.repl_state)) + FStar_Interactive_ReplState.repl_state)) = fun text -> fun st -> FStar_TypeChecker_Env.toggle_id_info - st.FStar_Interactive_JsonHelper.repl_env true; + st.FStar_Interactive_ReplState.repl_env true; (let frag = { FStar_Parser_ParseIt.frag_fname = - (st.FStar_Interactive_JsonHelper.repl_fname); + (st.FStar_Interactive_ReplState.repl_fname); FStar_Parser_ParseIt.frag_text = text; FStar_Parser_ParseIt.frag_line = Prims.int_one; FStar_Parser_ParseIt.frag_col = Prims.int_zero @@ -869,24 +868,26 @@ let (full_lax : | FStar_Pervasives.Inl (st1, deps) -> let names = add_module_completions - st1.FStar_Interactive_JsonHelper.repl_fname deps - st1.FStar_Interactive_JsonHelper.repl_names in + st1.FStar_Interactive_ReplState.repl_fname deps + st1.FStar_Interactive_ReplState.repl_names in repl_tx { - FStar_Interactive_JsonHelper.repl_line = - (st1.FStar_Interactive_JsonHelper.repl_line); - FStar_Interactive_JsonHelper.repl_column = - (st1.FStar_Interactive_JsonHelper.repl_column); - FStar_Interactive_JsonHelper.repl_fname = - (st1.FStar_Interactive_JsonHelper.repl_fname); - FStar_Interactive_JsonHelper.repl_deps_stack = - (st1.FStar_Interactive_JsonHelper.repl_deps_stack); - FStar_Interactive_JsonHelper.repl_curmod = - (st1.FStar_Interactive_JsonHelper.repl_curmod); - FStar_Interactive_JsonHelper.repl_env = - (st1.FStar_Interactive_JsonHelper.repl_env); - FStar_Interactive_JsonHelper.repl_stdin = - (st1.FStar_Interactive_JsonHelper.repl_stdin); - FStar_Interactive_JsonHelper.repl_names = names - } LaxCheck (FStar_Interactive_JsonHelper.PushFragment frag) + FStar_Interactive_ReplState.repl_line = + (st1.FStar_Interactive_ReplState.repl_line); + FStar_Interactive_ReplState.repl_column = + (st1.FStar_Interactive_ReplState.repl_column); + FStar_Interactive_ReplState.repl_fname = + (st1.FStar_Interactive_ReplState.repl_fname); + FStar_Interactive_ReplState.repl_deps_stack = + (st1.FStar_Interactive_ReplState.repl_deps_stack); + FStar_Interactive_ReplState.repl_curmod = + (st1.FStar_Interactive_ReplState.repl_curmod); + FStar_Interactive_ReplState.repl_env = + (st1.FStar_Interactive_ReplState.repl_env); + FStar_Interactive_ReplState.repl_stdin = + (st1.FStar_Interactive_ReplState.repl_stdin); + FStar_Interactive_ReplState.repl_names = names + } LaxCheck + (FStar_Interactive_ReplState.PushFragment + (FStar_Pervasives.Inl frag)) | FStar_Pervasives.Inr st1 -> (FStar_Pervasives_Native.None, st1)) \ No newline at end of file diff --git a/ocaml/fstar-lib/generated/FStar_Interactive_QueryHelper.ml b/ocaml/fstar-lib/generated/FStar_Interactive_QueryHelper.ml index f85ba89dd10..c69641c4a6d 100644 --- a/ocaml/fstar-lib/generated/FStar_Interactive_QueryHelper.ml +++ b/ocaml/fstar-lib/generated/FStar_Interactive_QueryHelper.ml @@ -171,7 +171,7 @@ let mod_filter : (pth, uu___2) in FStar_Pervasives_Native.Some uu___1 let (ck_completion : - FStar_Interactive_JsonHelper.repl_state -> + FStar_Interactive_ReplState.repl_state -> Prims.string -> FStar_Interactive_CompletionTable.completion_result Prims.list) = @@ -180,10 +180,10 @@ let (ck_completion : let needle = FStar_Compiler_Util.split search_term "." in let mods_and_nss = FStar_Interactive_CompletionTable.autocomplete_mod_or_ns - st.FStar_Interactive_JsonHelper.repl_names needle mod_filter in + st.FStar_Interactive_ReplState.repl_names needle mod_filter in let lids = FStar_Interactive_CompletionTable.autocomplete_lid - st.FStar_Interactive_JsonHelper.repl_names needle in + st.FStar_Interactive_ReplState.repl_names needle in FStar_Compiler_List.op_At lids mods_and_nss let (deflookup : FStar_TypeChecker_Env.env -> @@ -235,7 +235,7 @@ let (hoverlookup : ("value", (FStar_Compiler_Util.JsonStr hovertxt))]))]) | uu___1 -> FStar_Interactive_JsonHelper.nullResponse let (complookup : - FStar_Interactive_JsonHelper.repl_state -> + FStar_Interactive_ReplState.repl_state -> FStar_Interactive_JsonHelper.txdoc_pos -> FStar_Interactive_JsonHelper.assoct FStar_Pervasives_Native.option) = diff --git a/ocaml/fstar-lib/generated/FStar_Interactive_ReplState.ml b/ocaml/fstar-lib/generated/FStar_Interactive_ReplState.ml new file mode 100755 index 00000000000..1e318053bf6 --- /dev/null +++ b/ocaml/fstar-lib/generated/FStar_Interactive_ReplState.ml @@ -0,0 +1,184 @@ +open Prims +type repl_depth_t = (FStar_TypeChecker_Env.tcenv_depth_t * Prims.int) +type optmod_t = FStar_Syntax_Syntax.modul FStar_Pervasives_Native.option +type timed_fname = + { + tf_fname: Prims.string ; + tf_modtime: FStar_Compiler_Util.time } +let (__proj__Mktimed_fname__item__tf_fname : timed_fname -> Prims.string) = + fun projectee -> + match projectee with | { tf_fname; tf_modtime;_} -> tf_fname +let (__proj__Mktimed_fname__item__tf_modtime : + timed_fname -> FStar_Compiler_Util.time) = + fun projectee -> + match projectee with | { tf_fname; tf_modtime;_} -> tf_modtime +type repl_task = + | LDInterleaved of (timed_fname * timed_fname) + | LDSingle of timed_fname + | LDInterfaceOfCurrentFile of timed_fname + | PushFragment of (FStar_Parser_ParseIt.input_frag, FStar_Parser_AST.decl) + FStar_Pervasives.either + | Noop +let (uu___is_LDInterleaved : repl_task -> Prims.bool) = + fun projectee -> + match projectee with | LDInterleaved _0 -> true | uu___ -> false +let (__proj__LDInterleaved__item___0 : + repl_task -> (timed_fname * timed_fname)) = + fun projectee -> match projectee with | LDInterleaved _0 -> _0 +let (uu___is_LDSingle : repl_task -> Prims.bool) = + fun projectee -> + match projectee with | LDSingle _0 -> true | uu___ -> false +let (__proj__LDSingle__item___0 : repl_task -> timed_fname) = + fun projectee -> match projectee with | LDSingle _0 -> _0 +let (uu___is_LDInterfaceOfCurrentFile : repl_task -> Prims.bool) = + fun projectee -> + match projectee with + | LDInterfaceOfCurrentFile _0 -> true + | uu___ -> false +let (__proj__LDInterfaceOfCurrentFile__item___0 : repl_task -> timed_fname) = + fun projectee -> match projectee with | LDInterfaceOfCurrentFile _0 -> _0 +let (uu___is_PushFragment : repl_task -> Prims.bool) = + fun projectee -> + match projectee with | PushFragment _0 -> true | uu___ -> false +let (__proj__PushFragment__item___0 : + repl_task -> + (FStar_Parser_ParseIt.input_frag, FStar_Parser_AST.decl) + FStar_Pervasives.either) + = fun projectee -> match projectee with | PushFragment _0 -> _0 +let (uu___is_Noop : repl_task -> Prims.bool) = + fun projectee -> match projectee with | Noop -> true | uu___ -> false +type repl_state = + { + repl_line: Prims.int ; + repl_column: Prims.int ; + repl_fname: Prims.string ; + repl_deps_stack: (repl_depth_t * (repl_task * repl_state)) Prims.list ; + repl_curmod: optmod_t ; + repl_env: FStar_TypeChecker_Env.env ; + repl_stdin: FStar_Compiler_Util.stream_reader ; + repl_names: FStar_Interactive_CompletionTable.table } +let (__proj__Mkrepl_state__item__repl_line : repl_state -> Prims.int) = + fun projectee -> + match projectee with + | { repl_line; repl_column; repl_fname; repl_deps_stack; repl_curmod; + repl_env; repl_stdin; repl_names;_} -> repl_line +let (__proj__Mkrepl_state__item__repl_column : repl_state -> Prims.int) = + fun projectee -> + match projectee with + | { repl_line; repl_column; repl_fname; repl_deps_stack; repl_curmod; + repl_env; repl_stdin; repl_names;_} -> repl_column +let (__proj__Mkrepl_state__item__repl_fname : repl_state -> Prims.string) = + fun projectee -> + match projectee with + | { repl_line; repl_column; repl_fname; repl_deps_stack; repl_curmod; + repl_env; repl_stdin; repl_names;_} -> repl_fname +let (__proj__Mkrepl_state__item__repl_deps_stack : + repl_state -> (repl_depth_t * (repl_task * repl_state)) Prims.list) = + fun projectee -> + match projectee with + | { repl_line; repl_column; repl_fname; repl_deps_stack; repl_curmod; + repl_env; repl_stdin; repl_names;_} -> repl_deps_stack +let (__proj__Mkrepl_state__item__repl_curmod : repl_state -> optmod_t) = + fun projectee -> + match projectee with + | { repl_line; repl_column; repl_fname; repl_deps_stack; repl_curmod; + repl_env; repl_stdin; repl_names;_} -> repl_curmod +let (__proj__Mkrepl_state__item__repl_env : + repl_state -> FStar_TypeChecker_Env.env) = + fun projectee -> + match projectee with + | { repl_line; repl_column; repl_fname; repl_deps_stack; repl_curmod; + repl_env; repl_stdin; repl_names;_} -> repl_env +let (__proj__Mkrepl_state__item__repl_stdin : + repl_state -> FStar_Compiler_Util.stream_reader) = + fun projectee -> + match projectee with + | { repl_line; repl_column; repl_fname; repl_deps_stack; repl_curmod; + repl_env; repl_stdin; repl_names;_} -> repl_stdin +let (__proj__Mkrepl_state__item__repl_names : + repl_state -> FStar_Interactive_CompletionTable.table) = + fun projectee -> + match projectee with + | { repl_line; repl_column; repl_fname; repl_deps_stack; repl_curmod; + repl_env; repl_stdin; repl_names;_} -> repl_names +type repl_stack_entry_t = (repl_depth_t * (repl_task * repl_state)) +type repl_stack_t = (repl_depth_t * (repl_task * repl_state)) Prims.list +type grepl_state = + { + grepl_repls: repl_state FStar_Compiler_Util.psmap ; + grepl_stdin: FStar_Compiler_Util.stream_reader } +let (__proj__Mkgrepl_state__item__grepl_repls : + grepl_state -> repl_state FStar_Compiler_Util.psmap) = + fun projectee -> + match projectee with | { grepl_repls; grepl_stdin;_} -> grepl_repls +let (__proj__Mkgrepl_state__item__grepl_stdin : + grepl_state -> FStar_Compiler_Util.stream_reader) = + fun projectee -> + match projectee with | { grepl_repls; grepl_stdin;_} -> grepl_stdin +let (t0 : FStar_Compiler_Util.time) = FStar_Compiler_Util.now () +let (dummy_tf_of_fname : Prims.string -> timed_fname) = + fun fname -> { tf_fname = fname; tf_modtime = t0 } +let (string_of_timed_fname : timed_fname -> Prims.string) = + fun uu___ -> + match uu___ with + | { tf_fname = fname; tf_modtime = modtime;_} -> + if modtime = t0 + then FStar_Compiler_Util.format1 "{ %s }" fname + else + (let uu___2 = FStar_Compiler_Util.string_of_time modtime in + FStar_Compiler_Util.format2 "{ %s; %s }" fname uu___2) +let (string_of_repl_task : repl_task -> Prims.string) = + fun uu___ -> + match uu___ with + | LDInterleaved (intf, impl) -> + let uu___1 = string_of_timed_fname intf in + let uu___2 = string_of_timed_fname impl in + FStar_Compiler_Util.format2 "LDInterleaved (%s, %s)" uu___1 uu___2 + | LDSingle intf_or_impl -> + let uu___1 = string_of_timed_fname intf_or_impl in + FStar_Compiler_Util.format1 "LDSingle %s" uu___1 + | LDInterfaceOfCurrentFile intf -> + let uu___1 = string_of_timed_fname intf in + FStar_Compiler_Util.format1 "LDInterfaceOfCurrentFile %s" uu___1 + | PushFragment (FStar_Pervasives.Inl frag) -> + FStar_Compiler_Util.format1 "PushFragment { code = %s }" + frag.FStar_Parser_ParseIt.frag_text + | PushFragment (FStar_Pervasives.Inr d) -> + let uu___1 = FStar_Parser_AST.decl_to_string d in + FStar_Compiler_Util.format1 "PushFragment { decl = %s }" uu___1 + | Noop -> "Noop {}" +let (string_of_repl_stack_entry : repl_stack_entry_t -> Prims.string) = + fun uu___ -> + match uu___ with + | ((depth, i), (task, state)) -> + let uu___1 = + let uu___2 = FStar_Compiler_Util.string_of_int i in + let uu___3 = let uu___4 = string_of_repl_task task in [uu___4] in + uu___2 :: uu___3 in + FStar_Compiler_Util.format "{depth=%s; task=%s}" uu___1 +let (string_of_repl_stack : repl_stack_entry_t Prims.list -> Prims.string) = + fun s -> + let uu___ = FStar_Compiler_List.map string_of_repl_stack_entry s in + FStar_String.concat ";\n\t\t" uu___ +let (repl_state_to_string : repl_state -> Prims.string) = + fun r -> + let uu___ = + let uu___1 = FStar_Compiler_Util.string_of_int r.repl_line in + let uu___2 = + let uu___3 = FStar_Compiler_Util.string_of_int r.repl_column in + let uu___4 = + let uu___5 = + let uu___6 = + match r.repl_curmod with + | FStar_Pervasives_Native.None -> "None" + | FStar_Pervasives_Native.Some m -> + FStar_Ident.string_of_lid m.FStar_Syntax_Syntax.name in + let uu___7 = + let uu___8 = string_of_repl_stack r.repl_deps_stack in [uu___8] in + uu___6 :: uu___7 in + (r.repl_fname) :: uu___5 in + uu___3 :: uu___4 in + uu___1 :: uu___2 in + FStar_Compiler_Util.format + "{\n\trepl_line=%s;\n\trepl_column=%s;\n\trepl_fname=%s;\n\trepl_cur_mod=%s;\n\t\\ \r\n repl_deps_stack={%s}\n}" + uu___ \ No newline at end of file diff --git a/ocaml/fstar-lib/generated/FStar_Parser_AST.ml b/ocaml/fstar-lib/generated/FStar_Parser_AST.ml index b1b69b15d9a..074ea438dd4 100644 --- a/ocaml/fstar-lib/generated/FStar_Parser_AST.ml +++ b/ocaml/fstar-lib/generated/FStar_Parser_AST.ml @@ -738,17 +738,24 @@ type lift = { msource: FStar_Ident.lid ; mdest: FStar_Ident.lid ; - lift_op: lift_op } + lift_op: lift_op ; + braced: Prims.bool } let (__proj__Mklift__item__msource : lift -> FStar_Ident.lid) = fun projectee -> - match projectee with | { msource; mdest; lift_op = lift_op1;_} -> msource + match projectee with + | { msource; mdest; lift_op = lift_op1; braced;_} -> msource let (__proj__Mklift__item__mdest : lift -> FStar_Ident.lid) = fun projectee -> - match projectee with | { msource; mdest; lift_op = lift_op1;_} -> mdest + match projectee with + | { msource; mdest; lift_op = lift_op1; braced;_} -> mdest let (__proj__Mklift__item__lift_op : lift -> lift_op) = fun projectee -> match projectee with - | { msource; mdest; lift_op = lift_op1;_} -> lift_op1 + | { msource; mdest; lift_op = lift_op1; braced;_} -> lift_op1 +let (__proj__Mklift__item__braced : lift -> Prims.bool) = + fun projectee -> + match projectee with + | { msource; mdest; lift_op = lift_op1; braced;_} -> braced type pragma = | SetOptions of Prims.string | ResetOptions of Prims.string FStar_Pervasives_Native.option @@ -2420,4 +2427,14 @@ let (idents_of_binders : fun bs -> fun r -> FStar_Compiler_Effect.op_Bar_Greater bs - (FStar_Compiler_List.map (ident_of_binder r)) \ No newline at end of file + (FStar_Compiler_List.map (ident_of_binder r)) +let (decl_syntax_is_delimited : decl -> Prims.bool) = + fun d -> + match d.d with + | Pragma uu___ -> true + | NewEffect (DefineEffect uu___) -> true + | LayeredEffect (DefineEffect uu___) -> true + | SubEffect + { msource = uu___; mdest = uu___1; lift_op = uu___2; braced = true;_} + -> true + | uu___ -> false \ No newline at end of file diff --git a/ocaml/fstar-lib/generated/FStar_Parser_Dep.ml b/ocaml/fstar-lib/generated/FStar_Parser_Dep.ml index aeb17a7824b..a834706d1fb 100644 --- a/ocaml/fstar-lib/generated/FStar_Parser_Dep.ml +++ b/ocaml/fstar-lib/generated/FStar_Parser_Dep.ml @@ -1234,19 +1234,22 @@ let (collect_one : { FStar_Parser_AST.msource = uu___2; FStar_Parser_AST.mdest = uu___3; FStar_Parser_AST.lift_op = - FStar_Parser_AST.NonReifiableLift t;_} + FStar_Parser_AST.NonReifiableLift t; + FStar_Parser_AST.braced = uu___4;_} -> collect_term t | FStar_Parser_AST.SubEffect { FStar_Parser_AST.msource = uu___2; FStar_Parser_AST.mdest = uu___3; - FStar_Parser_AST.lift_op = FStar_Parser_AST.LiftForFree t;_} + FStar_Parser_AST.lift_op = FStar_Parser_AST.LiftForFree t; + FStar_Parser_AST.braced = uu___4;_} -> collect_term t | FStar_Parser_AST.Val (uu___2, t) -> collect_term t | FStar_Parser_AST.SubEffect { FStar_Parser_AST.msource = uu___2; FStar_Parser_AST.mdest = uu___3; FStar_Parser_AST.lift_op = FStar_Parser_AST.ReifiableLift - (t0, t1);_} + (t0, t1); + FStar_Parser_AST.braced = uu___4;_} -> (collect_term t0; collect_term t1) | FStar_Parser_AST.Tycon (uu___2, tc, ts) -> (if tc diff --git a/ocaml/fstar-lib/generated/FStar_Parser_Driver.ml b/ocaml/fstar-lib/generated/FStar_Parser_Driver.ml index 12772849058..95f21ce391e 100644 --- a/ocaml/fstar-lib/generated/FStar_Parser_Driver.ml +++ b/ocaml/fstar-lib/generated/FStar_Parser_Driver.ml @@ -27,6 +27,8 @@ let (parse_fragment : FStar_Parser_ParseIt.input_frag -> fragment) = Empty | FStar_Parser_ParseIt.ASTFragment (FStar_Pervasives.Inr decls, uu___1) -> Decls decls + | FStar_Parser_ParseIt.IncrementalFragment (decls, uu___1, uu___2) -> + Decls decls | FStar_Parser_ParseIt.ParseError (e, msg, r) -> FStar_Errors.raise_error (e, msg) r | FStar_Parser_ParseIt.Term uu___1 -> diff --git a/ocaml/fstar-lib/generated/FStar_Parser_Parse.ml b/ocaml/fstar-lib/generated/FStar_Parser_Parse.ml index f5cf1fc3ddc..0febb670360 100644 --- a/ocaml/fstar-lib/generated/FStar_Parser_Parse.ml +++ b/ocaml/fstar-lib/generated/FStar_Parser_Parse.ml @@ -403,80 +403,80 @@ let yytransl_block = [| 0|] let yylhs = "\255\255\ -\005\000\005\000\007\000\007\000\008\000\008\000\010\000\010\000\ -\012\000\012\000\013\000\013\000\015\000\015\000\016\000\016\000\ -\017\000\017\000\019\000\019\000\020\000\020\000\022\000\022\000\ -\024\000\024\000\026\000\026\000\028\000\028\000\030\000\030\000\ -\033\000\033\000\035\000\035\000\036\000\036\000\037\000\037\000\ -\039\000\039\000\041\000\041\000\043\000\043\000\045\000\045\000\ -\045\000\045\000\048\000\048\000\050\000\050\000\050\000\053\000\ -\053\000\054\000\054\000\056\000\056\000\058\000\058\000\060\000\ -\060\000\062\000\062\000\064\000\064\000\066\000\066\000\067\000\ -\067\000\069\000\069\000\071\000\071\000\073\000\073\000\074\000\ -\074\000\076\000\076\000\076\000\076\000\076\000\076\000\076\000\ -\076\000\077\000\077\000\079\000\079\000\081\000\081\000\083\000\ -\083\000\038\000\038\000\086\000\086\000\087\000\087\000\089\000\ -\089\000\091\000\091\000\093\000\093\000\095\000\095\000\096\000\ -\096\000\040\000\040\000\042\000\042\000\001\000\099\000\099\000\ -\099\000\099\000\099\000\099\000\100\000\100\000\065\000\065\000\ -\063\000\063\000\063\000\105\000\105\000\104\000\104\000\104\000\ -\104\000\104\000\104\000\104\000\104\000\104\000\104\000\104\000\ -\104\000\104\000\104\000\104\000\104\000\104\000\104\000\104\000\ -\104\000\104\000\104\000\104\000\104\000\104\000\082\000\108\000\ -\108\000\118\000\117\000\117\000\117\000\117\000\117\000\122\000\ -\025\000\025\000\025\000\061\000\061\000\057\000\057\000\044\000\ -\080\000\080\000\080\000\080\000\080\000\080\000\080\000\080\000\ -\111\000\111\000\125\000\112\000\113\000\097\000\114\000\114\000\ -\114\000\115\000\116\000\102\000\102\000\102\000\102\000\102\000\ -\102\000\102\000\102\000\102\000\102\000\102\000\102\000\102\000\ -\102\000\102\000\102\000\102\000\124\000\106\000\106\000\128\000\ -\128\000\128\000\121\000\130\000\084\000\088\000\088\000\088\000\ -\072\000\072\000\072\000\072\000\072\000\072\000\072\000\072\000\ -\072\000\072\000\072\000\072\000\072\000\072\000\072\000\072\000\ -\072\000\134\000\134\000\078\000\078\000\078\000\078\000\051\000\ -\051\000\052\000\052\000\052\000\119\000\068\000\068\000\068\000\ -\068\000\123\000\123\000\123\000\123\000\123\000\123\000\123\000\ -\123\000\123\000\123\000\123\000\123\000\123\000\123\000\123\000\ -\123\000\123\000\123\000\123\000\123\000\123\000\123\000\123\000\ -\123\000\070\000\070\000\070\000\070\000\049\000\107\000\136\000\ -\136\000\137\000\137\000\098\000\098\000\138\000\138\000\138\000\ -\138\000\138\000\138\000\139\000\139\000\140\000\140\000\135\000\ -\135\000\004\000\103\000\092\000\014\000\129\000\018\000\009\000\ -\023\000\021\000\141\000\002\000\002\000\002\000\002\000\029\000\ -\029\000\011\000\011\000\011\000\011\000\011\000\011\000\011\000\ -\011\000\011\000\011\000\011\000\011\000\011\000\011\000\011\000\ -\011\000\011\000\011\000\011\000\011\000\011\000\011\000\011\000\ -\011\000\011\000\011\000\011\000\011\000\011\000\011\000\011\000\ -\011\000\011\000\011\000\011\000\011\000\147\000\150\000\150\000\ -\150\000\059\000\006\000\006\000\006\000\151\000\151\000\152\000\ -\094\000\032\000\032\000\153\000\153\000\154\000\154\000\142\000\ -\142\000\155\000\155\000\156\000\156\000\156\000\156\000\156\000\ -\156\000\156\000\156\000\126\000\126\000\126\000\126\000\126\000\ -\126\000\126\000\126\000\131\000\131\000\157\000\157\000\157\000\ -\157\000\157\000\146\000\146\000\148\000\148\000\149\000\159\000\ -\159\000\159\000\159\000\159\000\159\000\159\000\159\000\159\000\ -\159\000\159\000\159\000\159\000\159\000\159\000\159\000\161\000\ -\161\000\161\000\161\000\161\000\161\000\161\000\161\000\161\000\ -\161\000\161\000\161\000\161\000\161\000\161\000\161\000\160\000\ +\006\000\006\000\008\000\008\000\009\000\009\000\011\000\011\000\ +\013\000\013\000\014\000\014\000\016\000\016\000\017\000\017\000\ +\018\000\018\000\020\000\020\000\021\000\021\000\023\000\023\000\ +\025\000\025\000\027\000\027\000\029\000\029\000\031\000\031\000\ +\034\000\034\000\036\000\036\000\037\000\037\000\038\000\038\000\ +\040\000\040\000\042\000\042\000\044\000\044\000\046\000\046\000\ +\046\000\046\000\049\000\049\000\051\000\051\000\051\000\054\000\ +\054\000\055\000\055\000\057\000\057\000\059\000\059\000\061\000\ +\061\000\063\000\063\000\065\000\065\000\067\000\067\000\068\000\ +\068\000\070\000\070\000\072\000\072\000\074\000\074\000\075\000\ +\075\000\077\000\077\000\077\000\077\000\077\000\077\000\077\000\ +\077\000\078\000\078\000\080\000\080\000\082\000\082\000\084\000\ +\084\000\039\000\039\000\087\000\087\000\088\000\088\000\090\000\ +\090\000\092\000\092\000\094\000\094\000\096\000\096\000\097\000\ +\097\000\041\000\041\000\043\000\043\000\001\000\002\000\002\000\ +\100\000\100\000\100\000\100\000\100\000\100\000\101\000\101\000\ +\066\000\066\000\064\000\064\000\064\000\106\000\106\000\105\000\ +\105\000\105\000\105\000\105\000\105\000\105\000\105\000\105\000\ +\105\000\105\000\105\000\105\000\105\000\105\000\105\000\105\000\ +\105\000\105\000\105\000\105\000\105\000\105\000\105\000\105\000\ +\083\000\109\000\109\000\119\000\118\000\118\000\118\000\118\000\ +\118\000\123\000\026\000\026\000\026\000\062\000\062\000\058\000\ +\058\000\045\000\081\000\081\000\081\000\081\000\081\000\081\000\ +\081\000\081\000\112\000\112\000\126\000\113\000\114\000\098\000\ +\115\000\115\000\115\000\116\000\117\000\103\000\103\000\103\000\ +\103\000\103\000\103\000\103\000\103\000\103\000\103\000\103\000\ +\103\000\103\000\103\000\103\000\103\000\103\000\125\000\107\000\ +\107\000\129\000\129\000\129\000\122\000\131\000\085\000\089\000\ +\089\000\089\000\073\000\073\000\073\000\073\000\073\000\073\000\ +\073\000\073\000\073\000\073\000\073\000\073\000\073\000\073\000\ +\073\000\073\000\073\000\135\000\135\000\079\000\079\000\079\000\ +\079\000\052\000\052\000\053\000\053\000\053\000\120\000\069\000\ +\069\000\069\000\069\000\124\000\124\000\124\000\124\000\124\000\ +\124\000\124\000\124\000\124\000\124\000\124\000\124\000\124\000\ +\124\000\124\000\124\000\124\000\124\000\124\000\124\000\124\000\ +\124\000\124\000\124\000\071\000\071\000\071\000\071\000\050\000\ +\108\000\137\000\137\000\138\000\138\000\099\000\099\000\139\000\ +\139\000\139\000\139\000\139\000\139\000\140\000\140\000\141\000\ +\141\000\136\000\136\000\005\000\104\000\093\000\015\000\130\000\ +\019\000\010\000\024\000\022\000\142\000\003\000\003\000\003\000\ +\003\000\030\000\030\000\012\000\012\000\012\000\012\000\012\000\ +\012\000\012\000\012\000\012\000\012\000\012\000\012\000\012\000\ +\012\000\012\000\012\000\012\000\012\000\012\000\012\000\012\000\ +\012\000\012\000\012\000\012\000\012\000\012\000\012\000\012\000\ +\012\000\012\000\012\000\012\000\012\000\012\000\012\000\148\000\ +\151\000\151\000\151\000\060\000\007\000\007\000\007\000\152\000\ +\152\000\153\000\095\000\033\000\033\000\154\000\154\000\155\000\ +\155\000\143\000\143\000\156\000\156\000\157\000\157\000\157\000\ +\157\000\157\000\157\000\157\000\157\000\127\000\127\000\127\000\ +\127\000\127\000\127\000\127\000\127\000\132\000\132\000\158\000\ +\158\000\158\000\158\000\158\000\147\000\147\000\149\000\149\000\ +\150\000\160\000\160\000\160\000\160\000\160\000\160\000\160\000\ \160\000\160\000\160\000\160\000\160\000\160\000\160\000\160\000\ +\160\000\162\000\162\000\162\000\162\000\162\000\162\000\162\000\ \162\000\162\000\162\000\162\000\162\000\162\000\162\000\162\000\ -\162\000\110\000\110\000\110\000\110\000\110\000\110\000\110\000\ +\162\000\161\000\161\000\161\000\161\000\161\000\161\000\161\000\ +\161\000\161\000\163\000\163\000\163\000\163\000\163\000\163\000\ +\163\000\163\000\163\000\111\000\111\000\111\000\111\000\111\000\ +\111\000\111\000\111\000\111\000\111\000\111\000\111\000\111\000\ +\111\000\111\000\111\000\111\000\111\000\159\000\091\000\128\000\ +\165\000\165\000\133\000\048\000\048\000\167\000\167\000\086\000\ +\164\000\047\000\047\000\047\000\168\000\168\000\056\000\056\000\ +\056\000\170\000\170\000\144\000\144\000\144\000\144\000\144\000\ +\144\000\144\000\144\000\144\000\144\000\144\000\144\000\172\000\ +\171\000\173\000\173\000\173\000\173\000\173\000\173\000\173\000\ +\173\000\173\000\028\000\174\000\175\000\032\000\110\000\110\000\ \110\000\110\000\110\000\110\000\110\000\110\000\110\000\110\000\ -\110\000\110\000\110\000\158\000\090\000\127\000\164\000\164\000\ -\132\000\047\000\047\000\166\000\166\000\085\000\163\000\046\000\ -\046\000\046\000\167\000\167\000\055\000\055\000\055\000\169\000\ -\169\000\143\000\143\000\143\000\143\000\143\000\143\000\143\000\ -\143\000\143\000\143\000\143\000\143\000\171\000\170\000\172\000\ -\172\000\172\000\172\000\172\000\172\000\172\000\172\000\172\000\ -\027\000\173\000\174\000\031\000\109\000\109\000\109\000\109\000\ -\109\000\109\000\109\000\109\000\109\000\109\000\109\000\109\000\ -\109\000\109\000\109\000\109\000\109\000\109\000\109\000\168\000\ -\176\000\176\000\176\000\075\000\075\000\075\000\075\000\003\000\ -\177\000\177\000\178\000\178\000\178\000\179\000\179\000\034\000\ -\175\000\180\000\180\000\180\000\101\000\101\000\101\000\181\000\ -\181\000\181\000\182\000\182\000\182\000\133\000\133\000\120\000\ -\120\000\165\000\165\000\145\000\145\000\145\000\144\000\144\000\ -\144\000\000\000\000\000\000\000" +\110\000\110\000\110\000\110\000\110\000\110\000\110\000\110\000\ +\110\000\169\000\177\000\177\000\177\000\076\000\076\000\076\000\ +\076\000\004\000\178\000\178\000\179\000\179\000\179\000\180\000\ +\180\000\035\000\176\000\181\000\181\000\181\000\102\000\102\000\ +\102\000\182\000\182\000\182\000\183\000\183\000\183\000\134\000\ +\134\000\121\000\121\000\166\000\166\000\146\000\146\000\146\000\ +\145\000\145\000\145\000\000\000\000\000\000\000\000\000" let yylen = "\002\000\ \000\000\002\000\000\000\002\000\000\000\002\000\000\000\003\000\ @@ -493,1204 +493,1334 @@ let yylen = "\002\000\ \004\000\001\000\002\000\001\000\003\000\001\000\003\000\001\000\ \003\000\001\000\003\000\001\000\003\000\001\000\003\000\001\000\ \003\000\001\000\003\000\001\000\003\000\001\000\003\000\001\000\ -\003\000\001\000\003\000\001\000\003\000\002\000\002\000\002\000\ -\002\000\001\000\001\000\001\000\003\000\003\000\001\000\001\000\ -\004\000\002\000\002\000\002\000\003\000\001\000\002\000\002\000\ -\002\000\004\000\002\000\002\000\002\000\005\000\003\000\002\000\ -\005\000\007\000\007\000\007\000\007\000\007\000\005\000\003\000\ -\002\000\002\000\002\000\002\000\002\000\002\000\004\000\001\000\ -\001\000\003\000\000\000\002\000\004\000\005\000\002\000\003\000\ -\002\000\002\000\004\000\003\000\004\000\002\000\003\000\003\000\ -\006\000\008\000\008\000\008\000\008\000\008\000\005\000\004\000\ -\001\000\001\000\003\000\008\000\006\000\004\000\005\000\008\000\ -\012\000\009\000\005\000\001\000\001\000\001\000\001\000\001\000\ +\003\000\001\000\003\000\001\000\003\000\002\000\001\000\001\000\ +\002\000\002\000\002\000\001\000\001\000\001\000\003\000\003\000\ +\001\000\001\000\004\000\002\000\002\000\002\000\003\000\001\000\ +\002\000\002\000\002\000\004\000\002\000\002\000\002\000\005\000\ +\003\000\002\000\005\000\007\000\007\000\007\000\007\000\007\000\ +\005\000\003\000\002\000\002\000\002\000\002\000\002\000\002\000\ +\004\000\001\000\001\000\003\000\000\000\002\000\004\000\005\000\ +\002\000\003\000\002\000\002\000\004\000\003\000\004\000\002\000\ +\003\000\003\000\006\000\008\000\008\000\008\000\008\000\008\000\ +\005\000\004\000\001\000\001\000\003\000\008\000\006\000\004\000\ +\005\000\008\000\012\000\009\000\005\000\001\000\001\000\001\000\ +\001\000\001\000\001\000\001\000\001\000\001\000\001\000\001\000\ \001\000\001\000\001\000\001\000\001\000\001\000\001\000\001\000\ -\001\000\001\000\001\000\001\000\001\000\001\000\000\000\004\000\ -\001\000\001\000\003\000\001\000\001\000\003\000\002\000\001\000\ -\006\000\003\000\003\000\005\000\003\000\001\000\003\000\003\000\ -\003\000\003\000\003\000\001\000\002\000\001\000\002\000\001\000\ -\001\000\003\000\001\000\005\000\003\000\001\000\007\000\001\000\ -\001\000\005\000\003\000\006\000\001\000\003\000\002\000\002\000\ -\001\000\003\000\005\000\005\000\005\000\005\000\005\000\002\000\ -\004\000\004\000\004\000\004\000\004\000\002\000\004\000\004\000\ -\004\000\004\000\004\000\001\000\003\000\003\000\003\000\003\000\ -\003\000\003\000\002\000\002\000\001\000\001\000\001\000\001\000\ -\003\000\001\000\003\000\001\000\001\000\001\000\003\000\003\000\ -\003\000\003\000\003\000\001\000\001\000\001\000\001\000\001\000\ +\000\000\004\000\001\000\001\000\003\000\001\000\001\000\003\000\ +\002\000\001\000\006\000\003\000\003\000\005\000\003\000\001\000\ +\003\000\003\000\003\000\003\000\003\000\001\000\002\000\001\000\ +\002\000\001\000\001\000\003\000\001\000\005\000\003\000\001\000\ +\007\000\001\000\001\000\005\000\003\000\006\000\001\000\003\000\ +\002\000\002\000\001\000\003\000\005\000\005\000\005\000\005\000\ +\005\000\002\000\004\000\004\000\004\000\004\000\004\000\002\000\ +\004\000\004\000\004\000\004\000\004\000\001\000\003\000\003\000\ +\003\000\003\000\003\000\003\000\002\000\002\000\001\000\001\000\ +\001\000\001\000\003\000\001\000\003\000\001\000\001\000\001\000\ +\003\000\003\000\003\000\003\000\003\000\001\000\001\000\001\000\ \001\000\001\000\001\000\001\000\001\000\001\000\001\000\001\000\ -\003\000\002\000\001\000\001\000\003\000\003\000\005\000\003\000\ -\003\000\001\000\004\000\004\000\006\000\006\000\006\000\006\000\ -\002\000\002\000\002\000\004\000\002\000\007\000\005\000\004\000\ -\005\000\005\000\006\000\007\000\005\000\002\000\002\000\003\000\ -\003\000\002\000\007\000\007\000\009\000\008\000\007\000\008\000\ -\007\000\011\000\006\000\014\000\010\000\001\000\001\000\003\000\ -\001\000\006\000\001\000\005\000\005\000\000\000\003\000\001\000\ -\001\000\001\000\004\000\001\000\001\000\003\000\005\000\003\000\ -\001\000\003\000\001\000\005\000\006\000\007\000\003\000\004\000\ -\004\000\005\000\001\000\005\000\006\000\007\000\003\000\004\000\ -\004\000\005\000\001\000\003\000\001\000\003\000\001\000\002\000\ -\002\000\003\000\003\000\001\000\003\000\001\000\001\000\003\000\ -\003\000\003\000\003\000\003\000\003\000\003\000\003\000\003\000\ -\003\000\002\000\002\000\002\000\002\000\002\000\001\000\003\000\ -\003\000\003\000\003\000\003\000\003\000\003\000\003\000\003\000\ -\003\000\002\000\002\000\002\000\002\000\002\000\001\000\003\000\ -\003\000\003\000\005\000\003\000\003\000\002\000\002\000\001\000\ -\003\000\003\000\003\000\005\000\003\000\003\000\002\000\002\000\ +\001\000\001\000\003\000\002\000\001\000\001\000\003\000\003\000\ +\005\000\003\000\003\000\001\000\004\000\004\000\006\000\006\000\ +\006\000\006\000\002\000\002\000\002\000\004\000\002\000\007\000\ +\005\000\004\000\005\000\005\000\006\000\007\000\005\000\002\000\ +\002\000\003\000\003\000\002\000\007\000\007\000\009\000\008\000\ +\007\000\008\000\007\000\011\000\006\000\014\000\010\000\001\000\ +\001\000\003\000\001\000\006\000\001\000\005\000\005\000\000\000\ +\003\000\001\000\001\000\001\000\004\000\001\000\001\000\003\000\ +\005\000\003\000\001\000\003\000\001\000\005\000\006\000\007\000\ +\003\000\004\000\004\000\005\000\001\000\005\000\006\000\007\000\ +\003\000\004\000\004\000\005\000\001\000\003\000\001\000\003\000\ +\001\000\002\000\002\000\003\000\003\000\001\000\003\000\001\000\ +\001\000\003\000\003\000\003\000\003\000\003\000\003\000\003\000\ +\003\000\003\000\003\000\002\000\002\000\002\000\002\000\002\000\ +\001\000\003\000\003\000\003\000\003\000\003\000\003\000\003\000\ +\003\000\003\000\003\000\002\000\002\000\002\000\002\000\002\000\ +\001\000\003\000\003\000\003\000\005\000\003\000\003\000\002\000\ +\002\000\001\000\003\000\003\000\003\000\005\000\003\000\003\000\ +\002\000\002\000\001\000\001\000\001\000\001\000\001\000\001\000\ \001\000\001\000\001\000\001\000\001\000\001\000\001\000\001\000\ \001\000\001\000\001\000\001\000\001\000\001\000\001\000\001\000\ -\001\000\001\000\001\000\001\000\001\000\001\000\004\000\001\000\ -\001\000\001\000\003\000\003\000\001\000\002\000\002\000\001\000\ -\002\000\001\000\002\000\001\000\001\000\001\000\001\000\001\000\ -\004\000\001\000\001\000\001\000\001\000\003\000\003\000\003\000\ -\003\000\003\000\005\000\002\000\003\000\002\000\002\000\001\000\ -\001\000\004\000\003\000\003\000\003\000\003\000\003\000\002\000\ -\003\000\002\000\002\000\001\000\001\000\001\000\001\000\001\000\ +\004\000\001\000\001\000\001\000\003\000\003\000\001\000\002\000\ +\002\000\001\000\002\000\001\000\002\000\001\000\001\000\001\000\ +\001\000\001\000\004\000\001\000\001\000\001\000\001\000\003\000\ +\003\000\003\000\003\000\003\000\005\000\002\000\003\000\002\000\ +\002\000\001\000\001\000\004\000\003\000\003\000\003\000\003\000\ +\003\000\002\000\003\000\002\000\002\000\001\000\001\000\001\000\ \001\000\001\000\001\000\001\000\001\000\001\000\001\000\001\000\ -\001\000\001\000\001\000\001\000\001\000\001\000\001\000\002\000\ -\001\000\003\000\002\000\001\000\001\000\001\000\003\000\002\000\ -\002\000\003\000\001\000\001\000\001\000\001\000\001\000\001\000\ -\001\000\000\000\001\000\003\000\000\000\001\000\003\000\000\000\ -\001\000\003\000\000\000\001\000\003\000\001\000\003\000\001\000\ -\003\000\001\000\003\000\000\000\001\000\003\000\001\000\002\000\ -\003\000\002\000\002\000\002\000" +\001\000\001\000\001\000\001\000\001\000\001\000\001\000\001\000\ +\001\000\002\000\001\000\003\000\002\000\001\000\001\000\001\000\ +\003\000\002\000\002\000\003\000\001\000\001\000\001\000\001\000\ +\001\000\001\000\001\000\000\000\001\000\003\000\000\000\001\000\ +\003\000\000\000\001\000\003\000\000\000\001\000\003\000\001\000\ +\003\000\001\000\003\000\001\000\003\000\000\000\001\000\003\000\ +\001\000\002\000\003\000\002\000\002\000\002\000\002\000" let yydefred = "\000\000\ -\000\000\000\000\000\000\000\000\000\000\203\000\197\000\199\000\ -\201\000\000\000\000\000\209\000\208\000\206\000\202\000\210\000\ -\205\000\212\000\211\000\204\000\200\000\198\000\207\000\074\002\ -\000\000\000\000\000\000\000\000\127\000\128\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\015\002\000\000\218\000\000\000\000\000\000\000\018\002\000\000\ -\000\000\000\000\000\000\000\000\046\001\047\001\014\002\023\002\ -\025\002\027\002\021\002\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\013\002\044\001\045\001\ -\000\000\051\001\000\000\000\000\000\000\030\002\019\002\029\002\ -\000\000\031\002\028\002\016\002\000\000\000\000\017\002\000\000\ -\052\001\022\002\024\002\026\002\020\002\000\000\075\002\032\001\ -\066\001\000\000\107\001\000\000\236\001\224\001\151\001\000\000\ -\243\001\000\000\000\000\000\000\244\001\000\000\000\000\000\000\ -\030\001\031\001\000\000\000\000\000\000\000\000\000\000\000\000\ -\150\001\000\000\000\000\000\000\000\000\201\001\000\000\238\001\ -\239\001\245\001\000\000\000\002\001\002\045\002\043\002\044\002\ -\076\002\000\000\000\000\000\000\050\001\000\000\242\001\000\000\ -\000\000\237\001\000\000\000\000\000\000\118\000\067\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\122\000\124\000\000\000\000\000\ -\123\000\000\000\000\000\000\000\000\000\000\000\134\000\130\000\ -\131\000\196\000\069\000\000\000\087\001\000\000\077\001\000\000\ -\000\000\180\001\181\001\182\001\199\001\000\000\040\000\000\000\ -\000\000\000\000\000\000\075\001\114\001\000\000\000\000\000\000\ -\000\000\074\001\048\001\000\000\000\000\049\001\253\000\000\000\ -\000\000\248\000\249\000\000\000\000\000\000\000\029\001\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\236\000\ -\001\001\240\000\246\000\000\000\000\000\230\000\000\000\241\000\ -\238\000\000\000\000\000\000\000\000\000\224\000\220\000\000\000\ -\221\000\000\000\000\000\000\000\000\000\071\002\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\226\001\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\214\000\000\000\000\000\000\000\000\000\217\001\216\001\212\001\ -\213\001\205\001\214\001\211\001\000\000\202\001\203\001\204\001\ -\206\001\207\001\208\001\209\001\210\001\000\000\219\001\218\001\ -\215\001\000\000\000\000\000\000\000\000\178\001\000\000\254\001\ -\255\001\000\000\179\001\073\001\090\001\200\001\000\000\000\000\ -\000\000\000\000\000\000\010\002\028\000\000\000\000\000\000\000\ -\000\000\008\002\000\000\049\002\011\002\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\235\001\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\230\001\000\000\232\001\234\001\000\000\252\001\ -\040\002\046\002\047\002\000\000\000\000\125\000\059\000\000\000\ -\126\000\036\001\132\000\000\000\037\001\000\000\000\000\155\000\ -\000\000\136\000\137\000\000\000\000\000\154\000\000\000\139\000\ -\000\000\140\000\000\000\153\000\186\000\185\000\135\000\000\000\ -\157\000\000\000\158\000\048\002\121\000\034\000\120\000\119\000\ -\000\000\000\000\156\000\141\000\000\000\000\000\000\000\144\000\ -\000\000\088\001\079\000\000\000\000\000\000\000\000\000\006\002\ -\000\000\253\001\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\192\001\000\000\000\000\000\000\054\000\055\000\ -\000\000\028\001\000\000\027\001\000\000\239\000\237\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\044\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\091\000\000\000\000\001\255\000\000\000\072\002\000\000\ -\000\000\000\000\223\000\000\000\116\001\117\001\000\000\000\000\ -\000\000\054\001\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\198\001\000\000\ -\000\000\000\000\000\000\004\002\219\000\003\002\000\000\000\000\ -\038\000\213\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\024\000\249\001\250\001\246\001\248\001\012\002\000\000\000\000\ -\247\001\000\000\000\000\000\000\005\002\000\000\089\001\053\001\ -\061\001\062\001\000\000\000\000\105\000\000\000\033\001\035\001\ -\000\000\007\002\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\030\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\127\001\149\001\120\001\122\001\ -\000\000\000\000\177\001\000\000\000\000\000\000\000\000\000\000\ -\176\001\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\233\001\000\000\037\002\000\000\036\002\038\002\032\002\048\000\ -\000\000\042\002\129\000\055\002\000\000\000\000\160\000\161\000\ -\000\000\000\000\000\000\152\000\133\000\000\000\143\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\042\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\056\001\006\000\000\000\000\000\000\000\000\000\099\000\000\000\ -\076\001\000\000\000\000\000\000\000\000\000\000\164\001\165\001\ -\166\001\190\001\000\000\000\000\162\001\163\001\191\001\144\001\ -\000\000\145\001\251\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\231\001\000\000\ -\075\000\000\000\000\000\026\001\000\000\000\000\000\000\227\000\ -\000\000\245\000\000\000\226\000\000\000\000\000\234\000\235\000\ -\231\000\233\000\000\000\000\000\000\000\229\000\232\000\115\001\ -\254\000\097\000\000\000\103\000\077\000\000\000\118\001\073\002\ -\216\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\038\001\227\001\228\001\000\000\ -\067\002\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\176\000\002\002\032\000\000\000\ -\000\000\000\000\000\000\009\002\000\000\241\001\128\001\000\000\ -\129\001\000\000\225\001\223\001\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\068\001\000\000\067\001\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\033\002\000\000\ -\000\000\052\000\000\000\000\000\000\000\000\000\022\000\000\000\ -\000\000\002\000\000\000\000\000\138\000\187\000\000\000\000\000\ -\000\000\000\000\000\000\095\000\000\000\000\000\000\000\000\000\ -\000\000\071\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\189\001\142\001\146\001\ -\000\000\140\001\000\000\000\000\161\001\000\000\000\000\000\000\ -\000\000\000\000\160\001\000\000\000\000\000\000\000\000\000\000\ -\000\000\057\000\000\000\000\000\112\001\000\000\113\001\000\000\ -\109\001\108\001\242\000\000\000\063\002\000\000\117\000\000\000\ -\000\000\073\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\124\001\ -\251\001\082\001\174\000\000\000\061\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\046\000\085\001\ -\000\000\000\000\000\000\000\000\057\001\000\000\000\000\000\000\ -\010\000\000\000\000\000\101\000\000\000\130\001\000\000\063\001\ -\064\001\065\001\000\000\069\002\000\000\055\001\020\000\018\000\ -\000\000\087\000\000\000\088\000\000\000\089\000\000\000\086\000\ -\000\000\000\000\049\000\035\002\000\000\000\000\039\002\162\000\ -\000\000\059\001\058\001\000\000\159\000\000\000\142\000\000\000\ -\093\000\000\000\000\000\151\000\115\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\145\000\000\000\105\001\000\000\ -\000\000\103\001\000\000\000\000\000\000\000\000\099\001\000\000\ -\250\000\000\000\000\000\000\000\000\000\111\001\000\000\244\000\ -\228\000\000\000\000\000\119\001\000\000\000\000\000\000\102\001\ -\000\000\000\000\042\001\043\001\039\001\041\001\040\001\061\002\ -\083\001\175\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\184\000\000\000\000\000\000\000\000\000\000\000\000\000\012\000\ -\135\001\000\000\000\000\000\000\125\001\000\000\008\000\000\000\ -\000\000\070\001\071\001\072\001\069\001\050\000\081\000\034\002\ -\107\000\000\000\000\000\164\000\167\000\000\000\000\000\000\000\ -\000\000\000\000\195\000\191\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\091\001\063\000\000\000\000\000\097\001\ -\000\000\000\000\000\000\252\000\111\000\109\000\052\002\000\000\ -\225\000\000\000\092\001\095\001\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\183\000\000\000\000\000\000\000\ -\136\001\000\000\137\001\126\001\084\001\070\002\078\001\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\065\000\000\000\189\000\000\000\000\000\000\000\149\000\150\000\ -\146\000\148\000\147\000\104\001\036\000\000\000\000\000\000\000\ -\000\000\247\000\000\000\094\001\096\001\000\000\000\000\000\000\ -\000\000\000\000\177\000\132\001\000\000\000\000\138\001\000\000\ -\000\000\000\000\172\000\026\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\165\000\000\000\014\001\000\000\000\000\000\000\ -\008\001\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\093\001\000\000\000\000\000\000\ -\000\000\000\000\000\000\133\001\169\000\000\000\020\001\000\000\ -\170\000\173\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\065\002\168\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\002\001\166\000\000\000\ -\188\000\000\000\000\000\192\000\000\000\000\000\000\000\000\000\ -\101\001\181\000\182\000\178\000\180\000\179\000\134\001\000\000\ -\000\000\000\000\000\000\000\000\000\000\018\001\019\001\015\001\ -\017\001\016\001\000\000\012\001\013\001\009\001\011\001\010\001\ -\000\000\000\000\000\000\000\000\000\000\000\000\113\000\194\000\ -\000\000\106\001\098\001\000\000\024\001\025\001\021\001\023\001\ -\022\001\000\000\171\000\058\002\006\001\007\001\003\001\005\001\ -\004\001\190\000\000\000\000\000\004\000\000\000\000\000\193\000\ -\100\001" +\000\000\000\000\000\000\000\000\000\000\000\000\205\000\199\000\ +\201\000\203\000\000\000\000\000\211\000\210\000\208\000\204\000\ +\212\000\207\000\214\000\213\000\206\000\202\000\200\000\209\000\ +\076\002\000\000\000\000\000\000\000\000\129\000\130\000\119\000\ +\077\002\120\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\017\002\000\000\220\000\000\000\ +\000\000\000\000\020\002\000\000\000\000\000\000\000\000\000\000\ +\048\001\049\001\016\002\025\002\027\002\029\002\023\002\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\015\002\046\001\047\001\000\000\053\001\000\000\000\000\ +\000\000\032\002\021\002\031\002\000\000\033\002\030\002\018\002\ +\000\000\000\000\019\002\000\000\054\001\024\002\026\002\028\002\ +\022\002\000\000\078\002\034\001\068\001\000\000\109\001\000\000\ +\238\001\226\001\153\001\000\000\245\001\000\000\000\000\000\000\ +\246\001\000\000\000\000\000\000\032\001\033\001\000\000\000\000\ +\000\000\000\000\000\000\000\000\152\001\000\000\000\000\000\000\ +\000\000\203\001\000\000\240\001\241\001\247\001\000\000\002\002\ +\003\002\047\002\045\002\046\002\079\002\000\000\000\000\000\000\ +\052\001\000\000\244\001\000\000\000\000\239\001\000\000\000\000\ +\000\000\118\000\067\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\124\000\126\000\000\000\000\000\125\000\000\000\000\000\000\000\ +\000\000\000\000\136\000\132\000\133\000\198\000\069\000\000\000\ +\089\001\000\000\079\001\000\000\000\000\182\001\183\001\184\001\ +\201\001\000\000\040\000\000\000\000\000\000\000\000\000\077\001\ +\116\001\000\000\000\000\000\000\000\000\076\001\050\001\000\000\ +\000\000\051\001\255\000\000\000\000\000\250\000\251\000\000\000\ +\000\000\000\000\031\001\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\238\000\003\001\242\000\248\000\000\000\ +\000\000\232\000\000\000\243\000\240\000\000\000\000\000\000\000\ +\000\000\226\000\222\000\000\000\223\000\000\000\000\000\000\000\ +\000\000\073\002\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\228\001\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\216\000\000\000\000\000\000\000\ +\000\000\219\001\218\001\214\001\215\001\207\001\216\001\213\001\ +\000\000\204\001\205\001\206\001\208\001\209\001\210\001\211\001\ +\212\001\000\000\221\001\220\001\217\001\000\000\000\000\000\000\ +\000\000\180\001\000\000\000\002\001\002\000\000\181\001\075\001\ +\092\001\202\001\000\000\000\000\000\000\000\000\000\000\012\002\ +\028\000\000\000\000\000\000\000\000\000\010\002\000\000\051\002\ +\013\002\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\237\001\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\232\001\000\000\ +\234\001\236\001\000\000\254\001\042\002\048\002\049\002\000\000\ +\000\000\127\000\059\000\000\000\128\000\038\001\134\000\000\000\ +\039\001\000\000\000\000\157\000\000\000\138\000\139\000\000\000\ +\000\000\156\000\000\000\141\000\000\000\142\000\000\000\155\000\ +\188\000\187\000\137\000\000\000\159\000\000\000\160\000\050\002\ +\123\000\034\000\122\000\121\000\000\000\000\000\158\000\143\000\ +\000\000\000\000\000\000\146\000\000\000\090\001\079\000\000\000\ +\000\000\000\000\000\000\008\002\000\000\255\001\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\194\001\000\000\ +\000\000\000\000\054\000\055\000\000\000\030\001\000\000\029\001\ +\000\000\241\000\239\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\044\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\091\000\000\000\002\001\ +\001\001\000\000\074\002\000\000\000\000\000\000\225\000\000\000\ +\118\001\119\001\000\000\000\000\000\000\056\001\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\200\001\000\000\000\000\000\000\000\000\006\002\ +\221\000\005\002\000\000\000\000\038\000\215\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\024\000\251\001\252\001\248\001\ +\250\001\014\002\000\000\000\000\249\001\000\000\000\000\000\000\ +\007\002\000\000\091\001\055\001\063\001\064\001\000\000\000\000\ +\105\000\000\000\035\001\037\001\000\000\009\002\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\030\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\129\001\151\001\122\001\124\001\000\000\000\000\179\001\000\000\ +\000\000\000\000\000\000\000\000\178\001\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\235\001\000\000\039\002\000\000\ +\038\002\040\002\034\002\048\000\000\000\044\002\131\000\057\002\ +\000\000\000\000\162\000\163\000\000\000\000\000\000\000\154\000\ +\135\000\000\000\145\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\042\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\058\001\006\000\000\000\000\000\ +\000\000\000\000\099\000\000\000\078\001\000\000\000\000\000\000\ +\000\000\000\000\166\001\167\001\168\001\192\001\000\000\000\000\ +\164\001\165\001\193\001\146\001\000\000\147\001\253\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\233\001\000\000\075\000\000\000\000\000\028\001\ +\000\000\000\000\000\000\229\000\000\000\247\000\000\000\228\000\ +\000\000\000\000\236\000\237\000\233\000\235\000\000\000\000\000\ +\000\000\231\000\234\000\117\001\000\001\097\000\000\000\103\000\ +\077\000\000\000\120\001\075\002\218\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\040\001\229\001\230\001\000\000\069\002\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\178\000\004\002\032\000\000\000\000\000\000\000\000\000\011\002\ +\000\000\243\001\130\001\000\000\131\001\000\000\227\001\225\001\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\070\001\ +\000\000\069\001\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\035\002\000\000\000\000\052\000\000\000\000\000\ +\000\000\000\000\022\000\000\000\000\000\002\000\000\000\000\000\ +\140\000\189\000\000\000\000\000\000\000\000\000\000\000\095\000\ +\000\000\000\000\000\000\000\000\000\000\071\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\191\001\144\001\148\001\000\000\142\001\000\000\000\000\ +\163\001\000\000\000\000\000\000\000\000\000\000\162\001\000\000\ +\000\000\000\000\000\000\000\000\000\000\057\000\000\000\000\000\ +\114\001\000\000\115\001\000\000\111\001\110\001\244\000\000\000\ +\065\002\000\000\117\000\000\000\000\000\073\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\126\001\253\001\084\001\176\000\000\000\ +\061\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\046\000\087\001\000\000\000\000\000\000\000\000\ +\059\001\000\000\000\000\000\000\010\000\000\000\000\000\101\000\ +\000\000\132\001\000\000\065\001\066\001\067\001\000\000\071\002\ +\000\000\057\001\020\000\018\000\000\000\087\000\000\000\088\000\ +\000\000\089\000\000\000\086\000\000\000\000\000\049\000\037\002\ +\000\000\000\000\041\002\164\000\000\000\061\001\060\001\000\000\ +\161\000\000\000\144\000\000\000\093\000\000\000\000\000\153\000\ +\115\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\147\000\000\000\107\001\000\000\000\000\105\001\000\000\000\000\ +\000\000\000\000\101\001\000\000\252\000\000\000\000\000\000\000\ +\000\000\113\001\000\000\246\000\230\000\000\000\000\000\121\001\ +\000\000\000\000\000\000\104\001\000\000\000\000\044\001\045\001\ +\041\001\043\001\042\001\063\002\085\001\177\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\186\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\012\000\137\001\000\000\000\000\000\000\ +\127\001\000\000\008\000\000\000\000\000\072\001\073\001\074\001\ +\071\001\050\000\081\000\036\002\107\000\000\000\000\000\166\000\ +\169\000\000\000\000\000\000\000\000\000\000\000\197\000\193\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\093\001\ +\063\000\000\000\000\000\099\001\000\000\000\000\000\000\254\000\ +\111\000\109\000\054\002\000\000\227\000\000\000\094\001\097\001\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\185\000\000\000\000\000\000\000\138\001\000\000\139\001\128\001\ +\086\001\072\002\080\001\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\065\000\000\000\191\000\000\000\ +\000\000\000\000\151\000\152\000\148\000\150\000\149\000\106\001\ +\036\000\000\000\000\000\000\000\000\000\249\000\000\000\096\001\ +\098\001\000\000\000\000\000\000\000\000\000\000\179\000\134\001\ +\000\000\000\000\140\001\000\000\000\000\000\000\174\000\026\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\167\000\000\000\ +\016\001\000\000\000\000\000\000\010\001\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\095\001\000\000\000\000\000\000\000\000\000\000\000\000\135\001\ +\171\000\000\000\022\001\000\000\172\000\175\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\067\002\170\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\004\001\168\000\000\000\190\000\000\000\000\000\194\000\ +\000\000\000\000\000\000\000\000\103\001\183\000\184\000\180\000\ +\182\000\181\000\136\001\000\000\000\000\000\000\000\000\000\000\ +\000\000\020\001\021\001\017\001\019\001\018\001\000\000\014\001\ +\015\001\011\001\013\001\012\001\000\000\000\000\000\000\000\000\ +\000\000\000\000\113\000\196\000\000\000\108\001\100\001\000\000\ +\026\001\027\001\023\001\025\001\024\001\000\000\173\000\060\002\ +\008\001\009\001\005\001\007\001\006\001\192\000\000\000\000\000\ +\004\000\000\000\000\000\195\000\102\001" -let yydgoto = "\004\000\ -\024\000\035\001\137\000\096\000\092\002\097\000\195\004\154\001\ -\114\002\243\002\098\000\229\002\117\003\023\002\041\002\042\002\ -\254\002\135\003\252\002\014\003\015\003\008\002\009\002\083\004\ -\084\004\052\001\053\001\043\002\044\002\015\002\016\002\099\000\ -\141\001\142\001\062\004\002\002\190\000\191\000\101\002\102\002\ -\206\001\207\001\006\002\011\001\099\001\100\001\253\000\104\001\ -\100\000\207\000\208\000\209\000\159\002\144\000\101\000\214\002\ -\215\002\176\003\177\003\245\003\246\003\025\000\026\000\027\000\ -\028\000\112\002\179\002\226\000\189\001\210\000\227\001\238\000\ -\183\000\148\003\007\003\075\001\228\000\229\000\095\002\096\002\ -\148\001\149\001\239\000\240\000\102\000\028\002\241\000\242\000\ -\103\000\104\000\011\003\105\000\069\003\070\003\071\003\101\004\ -\102\004\116\001\175\000\106\000\148\000\030\000\107\000\176\000\ -\177\000\010\001\108\000\086\002\109\000\036\001\132\001\126\001\ -\120\001\147\001\137\001\139\001\157\003\087\002\200\003\044\004\ -\110\000\046\004\047\004\004\002\134\001\225\002\226\002\111\000\ -\235\001\244\000\181\001\244\002\202\001\076\003\149\000\113\000\ -\114\000\000\001\115\000\116\000\155\003\117\000\185\000\245\000\ -\131\003\119\000\201\003\120\000\121\000\179\003\163\002\072\003\ -\232\001\246\000\122\000\123\000\183\001\184\001\185\001\186\001\ -\124\000\125\000\187\001\126\000\001\001\002\001\127\000\102\001\ -\128\000\129\000\130\000\131\000\132\000\133\000\061\001\009\003\ -\138\000\139\000\108\001\077\003\134\004\209\002" +let yydgoto = "\005\000\ +\025\000\033\000\039\001\141\000\100\000\096\002\101\000\199\004\ +\158\001\118\002\247\002\102\000\233\002\121\003\027\002\045\002\ +\046\002\002\003\139\003\000\003\018\003\019\003\012\002\013\002\ +\087\004\088\004\056\001\057\001\047\002\048\002\019\002\020\002\ +\103\000\145\001\146\001\066\004\006\002\194\000\195\000\105\002\ +\106\002\210\001\211\001\010\002\015\001\103\001\104\001\001\001\ +\108\001\104\000\211\000\212\000\213\000\163\002\148\000\105\000\ +\218\002\219\002\180\003\181\003\249\003\250\003\026\000\027\000\ +\028\000\029\000\116\002\183\002\230\000\193\001\214\000\231\001\ +\242\000\187\000\152\003\011\003\079\001\232\000\233\000\099\002\ +\100\002\152\001\153\001\243\000\244\000\106\000\032\002\245\000\ +\246\000\107\000\108\000\015\003\109\000\073\003\074\003\075\003\ +\105\004\106\004\120\001\179\000\110\000\152\000\031\000\111\000\ +\180\000\181\000\014\001\112\000\090\002\113\000\040\001\136\001\ +\130\001\124\001\151\001\141\001\143\001\161\003\091\002\204\003\ +\048\004\114\000\050\004\051\004\008\002\138\001\229\002\230\002\ +\115\000\239\001\248\000\185\001\248\002\206\001\080\003\153\000\ +\117\000\118\000\004\001\119\000\120\000\159\003\121\000\189\000\ +\249\000\135\003\123\000\205\003\124\000\125\000\183\003\167\002\ +\076\003\236\001\250\000\126\000\127\000\187\001\188\001\189\001\ +\190\001\128\000\129\000\191\001\130\000\005\001\006\001\131\000\ +\106\001\132\000\133\000\134\000\135\000\136\000\137\000\065\001\ +\013\003\142\000\143\000\112\001\081\003\138\004\213\002" -let yysindex = "\207\001\ -\235\004\197\056\172\000\000\000\135\255\000\000\000\000\000\000\ -\000\000\079\081\197\056\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\040\001\235\004\119\005\058\006\000\000\000\000\079\081\079\081\ -\079\081\077\070\079\081\079\081\079\081\079\081\197\056\079\081\ -\000\000\097\065\000\000\181\068\091\066\112\255\000\000\112\255\ -\101\074\207\074\008\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\027\069\181\081\077\070\197\056\197\056\ -\197\056\077\070\049\255\225\076\204\053\000\000\000\000\000\000\ -\077\070\000\000\027\082\197\056\077\070\000\000\000\000\000\000\ -\091\066\000\000\000\000\000\000\079\081\079\081\000\000\197\056\ -\000\000\000\000\000\000\000\000\000\000\072\000\000\000\000\000\ -\000\000\028\000\000\000\217\255\000\000\000\000\000\000\157\000\ -\000\000\141\000\196\000\048\255\000\000\077\070\129\069\249\255\ -\000\000\000\000\197\056\197\056\154\255\170\002\078\255\219\000\ -\000\000\204\000\217\000\037\003\198\000\000\000\245\071\000\000\ -\000\000\000\000\234\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\016\002\060\255\002\001\000\000\204\053\000\000\178\000\ -\079\081\000\000\214\000\201\000\083\001\000\000\000\000\057\255\ -\222\255\135\255\135\255\135\255\254\000\059\001\254\000\057\255\ -\026\000\135\255\045\001\135\255\000\000\000\000\252\000\252\000\ -\000\000\252\000\074\001\135\255\057\255\223\010\000\000\000\000\ -\000\000\000\000\000\000\132\001\000\000\079\081\000\000\000\000\ -\238\002\000\000\000\000\000\000\000\000\025\001\000\000\134\001\ -\118\001\091\001\197\056\000\000\000\000\112\255\112\255\235\255\ -\141\001\000\000\000\000\085\067\062\255\000\000\000\000\112\255\ -\112\255\000\000\000\000\152\001\094\255\218\255\000\000\156\001\ -\079\081\029\255\057\255\085\067\225\076\225\076\153\055\000\000\ -\000\000\000\000\000\000\042\001\101\074\000\000\165\001\000\000\ -\000\000\121\001\086\000\225\076\153\055\000\000\000\000\171\001\ -\000\000\006\001\225\076\239\254\174\001\000\000\149\073\112\255\ -\112\255\132\000\173\001\107\054\070\001\217\255\043\001\180\001\ -\000\000\085\001\123\255\098\001\100\001\217\001\210\001\197\056\ -\000\000\095\001\235\001\218\001\109\001\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\057\075\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\133\078\000\000\000\000\ -\000\000\235\078\111\001\124\001\129\069\000\000\027\082\000\000\ -\000\000\148\001\000\000\000\000\000\000\000\000\093\001\079\081\ -\197\056\197\056\079\081\000\000\000\000\077\070\254\000\057\255\ -\197\056\000\000\121\001\000\000\000\000\105\255\077\070\117\255\ -\079\081\197\056\009\002\009\002\210\067\210\067\197\056\197\056\ -\197\056\197\056\000\000\077\070\210\067\077\070\210\067\210\067\ -\077\070\077\070\077\070\077\070\077\070\077\070\077\070\077\070\ -\077\070\077\070\149\073\149\073\149\073\149\073\149\073\081\079\ -\181\081\110\255\000\000\245\071\000\000\000\000\057\255\000\000\ -\000\000\000\000\000\000\172\000\197\056\000\000\000\000\197\056\ -\000\000\000\000\000\000\103\255\000\000\135\255\103\255\000\000\ -\183\001\000\000\000\000\095\001\135\255\000\000\095\001\000\000\ -\021\001\000\000\240\001\000\000\000\000\000\000\000\000\135\255\ -\000\000\146\001\000\000\000\000\000\000\000\000\000\000\000\000\ -\057\255\151\001\000\000\000\000\037\002\091\082\075\000\000\000\ -\091\066\000\000\000\000\197\056\197\056\197\056\197\056\000\000\ -\079\081\000\000\197\056\170\001\029\002\030\002\077\070\077\070\ -\077\070\179\070\079\081\079\081\079\081\181\081\179\070\179\070\ -\179\070\079\081\179\070\231\069\049\002\047\002\201\001\000\000\ -\236\004\137\001\000\000\095\072\051\002\062\255\000\000\000\000\ -\002\002\000\000\094\255\000\000\002\002\000\000\000\000\039\002\ -\053\002\216\001\212\001\074\002\082\002\234\001\000\000\225\001\ -\041\001\230\001\231\001\236\001\239\001\187\000\064\255\241\001\ -\197\056\000\000\135\255\000\000\000\000\121\001\000\000\225\076\ -\225\076\225\076\000\000\225\076\000\000\000\000\077\070\197\056\ -\225\076\000\000\244\001\198\000\092\002\098\002\077\070\077\070\ -\077\070\253\001\163\075\183\079\029\080\000\002\000\000\084\000\ -\197\056\084\000\013\002\000\000\000\000\000\000\077\070\085\002\ -\000\000\000\000\073\000\075\077\225\076\095\002\201\072\089\002\ -\000\000\000\000\000\000\000\000\000\000\000\000\020\002\216\066\ -\000\000\077\070\182\255\236\254\000\000\207\074\000\000\000\000\ -\000\000\000\000\128\002\017\002\000\000\095\001\000\000\000\000\ -\028\002\000\000\210\067\161\255\210\067\104\002\040\002\121\001\ -\042\002\056\002\014\002\000\000\045\002\188\002\191\002\086\002\ -\194\002\135\002\087\002\219\000\000\000\000\000\000\000\000\000\ -\044\002\052\001\000\000\014\001\204\001\052\001\242\000\034\001\ -\000\000\034\001\107\000\142\001\198\000\122\255\117\002\181\081\ -\000\000\099\002\000\000\037\255\000\000\000\000\000\000\000\000\ -\234\000\000\000\000\000\000\000\076\002\206\002\000\000\000\000\ -\112\255\199\002\091\066\000\000\000\000\112\255\000\000\246\002\ -\135\255\216\066\226\002\135\255\125\002\000\000\119\002\135\255\ -\057\255\118\002\126\002\130\002\131\002\137\002\075\000\240\002\ -\000\000\000\000\147\002\249\002\202\002\150\002\000\000\146\002\ -\000\000\197\056\197\056\080\255\041\255\164\002\000\000\000\000\ -\000\000\000\000\161\002\029\003\000\000\000\000\000\000\000\000\ -\179\070\000\000\000\000\079\068\079\068\179\070\179\070\179\070\ -\179\070\179\070\179\070\179\070\179\070\179\070\179\070\251\073\ -\251\073\251\073\251\073\251\073\079\081\095\072\000\000\079\068\ -\000\000\079\081\197\056\000\000\197\056\225\076\057\255\000\000\ -\057\255\000\000\079\068\000\000\225\076\225\076\000\000\000\000\ -\000\000\000\000\023\003\187\000\079\068\000\000\000\000\000\000\ -\000\000\000\000\024\003\000\000\000\000\076\255\000\000\000\000\ -\000\000\197\056\197\056\056\255\051\255\133\002\000\000\000\000\ -\000\000\000\000\000\000\125\082\000\000\000\000\000\000\166\002\ -\000\000\210\067\231\002\197\056\095\001\000\003\073\000\054\003\ -\040\056\101\074\229\000\235\001\197\056\149\073\204\053\149\073\ -\042\003\192\002\047\073\197\056\000\000\000\000\000\000\001\000\ -\197\002\174\001\079\081\000\000\073\000\000\000\000\000\210\067\ -\000\000\197\056\000\000\000\000\197\056\000\000\210\067\210\067\ -\225\076\197\056\091\066\000\000\091\066\000\000\008\001\066\001\ -\233\001\153\002\149\073\193\002\245\071\000\000\000\000\110\255\ -\006\255\000\000\168\002\045\003\201\072\030\003\000\000\159\002\ -\091\066\000\000\057\003\095\001\000\000\000\000\135\255\040\003\ -\079\081\057\255\010\000\000\000\075\000\075\000\075\000\075\000\ -\075\000\000\000\091\066\238\002\238\002\238\002\238\002\228\064\ -\204\002\172\002\197\056\197\056\197\056\000\000\000\000\000\000\ -\071\003\000\000\056\004\113\001\000\000\145\004\024\002\113\001\ -\147\001\249\001\000\000\249\001\124\000\172\001\137\001\037\000\ -\239\002\000\000\104\002\209\002\000\000\063\003\000\000\222\002\ -\000\000\000\000\000\000\213\002\000\000\081\003\000\000\018\003\ -\079\068\000\000\104\002\197\056\184\002\189\002\010\003\112\255\ -\197\056\224\002\233\002\234\002\235\002\236\002\084\000\000\000\ -\000\000\000\000\000\000\197\056\000\000\095\001\241\002\242\002\ -\243\002\244\002\247\002\218\001\197\056\062\003\000\000\000\000\ -\104\003\129\069\004\003\079\081\000\000\201\072\149\073\006\003\ -\000\000\008\003\210\067\000\000\079\003\000\000\016\003\000\000\ -\000\000\000\000\130\003\000\000\107\003\000\000\000\000\000\000\ -\197\056\000\000\197\056\000\000\197\056\000\000\197\056\000\000\ -\117\002\245\071\000\000\000\000\110\255\037\255\000\000\000\000\ -\076\002\000\000\000\000\222\065\000\000\149\073\000\000\201\072\ -\000\000\012\003\216\066\000\000\000\000\216\066\100\003\132\003\ -\134\003\138\003\139\003\143\003\000\000\057\255\000\000\044\003\ -\228\064\000\000\097\003\197\056\079\081\009\003\000\000\011\003\ -\000\000\251\073\032\003\079\081\079\081\000\000\057\255\000\000\ -\000\000\104\002\036\003\000\000\079\081\197\056\197\056\000\000\ -\150\003\179\003\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\101\074\101\074\101\074\101\074\101\074\144\003\ -\000\000\197\056\060\003\077\070\050\003\198\000\201\072\000\000\ -\000\000\066\003\201\072\210\067\000\000\197\056\000\000\225\076\ -\197\056\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\111\000\031\078\000\000\000\000\174\003\025\071\068\003\ -\027\003\082\003\000\000\000\000\157\003\091\066\091\066\091\066\ -\091\066\091\066\190\003\000\000\000\000\197\056\031\003\000\000\ -\112\255\112\255\239\002\000\000\000\000\000\000\000\000\067\003\ -\000\000\204\003\000\000\000\000\197\056\197\056\218\001\218\001\ -\218\001\218\001\218\001\197\056\000\000\201\072\072\003\084\003\ -\000\000\201\072\000\000\000\000\000\000\000\000\000\000\023\000\ -\135\255\010\055\000\000\092\003\043\000\070\003\189\003\007\000\ -\000\000\031\078\000\000\121\001\135\255\216\066\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\094\003\112\255\194\003\ -\199\003\000\000\197\056\000\000\000\000\191\003\193\003\195\003\ -\196\003\200\003\000\000\000\000\109\003\201\072\000\000\091\066\ -\169\000\091\066\000\000\000\000\023\000\103\003\013\076\131\080\ -\233\080\105\003\000\000\214\082\000\000\169\000\091\066\242\082\ -\000\000\140\000\116\003\112\255\117\003\108\003\208\003\207\255\ -\197\056\221\003\197\056\197\056\000\000\197\056\197\056\197\056\ -\197\056\197\056\201\072\000\000\000\000\075\083\000\000\126\003\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\118\003\ -\121\003\122\003\123\003\124\003\125\003\000\000\000\000\127\003\ -\128\003\129\003\131\003\142\003\103\083\000\000\000\000\226\003\ -\000\000\121\001\216\066\000\000\211\003\145\003\197\056\021\004\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\147\003\ -\148\003\149\003\151\003\155\003\003\004\000\000\000\000\000\000\ -\000\000\000\000\169\000\000\000\000\000\000\000\000\000\000\000\ -\156\003\159\003\160\003\162\003\164\003\216\066\000\000\000\000\ -\246\003\000\000\000\000\112\255\000\000\000\000\000\000\000\000\ -\000\000\091\066\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\216\066\008\004\000\000\175\003\197\056\000\000\ -\000\000" +let yysindex = "\105\003\ +\061\006\035\004\046\057\010\001\000\000\020\255\000\000\000\000\ +\000\000\000\000\078\081\046\057\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\099\000\061\006\199\008\216\007\000\000\000\000\000\000\ +\000\000\000\000\078\081\078\081\078\081\182\070\078\081\078\081\ +\078\081\078\081\046\057\078\081\000\000\202\065\000\000\030\069\ +\196\066\027\000\000\000\027\000\233\009\206\074\076\255\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\132\069\ +\180\081\182\070\046\057\046\057\046\057\182\070\196\255\224\076\ +\053\054\000\000\000\000\000\000\182\070\000\000\026\082\046\057\ +\182\070\000\000\000\000\000\000\196\066\000\000\000\000\000\000\ +\078\081\078\081\000\000\046\057\000\000\000\000\000\000\000\000\ +\000\000\140\255\000\000\000\000\000\000\122\000\000\000\008\255\ +\000\000\000\000\000\000\160\255\000\000\112\255\179\255\034\255\ +\000\000\182\070\234\069\005\000\000\000\000\000\046\057\046\057\ +\059\255\154\002\081\255\007\000\000\000\028\000\048\000\225\000\ +\062\000\000\000\094\072\000\000\000\000\000\000\113\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\125\001\045\255\134\000\ +\000\000\053\054\000\000\090\000\078\081\000\000\117\000\192\000\ +\027\001\000\000\000\000\014\000\017\000\020\255\020\255\020\255\ +\193\000\012\001\193\000\014\000\032\000\020\255\016\001\020\255\ +\000\000\000\000\246\000\246\000\000\000\246\000\036\001\020\255\ +\014\000\147\005\000\000\000\000\000\000\000\000\000\000\117\001\ +\000\000\078\081\000\000\000\000\189\002\000\000\000\000\000\000\ +\000\000\009\001\000\000\115\001\105\001\079\001\046\057\000\000\ +\000\000\027\000\027\000\244\255\142\001\000\000\000\000\190\067\ +\053\255\000\000\000\000\027\000\027\000\000\000\000\000\120\001\ +\089\255\101\255\000\000\135\001\078\081\107\255\014\000\190\067\ +\224\076\224\076\002\056\000\000\000\000\000\000\000\000\047\001\ +\233\009\000\000\143\001\000\000\000\000\127\001\042\000\224\076\ +\002\056\000\000\000\000\166\001\000\000\062\001\224\076\000\255\ +\168\001\000\000\254\073\027\000\027\000\242\000\164\001\212\054\ +\074\001\008\255\024\001\162\001\000\000\070\001\235\000\087\001\ +\103\001\220\001\213\001\046\057\000\000\102\001\242\001\226\001\ +\112\001\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\056\075\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\132\078\000\000\000\000\000\000\234\078\108\001\113\001\ +\234\069\000\000\026\082\000\000\000\000\126\001\000\000\000\000\ +\000\000\000\000\093\001\078\081\046\057\046\057\078\081\000\000\ +\000\000\182\070\193\000\014\000\046\057\000\000\127\001\000\000\ +\000\000\105\255\182\070\134\255\078\081\046\057\004\002\004\002\ +\059\068\059\068\046\057\046\057\046\057\046\057\000\000\182\070\ +\059\068\182\070\059\068\059\068\182\070\182\070\182\070\182\070\ +\182\070\182\070\182\070\182\070\182\070\182\070\254\073\254\073\ +\254\073\254\073\254\073\080\079\180\081\068\255\000\000\094\072\ +\000\000\000\000\014\000\000\000\000\000\000\000\000\000\010\001\ +\046\057\000\000\000\000\046\057\000\000\000\000\000\000\095\255\ +\000\000\020\255\095\255\000\000\178\001\000\000\000\000\102\001\ +\020\255\000\000\102\001\000\000\011\001\000\000\234\001\000\000\ +\000\000\000\000\000\000\020\255\000\000\131\001\000\000\000\000\ +\000\000\000\000\000\000\000\000\014\000\139\001\000\000\000\000\ +\029\002\078\008\231\000\000\000\196\066\000\000\000\000\046\057\ +\046\057\046\057\046\057\000\000\078\081\000\000\046\057\172\001\ +\008\002\017\002\182\070\182\070\182\070\028\071\078\081\078\081\ +\078\081\180\081\028\071\028\071\028\071\078\081\028\071\080\070\ +\018\002\022\002\180\001\000\000\215\002\098\000\000\000\200\072\ +\030\002\053\255\000\000\000\000\237\001\000\000\089\255\000\000\ +\237\001\000\000\000\000\021\002\033\002\196\001\186\001\057\002\ +\052\002\205\001\000\000\201\001\083\001\197\001\209\001\215\001\ +\216\001\119\000\046\255\218\001\046\057\000\000\020\255\000\000\ +\000\000\127\001\000\000\224\076\224\076\224\076\000\000\224\076\ +\000\000\000\000\182\070\046\057\224\076\000\000\228\001\062\000\ +\070\002\073\002\182\070\182\070\182\070\231\001\162\075\182\079\ +\028\080\236\001\000\000\219\255\046\057\219\255\255\001\000\000\ +\000\000\000\000\182\070\072\002\000\000\000\000\169\000\074\077\ +\224\076\075\002\050\073\085\002\000\000\000\000\000\000\000\000\ +\000\000\000\000\248\001\065\067\000\000\182\070\109\255\195\255\ +\000\000\206\074\000\000\000\000\000\000\000\000\109\002\235\001\ +\000\000\102\001\000\000\000\000\001\002\000\000\059\068\148\255\ +\059\068\067\002\006\002\127\001\005\002\016\002\229\001\000\000\ +\011\002\136\002\145\002\049\002\162\002\101\002\068\002\007\000\ +\000\000\000\000\000\000\000\000\059\001\116\001\000\000\040\003\ +\166\002\116\001\026\001\015\001\000\000\015\001\088\000\120\000\ +\062\000\180\255\100\002\180\081\000\000\081\002\000\000\088\255\ +\000\000\000\000\000\000\000\000\113\000\000\000\000\000\000\000\ +\059\002\194\002\000\000\000\000\027\000\177\002\196\066\000\000\ +\000\000\027\000\000\000\216\002\020\255\065\067\201\002\020\255\ +\105\002\000\000\098\002\020\255\014\000\092\002\106\002\113\002\ +\115\002\116\002\231\000\220\002\000\000\000\000\118\002\221\002\ +\174\002\120\002\000\000\122\002\000\000\046\057\046\057\071\255\ +\025\255\123\002\000\000\000\000\000\000\000\000\149\002\001\003\ +\000\000\000\000\000\000\000\000\028\071\000\000\000\000\184\068\ +\184\068\028\071\028\071\028\071\028\071\028\071\028\071\028\071\ +\028\071\028\071\028\071\100\074\100\074\100\074\100\074\100\074\ +\078\081\200\072\000\000\184\068\000\000\078\081\046\057\000\000\ +\046\057\224\076\014\000\000\000\014\000\000\000\184\068\000\000\ +\224\076\224\076\000\000\000\000\000\000\000\000\000\003\119\000\ +\184\068\000\000\000\000\000\000\000\000\000\000\004\003\000\000\ +\000\000\031\255\000\000\000\000\000\000\046\057\046\057\033\255\ +\035\255\126\002\000\000\000\000\000\000\000\000\000\000\193\082\ +\000\000\000\000\000\000\146\002\000\000\059\068\205\002\046\057\ +\102\001\242\002\169\000\042\003\145\056\233\009\035\001\242\001\ +\046\057\254\073\053\054\254\073\025\003\176\002\152\073\046\057\ +\000\000\000\000\000\000\185\255\179\002\168\001\078\081\000\000\ +\169\000\000\000\000\000\059\068\000\000\046\057\000\000\000\000\ +\046\057\000\000\059\068\059\068\224\076\046\057\196\066\000\000\ +\196\066\000\000\058\001\208\001\088\002\141\002\254\073\180\002\ +\094\072\000\000\000\000\068\255\234\254\000\000\147\002\032\003\ +\050\073\009\003\000\000\153\002\196\066\000\000\045\003\102\001\ +\000\000\000\000\020\255\030\003\078\081\014\000\146\000\000\000\ +\231\000\231\000\231\000\231\000\231\000\000\000\196\066\189\002\ +\189\002\189\002\189\002\077\065\195\002\169\002\046\057\046\057\ +\046\057\000\000\000\000\000\000\060\003\000\000\047\002\253\002\ +\000\000\087\003\186\002\253\002\043\001\088\001\000\000\088\001\ +\108\000\189\000\098\000\192\255\234\002\000\000\067\002\203\002\ +\000\000\058\003\000\000\217\002\000\000\000\000\000\000\207\002\ +\000\000\077\003\000\000\018\003\184\068\000\000\067\002\046\057\ +\190\002\191\002\015\003\027\000\046\057\230\002\236\002\237\002\ +\238\002\241\002\219\255\000\000\000\000\000\000\000\000\046\057\ +\000\000\102\001\244\002\246\002\252\002\254\002\002\003\226\001\ +\046\057\096\003\000\000\000\000\123\003\234\069\019\003\078\081\ +\000\000\050\073\254\073\021\003\000\000\022\003\059\068\000\000\ +\094\003\000\000\027\003\000\000\000\000\000\000\138\003\000\000\ +\119\003\000\000\000\000\000\000\046\057\000\000\046\057\000\000\ +\046\057\000\000\046\057\000\000\100\002\094\072\000\000\000\000\ +\068\255\088\255\000\000\000\000\059\002\000\000\000\000\071\066\ +\000\000\254\073\000\000\050\073\000\000\026\003\065\067\000\000\ +\000\000\065\067\108\003\142\003\146\003\147\003\148\003\150\003\ +\000\000\014\000\000\000\047\003\077\065\000\000\111\003\046\057\ +\078\081\006\003\000\000\014\003\000\000\100\074\048\003\078\081\ +\078\081\000\000\014\000\000\000\000\000\067\002\059\003\000\000\ +\078\081\046\057\046\057\000\000\153\003\192\003\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\233\009\233\009\ +\233\009\233\009\233\009\158\003\000\000\046\057\076\003\182\070\ +\066\003\062\000\050\073\000\000\000\000\081\003\050\073\059\068\ +\000\000\046\057\000\000\224\076\046\057\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\066\001\030\078\000\000\ +\000\000\197\003\130\071\085\003\044\003\103\003\000\000\000\000\ +\174\003\196\066\196\066\196\066\196\066\196\066\207\003\000\000\ +\000\000\046\057\050\003\000\000\027\000\027\000\234\002\000\000\ +\000\000\000\000\000\000\080\003\000\000\219\003\000\000\000\000\ +\046\057\046\057\226\001\226\001\226\001\226\001\226\001\046\057\ +\000\000\050\073\089\003\098\003\000\000\050\073\000\000\000\000\ +\000\000\000\000\000\000\022\000\020\255\115\055\000\000\101\003\ +\251\255\095\003\210\003\177\000\000\000\030\078\000\000\127\001\ +\020\255\065\067\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\110\003\027\000\204\003\209\003\000\000\046\057\000\000\ +\000\000\198\003\200\003\202\003\203\003\205\003\000\000\000\000\ +\112\003\050\073\000\000\196\066\222\000\196\066\000\000\000\000\ +\022\000\115\003\012\076\130\080\232\080\118\003\000\000\221\082\ +\000\000\222\000\196\066\054\083\000\000\103\000\120\003\027\000\ +\122\003\107\003\216\003\024\000\046\057\227\003\046\057\046\057\ +\000\000\046\057\046\057\046\057\046\057\046\057\050\073\000\000\ +\000\000\082\083\000\000\134\003\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\129\003\130\003\131\003\132\003\133\003\ +\144\003\000\000\000\000\152\003\154\003\155\003\156\003\159\003\ +\171\083\000\000\000\000\230\003\000\000\127\001\065\067\000\000\ +\232\003\157\003\046\057\025\004\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\167\003\168\003\170\003\172\003\173\003\ +\012\004\000\000\000\000\000\000\000\000\000\000\222\000\000\000\ +\000\000\000\000\000\000\000\000\176\003\177\003\178\003\180\003\ +\181\003\065\067\000\000\000\000\014\004\000\000\000\000\027\000\ +\000\000\000\000\000\000\000\000\000\000\196\066\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\065\067\033\004\ +\000\000\194\003\046\057\000\000\000\000" let yyrindex = "\000\000\ -\158\003\000\000\000\000\000\000\040\053\000\000\000\000\000\000\ -\000\000\176\003\181\003\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\158\003\000\000\231\007\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\180\003\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\023\004\000\000\023\004\ -\000\000\000\000\127\071\211\014\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\181\003\181\003\ -\039\004\000\000\119\076\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\181\003\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\205\016\000\000\000\000\ -\000\000\178\045\000\000\218\012\000\000\000\000\000\000\227\035\ -\000\000\000\000\130\013\124\015\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\030\039\192\020\243\011\140\036\ -\000\000\119\038\221\037\146\034\166\028\000\000\179\024\000\000\ -\000\000\000\000\043\014\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\184\039\000\000\075\255\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\119\076\000\000\119\076\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\047\051\047\051\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\082\040\000\000\236\040\000\000\036\016\ -\117\017\000\000\000\000\000\000\000\000\000\000\000\000\185\003\ -\000\000\000\000\000\000\000\000\000\000\023\004\023\004\000\000\ -\186\000\000\000\000\000\000\000\000\000\000\000\000\000\097\047\ -\097\047\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\175\000\000\000\000\000\186\003\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\165\255\000\000\255\047\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\253\254\ -\000\000\086\003\063\001\000\000\079\043\000\000\000\000\023\004\ -\023\004\000\000\186\000\000\000\000\000\142\058\000\000\167\000\ -\000\000\188\003\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\181\077\006\004\137\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\202\003\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\119\076\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\017\000\120\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\179\024\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\070\005\000\000\000\000\000\000\081\255\ -\000\000\000\000\000\000\149\048\000\000\000\000\033\004\000\000\ -\148\051\000\000\000\000\181\077\000\000\000\000\181\077\000\000\ -\249\051\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\203\003\000\000\000\000\000\000\094\052\000\000\054\004\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\115\255\ -\253\255\242\004\000\000\010\024\000\000\062\004\000\000\000\000\ -\098\057\000\000\000\000\000\000\098\057\000\000\000\000\210\000\ -\000\000\000\000\206\003\000\000\000\000\000\000\000\000\213\003\ -\000\000\000\000\000\000\000\000\000\000\102\255\000\000\000\000\ +\161\003\114\082\000\000\000\000\000\000\132\012\000\000\000\000\ +\000\000\000\000\196\003\199\003\000\000\000\000\000\000\000\000\ \000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\161\003\000\000\114\082\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\195\003\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\038\004\000\000\038\004\000\000\000\000\232\071\153\015\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\199\003\199\003\058\004\000\000\118\076\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\199\003\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\146\017\000\000\000\000\000\000\016\047\000\000\154\013\ +\000\000\000\000\000\000\169\036\000\000\000\000\070\014\065\016\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\124\040\133\021\148\038\081\037\000\000\213\039\059\039\088\035\ +\107\029\000\000\120\025\000\000\000\000\000\000\240\014\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\022\041\000\000\054\255\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\118\076\000\000\118\076\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\253\051\253\051\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\176\041\ +\000\000\074\042\000\000\234\016\059\018\000\000\000\000\000\000\ +\000\000\000\000\000\000\206\003\000\000\000\000\000\000\000\000\ +\000\000\038\004\038\004\000\000\019\001\000\000\000\000\000\000\ +\000\000\000\000\000\000\047\048\047\048\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\002\001\000\000\000\000\ +\214\003\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\243\255\000\000\205\048\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\197\255\000\000\122\004\061\001\000\000\ +\173\044\000\000\000\000\038\004\038\004\000\000\019\001\000\000\ +\000\000\247\058\000\000\095\000\000\000\211\003\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\180\077\019\004\226\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\189\003\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\118\076\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\165\000\074\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\120\025\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\077\005\ +\000\000\000\000\000\000\070\255\000\000\000\000\000\000\099\049\ +\000\000\000\000\039\004\000\000\098\052\000\000\000\000\180\077\ +\000\000\000\000\180\077\000\000\199\052\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\215\003\000\000\000\000\000\000\ +\044\053\000\000\057\004\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\122\255\188\255\104\004\000\000\208\024\ +\000\000\067\004\000\000\000\000\203\057\000\000\000\000\000\000\ +\203\057\000\000\000\000\130\000\000\000\000\000\220\003\000\000\ +\000\000\000\000\000\000\218\003\000\000\000\000\000\000\000\000\ +\000\000\091\255\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\084\004\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\166\047\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\224\003\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\041\004\000\000\ +\000\000\000\000\000\000\108\255\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\201\003\000\000\ +\000\000\180\077\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\033\026\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\228\042\126\043\000\000\000\000\000\000\000\000\250\037\ +\000\000\000\000\000\000\000\000\000\036\182\032\000\000\175\034\ +\007\034\094\033\013\032\188\030\000\000\101\031\195\028\000\000\ +\020\030\026\028\201\026\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\240\014\000\000\000\000\000\000\ +\000\000\217\049\000\000\000\000\187\003\000\000\000\000\000\000\ +\000\000\077\004\000\000\145\053\000\000\000\000\000\000\000\000\ +\000\000\000\000\233\003\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\057\004\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\052\001\ +\000\000\182\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\208\024\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\241\003\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\091\004\ +\000\000\000\000\000\000\000\000\000\000\000\000\045\005\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\052\001\ +\000\000\182\000\123\062\017\063\167\063\061\064\211\064\000\000\ +\000\000\000\000\000\000\243\003\000\000\000\000\000\000\000\000\ +\180\077\000\000\041\004\000\000\000\000\130\001\000\000\019\004\ +\000\000\000\000\000\000\000\000\255\000\132\011\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\071\045\000\000\000\000\ +\041\004\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\209\000\000\000\000\000\024\044\000\000\000\000\000\000\ +\000\000\000\000\046\022\214\022\127\023\039\024\000\000\000\000\ +\120\025\144\000\000\000\000\000\000\000\000\000\000\000\221\003\ +\000\000\051\051\000\000\000\000\000\000\000\000\000\000\180\077\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\057\004\057\004\057\004\057\004\057\004\000\000\000\000\227\018\ +\140\019\052\020\221\020\244\003\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\080\000\249\001\ +\000\000\016\000\118\255\006\005\185\001\250\002\000\000\186\003\ +\091\002\000\000\251\006\153\005\212\010\000\000\235\003\069\255\ +\000\000\249\003\000\000\000\000\000\000\000\000\000\000\251\003\ +\000\000\000\000\000\000\000\000\000\000\000\000\235\003\000\000\ \000\000\000\000\000\000\038\004\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\072\046\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\217\003\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\034\004\000\000\000\000\000\000\000\000\064\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\192\003\000\000\000\000\181\077\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\091\025\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\134\041\032\042\000\000\ -\000\000\000\000\000\000\052\037\000\000\000\000\000\000\000\000\ -\059\035\240\031\000\000\234\033\065\033\153\032\072\031\247\029\ -\000\000\159\030\253\027\000\000\078\029\085\027\004\026\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\043\014\000\000\000\000\000\000\000\000\011\049\000\000\000\000\ -\178\003\000\000\000\000\000\000\000\000\073\004\000\000\195\052\ -\000\000\000\000\000\000\000\000\000\000\000\000\227\003\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\054\004\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\213\000\000\000\224\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\010\024\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\231\003\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\081\004\000\000\000\000\000\000\000\000\ -\000\000\000\000\082\004\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\213\000\000\000\224\000\018\062\168\062\ -\062\063\212\063\106\064\000\000\000\000\000\000\000\000\235\003\ -\000\000\000\000\000\000\000\000\181\077\000\000\034\004\000\000\ -\000\000\157\001\000\000\006\004\000\000\000\000\000\000\000\000\ -\191\000\204\046\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\233\043\000\000\000\000\034\004\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\166\000\000\000\000\000\ -\186\042\000\000\000\000\000\000\000\000\000\000\104\021\017\022\ -\185\022\098\023\000\000\000\000\179\024\036\000\000\000\000\000\ -\000\000\000\000\000\000\215\003\000\000\101\050\000\000\000\000\ -\000\000\000\000\000\000\181\077\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\054\004\054\004\054\004\054\004\ -\054\004\000\000\000\000\030\018\198\018\111\019\023\020\248\003\ +\000\000\000\000\224\003\000\000\000\000\000\000\000\000\000\000\ +\000\000\180\077\000\000\000\000\000\000\000\000\000\000\082\004\ \000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\005\000\013\000\000\000\151\255\183\255\090\002\ -\225\000\143\002\000\000\248\006\101\004\000\000\200\005\035\008\ -\228\006\000\000\236\003\143\255\000\000\250\003\000\000\000\000\ -\000\000\000\000\000\000\251\003\000\000\000\000\000\000\000\000\ -\000\000\000\000\236\003\000\000\000\000\000\000\000\000\023\004\ -\000\000\000\000\000\000\000\000\000\000\000\000\217\003\000\000\ -\000\000\000\000\000\000\000\000\000\000\181\077\000\000\000\000\ -\000\000\000\000\000\000\085\004\000\000\000\000\000\000\000\000\ \000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\220\045\000\000\ +\118\046\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\114\027\120\025\000\000\000\000\ +\050\255\000\000\000\000\000\000\000\000\000\000\000\000\152\051\ \000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\126\044\000\000\024\045\000\000\000\000\000\000\ \000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\172\026\179\024\000\000\000\000\045\255\000\000\000\000\000\000\ -\000\000\000\000\000\000\202\050\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\244\003\000\000\000\000\000\000\ +\022\041\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\241\003\000\000\000\000\235\003\000\000\000\000\ +\121\004\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\161\001\210\001\ +\000\002\013\002\062\002\000\000\000\000\000\000\000\000\000\000\ +\000\000\078\007\000\000\000\000\000\000\000\000\000\000\000\000\ \000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ \000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\248\003\000\000\000\000\000\000\184\039\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\231\003\000\000\ -\000\000\236\003\000\000\000\000\124\004\000\000\000\000\000\000\ +\000\000\152\051\000\000\000\000\000\000\000\000\000\000\000\000\ \000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\068\002\078\002\211\002\225\002\017\003\000\000\ -\000\000\000\000\000\000\000\000\000\000\009\004\000\000\000\000\ +\000\000\004\004\000\000\000\000\038\004\038\004\011\012\000\000\ \000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\082\004\082\004\082\004\082\004\082\004\000\000\ \000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\202\050\000\000\000\000\ +\000\000\000\000\000\000\083\050\000\000\000\000\097\058\000\000\ +\000\000\006\004\000\000\000\000\000\000\000\000\000\000\000\000\ \000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\001\004\000\000\000\000\ -\023\004\023\004\079\011\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\085\004\085\004\ -\085\004\085\004\085\004\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\133\049\ -\000\000\000\000\248\057\000\000\000\000\002\004\000\000\000\000\ +\000\000\000\000\038\004\000\000\000\000\000\000\000\000\000\000\ \000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\023\004\000\000\ \000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\083\050\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\007\004\000\000\000\000\000\000\000\000\000\000\039\004\ +\000\000\008\004\000\000\000\000\000\000\000\000\000\000\000\000\ \000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\133\049\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\004\004\000\000\000\000\ -\000\000\000\000\000\000\033\004\000\000\005\004\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\141\059\035\060\ +\185\060\079\061\229\061\000\000\000\000\000\000\000\000\000\000\ +\011\004\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ \000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ \000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\036\059\186\059\080\060\230\060\124\061\000\000\ -\000\000\000\000\000\000\000\000\010\004\000\000\000\000\000\000\ \000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\205\050\000\000\000\000\000\000\000\000\000\000\007\004\000\000\ \000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\038\004\ \000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\255\049\000\000\000\000\000\000\ -\000\000\000\000\004\004\000\000\000\000\000\000\000\000\000\000\ \000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\023\004\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000" +\000\000\000\000\000\000\000\000\000\000" let yygindex = "\000\000\ -\000\000\209\000\000\000\243\007\000\000\050\000\000\000\000\000\ -\000\000\000\000\012\000\000\000\000\000\110\253\000\000\000\000\ -\000\000\132\002\000\000\000\000\000\000\220\252\169\002\050\001\ -\000\000\000\000\169\255\070\004\000\000\000\000\000\000\001\254\ -\227\004\226\004\000\000\000\000\000\000\237\003\000\000\127\002\ -\000\000\251\002\187\002\166\003\168\254\116\254\200\254\088\003\ -\213\255\154\001\000\000\003\255\250\002\121\255\046\002\099\254\ -\000\000\251\001\000\000\186\001\000\000\134\005\000\000\145\005\ -\000\000\204\253\002\003\052\255\249\003\090\255\212\003\212\255\ -\005\005\041\002\171\254\064\003\059\255\000\000\171\002\247\254\ -\083\003\040\005\225\003\191\255\096\000\215\002\087\254\070\255\ -\002\255\129\005\057\002\109\006\008\002\000\000\010\002\060\001\ -\000\000\137\254\000\000\014\000\238\255\000\000\251\008\000\000\ -\000\000\177\255\136\009\093\004\203\009\196\255\000\000\054\005\ -\000\000\000\000\000\000\000\000\000\000\000\000\224\255\248\252\ -\209\003\004\252\000\000\000\000\000\000\031\253\011\255\031\005\ -\000\000\000\000\042\255\027\253\000\000\254\004\004\002\224\254\ -\010\255\000\000\000\000\000\000\000\000\074\000\254\255\198\003\ -\000\000\216\255\040\252\239\255\246\254\000\000\024\004\000\000\ -\035\003\027\255\148\004\204\254\000\000\141\254\139\000\155\253\ -\016\000\055\004\167\004\000\000\013\004\034\254\158\008\000\000\ -\192\255\000\000\000\000\000\000\000\000\000\000\000\000\088\002\ -\134\004\000\000\000\000\059\002\086\001\163\002" +\000\000\000\000\241\007\000\000\242\007\000\000\049\000\000\000\ +\000\000\000\000\000\000\011\000\000\000\000\000\118\253\000\000\ +\000\000\000\000\137\002\000\000\000\000\000\000\250\252\173\002\ +\053\001\000\000\000\000\158\255\069\004\000\000\000\000\000\000\ +\072\254\229\004\230\004\000\000\000\000\000\000\238\003\000\000\ +\125\002\000\000\235\002\192\002\149\003\167\254\093\254\185\254\ +\072\003\193\255\140\001\000\000\241\254\255\002\110\255\200\000\ +\033\254\000\000\239\001\000\000\165\001\000\000\140\005\166\005\ +\142\005\000\000\188\253\247\002\042\255\231\003\073\255\212\003\ +\208\255\242\004\023\002\165\254\002\002\051\255\000\000\150\002\ +\002\255\069\003\023\005\222\003\190\255\018\000\200\002\068\254\ +\055\255\229\254\116\005\034\002\108\006\250\001\000\000\254\001\ +\045\001\000\000\114\254\000\000\019\000\238\255\000\000\038\010\ +\000\000\000\000\183\255\184\010\070\004\106\010\150\255\000\000\ +\031\005\000\000\000\000\000\000\000\000\000\000\000\000\223\255\ +\100\253\208\003\091\252\000\000\000\000\000\000\044\253\009\255\ +\030\005\000\000\000\000\056\255\033\253\000\000\241\004\003\002\ +\214\254\240\254\000\000\000\000\000\000\000\000\073\000\253\255\ +\226\003\000\000\235\255\041\252\247\255\209\254\000\000\017\004\ +\000\000\028\003\026\255\128\004\187\254\000\000\129\254\137\255\ +\171\253\046\000\054\004\144\004\000\000\228\003\024\254\239\007\ +\000\000\178\255\000\000\000\000\000\000\000\000\000\000\000\000\ +\061\002\111\004\000\000\000\000\032\002\055\001\129\002" -let yytablesize = 21756 -let yytable = "\118\000\ -\003\002\234\001\012\001\200\000\227\000\204\001\223\001\146\000\ -\118\000\111\001\041\001\080\002\079\002\212\000\029\000\216\000\ -\231\002\254\000\214\001\250\000\060\001\003\001\147\000\031\002\ -\053\002\103\002\201\000\208\002\146\000\146\000\146\000\218\001\ -\146\000\146\000\146\000\209\001\118\000\146\000\190\001\029\000\ -\074\002\029\000\251\000\154\003\004\001\005\001\006\001\158\002\ -\064\004\186\000\061\003\062\003\063\003\064\003\065\003\029\002\ -\188\002\042\001\034\003\132\002\118\000\118\000\118\000\136\002\ -\138\002\032\002\118\000\054\002\076\001\062\001\064\001\216\003\ -\040\001\118\000\147\000\147\000\147\000\124\001\076\001\127\001\ -\078\001\057\001\146\000\146\000\181\002\118\000\141\000\147\000\ -\038\001\054\002\043\000\196\000\043\001\187\003\202\000\053\002\ -\075\002\133\004\022\003\058\001\059\001\247\000\106\004\076\001\ -\078\001\076\001\141\000\229\001\150\003\195\003\051\000\203\000\ -\118\000\118\000\093\002\197\000\128\001\131\002\197\000\106\001\ -\076\002\096\000\240\000\240\000\230\001\240\000\074\000\068\001\ -\141\001\092\001\044\001\043\000\076\001\192\000\164\003\051\001\ -\251\001\064\000\096\000\118\000\043\000\151\003\146\000\203\000\ -\076\001\008\001\074\000\080\000\225\003\231\001\076\001\051\000\ -\203\000\111\002\197\000\208\001\255\000\215\001\126\002\141\000\ -\051\000\203\000\216\001\096\000\155\001\165\001\166\001\058\001\ -\059\001\075\002\110\000\215\001\204\000\155\001\133\004\200\001\ -\216\001\009\001\064\000\146\000\080\000\204\000\107\001\141\001\ -\227\000\090\000\205\000\064\000\076\001\199\001\249\003\246\001\ -\118\000\076\002\069\001\205\000\156\001\077\002\228\001\182\002\ -\229\001\054\002\077\001\051\001\090\000\156\001\164\001\053\002\ -\044\003\076\001\095\000\012\001\008\003\043\003\146\000\237\001\ -\238\001\230\001\088\003\155\001\016\004\041\001\206\000\087\003\ -\074\000\224\003\095\001\204\004\032\002\035\002\198\002\030\002\ -\168\003\169\003\170\003\171\003\172\003\240\000\239\002\143\001\ -\241\002\037\002\070\004\071\004\072\004\073\004\074\004\193\000\ -\155\001\118\000\141\001\156\001\019\002\033\004\206\000\089\000\ -\211\002\035\004\085\002\192\002\080\003\118\000\167\001\206\000\ -\089\000\180\002\220\001\203\000\110\000\065\001\077\002\004\003\ -\206\000\158\002\153\001\155\001\237\002\048\003\036\002\190\001\ -\156\001\156\001\152\001\040\001\076\001\240\002\155\001\146\000\ -\168\001\090\000\118\001\152\001\040\001\064\000\187\002\150\000\ -\047\001\099\004\070\001\080\004\153\002\146\000\118\000\118\000\ -\146\000\166\003\037\002\156\001\152\001\254\000\118\000\074\000\ -\141\000\101\003\052\002\081\002\076\004\233\002\156\001\118\000\ -\079\004\220\001\009\000\067\001\118\000\118\000\118\000\118\000\ -\120\004\153\001\213\002\066\001\148\004\039\002\167\003\125\003\ -\247\000\152\001\064\000\088\002\011\004\036\001\088\002\149\004\ -\048\001\110\002\096\004\081\004\141\000\084\002\125\001\036\001\ -\057\002\058\002\059\002\060\002\061\002\062\002\063\002\064\002\ -\065\002\066\002\118\000\091\001\116\004\118\000\152\001\152\001\ -\152\001\111\002\092\001\074\000\051\001\082\004\009\000\036\001\ -\083\002\220\001\206\000\147\000\152\002\240\002\092\004\125\002\ -\208\002\153\001\254\000\153\002\220\001\141\000\031\002\141\000\ -\122\003\152\001\038\002\023\000\153\001\156\002\046\002\047\002\ -\204\000\159\004\010\000\011\000\152\001\124\002\013\000\015\000\ -\055\002\118\000\118\000\118\000\118\000\096\003\205\000\239\001\ -\118\000\064\000\103\002\251\003\049\001\050\001\252\003\204\002\ -\146\000\146\000\146\000\038\002\147\003\074\000\120\002\146\000\ -\023\000\012\001\149\003\041\001\054\001\049\003\050\003\228\001\ -\029\000\240\001\064\000\126\003\060\001\141\000\190\002\023\000\ -\255\000\011\000\219\002\012\001\036\001\043\000\091\001\197\002\ -\074\000\067\003\113\002\099\003\205\002\092\001\205\002\001\000\ -\002\000\003\000\094\001\095\001\078\003\148\001\118\000\043\000\ -\000\002\051\000\141\000\093\001\055\001\196\002\083\003\141\004\ -\217\000\056\001\197\000\155\002\156\002\118\000\011\000\180\002\ -\113\003\232\002\115\003\051\000\141\000\120\003\158\001\148\001\ -\147\001\040\001\146\000\078\001\064\000\011\000\118\000\158\001\ -\149\001\007\002\217\000\149\001\118\004\013\000\015\000\079\001\ -\192\000\025\002\026\002\134\000\207\002\112\000\064\000\103\001\ -\158\001\033\002\147\001\029\000\109\003\255\000\080\001\105\001\ -\216\002\135\000\136\000\132\003\108\003\149\001\109\001\048\002\ -\049\002\050\002\051\002\225\001\254\000\226\001\008\003\111\002\ -\111\002\111\002\111\002\111\002\229\001\158\001\156\001\157\001\ -\158\001\159\001\112\000\014\000\016\000\094\001\095\001\229\001\ -\110\001\215\000\056\001\215\000\127\002\082\001\104\004\145\000\ -\016\003\238\003\133\002\134\002\097\002\019\003\225\001\149\003\ -\174\002\241\000\158\001\158\001\158\001\158\001\229\003\113\001\ -\112\000\083\001\137\003\241\000\180\000\181\000\182\000\243\000\ -\187\000\188\000\189\000\241\000\241\000\194\000\241\000\088\001\ -\089\001\197\000\243\000\112\000\210\003\158\001\112\001\090\001\ -\156\001\157\001\158\001\159\001\075\003\083\001\241\000\065\001\ -\158\001\191\001\192\001\208\001\115\002\116\002\117\002\118\002\ -\085\001\086\001\087\001\088\001\089\001\241\000\112\000\118\000\ -\118\000\083\001\194\003\090\001\221\003\200\001\009\001\125\001\ -\136\001\226\003\045\001\046\001\139\003\041\003\042\003\088\001\ -\089\001\152\002\140\001\083\001\018\003\241\000\091\001\094\003\ -\153\002\112\000\145\001\184\004\153\001\003\003\160\001\215\001\ -\087\001\088\001\089\001\162\001\107\003\161\001\154\002\163\001\ -\118\000\090\001\118\000\093\001\197\000\169\001\217\001\255\000\ -\248\003\184\002\141\000\197\000\152\002\227\000\073\003\036\004\ -\074\003\001\001\001\001\186\003\001\001\193\001\202\004\224\001\ -\191\002\197\001\233\001\241\000\241\000\241\000\145\000\118\000\ -\118\000\154\002\219\001\247\001\001\001\241\001\241\000\241\000\ -\144\002\215\000\215\000\206\004\241\000\085\003\086\003\182\001\ -\215\000\118\000\248\001\215\000\215\000\148\002\149\002\150\002\ -\194\001\196\001\118\000\249\001\118\000\250\001\151\002\205\001\ -\252\001\118\000\253\001\182\000\216\002\241\000\031\004\254\001\ -\146\000\255\001\144\002\178\003\001\002\005\002\007\002\118\000\ -\155\002\156\002\118\000\082\001\010\002\094\001\095\001\118\000\ -\149\002\150\002\216\002\215\000\215\000\127\003\014\002\112\000\ -\151\002\068\003\038\004\017\002\022\002\133\003\198\001\156\001\ -\157\001\158\001\159\001\112\000\040\002\023\004\024\004\025\004\ -\026\004\027\004\021\002\155\002\156\002\091\002\146\000\098\002\ -\051\003\052\003\053\003\054\003\055\003\056\003\057\003\058\003\ -\059\003\060\003\100\002\083\001\104\002\146\000\105\002\121\002\ -\118\000\118\000\118\000\141\003\134\003\008\004\134\003\086\001\ -\087\001\088\001\089\001\205\002\112\000\112\000\182\003\183\003\ -\184\003\090\001\122\002\123\002\112\000\018\004\139\002\143\002\ -\129\003\130\003\159\003\140\002\197\000\141\002\197\000\160\002\ -\162\002\064\001\112\000\112\000\112\000\112\000\166\002\046\001\ -\144\002\118\000\167\002\082\001\173\003\168\002\118\000\170\002\ -\234\000\234\000\197\000\234\000\169\002\024\002\149\002\150\002\ -\027\002\118\000\235\000\235\000\202\003\235\000\171\002\157\001\ -\172\002\173\002\118\000\234\000\197\000\175\002\176\002\144\002\ -\157\001\146\000\193\002\177\002\178\003\235\000\178\002\215\000\ -\183\002\194\002\215\000\147\002\148\002\149\002\150\002\195\002\ -\228\002\157\001\003\004\083\001\199\002\151\002\118\000\203\002\ -\118\000\210\002\118\000\212\002\118\000\084\001\085\001\086\001\ -\087\001\088\001\089\001\200\001\234\003\221\002\235\003\235\002\ -\236\003\090\001\237\003\230\002\159\001\066\000\157\001\112\000\ -\112\000\112\000\112\000\238\002\098\003\159\001\227\000\227\000\ -\227\000\227\000\227\000\236\002\242\002\112\003\146\000\247\002\ -\245\002\118\000\146\000\036\002\121\003\249\002\159\001\156\001\ -\157\001\158\001\159\001\157\001\157\001\157\001\248\002\007\004\ -\250\002\215\000\146\000\118\000\118\000\128\003\164\002\254\000\ -\071\001\072\001\073\001\074\001\251\002\244\003\062\001\253\002\ -\000\003\019\004\020\004\159\001\255\002\001\003\157\001\118\000\ -\128\002\129\002\130\002\143\003\112\000\095\001\002\003\135\002\ -\005\003\157\001\013\003\118\000\089\000\197\000\118\000\231\000\ -\231\000\065\004\231\000\112\000\197\000\090\004\017\003\197\000\ -\159\001\159\001\159\001\159\001\039\004\233\000\233\000\020\003\ -\233\000\023\003\231\000\025\003\159\001\029\003\213\000\026\003\ -\213\000\234\000\234\000\118\000\035\003\030\003\254\000\037\003\ -\233\000\031\003\032\003\159\001\156\001\157\001\158\001\159\001\ -\033\003\036\003\118\000\118\000\234\000\012\001\159\001\041\001\ -\038\003\118\000\040\003\068\003\068\003\039\003\046\003\132\004\ -\068\004\069\004\046\001\140\004\196\003\232\000\232\000\118\000\ -\232\000\045\003\047\003\081\003\089\003\225\001\095\003\055\004\ -\056\004\057\004\058\004\059\004\209\003\097\003\100\003\102\003\ -\232\000\164\004\116\003\081\001\118\003\217\003\146\003\063\001\ -\118\000\123\003\152\003\144\004\153\003\156\003\158\003\197\000\ -\197\000\197\000\197\000\197\000\082\001\160\003\109\004\163\003\ -\181\004\180\003\255\000\181\003\185\003\040\001\146\000\156\002\ -\102\000\188\003\189\003\190\003\215\000\191\003\192\003\197\003\ -\193\003\215\000\102\000\199\003\198\003\218\003\118\000\203\003\ -\118\000\118\000\102\000\118\000\118\000\118\000\118\000\118\000\ -\204\003\205\003\206\003\207\003\150\004\219\003\152\004\153\004\ -\211\003\212\003\213\003\214\003\083\001\102\000\215\003\197\000\ -\223\003\117\004\227\003\121\004\228\003\230\003\084\001\085\001\ -\086\001\087\001\088\001\089\001\102\000\231\003\232\003\233\003\ -\135\004\255\000\090\001\250\003\118\000\253\003\213\000\213\000\ -\254\003\197\000\255\003\197\000\179\001\213\000\000\004\001\004\ -\213\000\213\000\187\004\002\004\102\000\006\004\195\001\012\004\ -\197\000\004\004\029\004\017\004\179\001\234\000\234\000\234\000\ -\009\004\068\000\010\004\021\004\022\004\234\000\037\004\028\004\ -\030\004\032\004\242\003\222\001\234\000\234\000\034\004\053\004\ -\068\000\051\004\052\004\234\000\054\004\060\004\063\004\068\000\ -\213\000\213\000\066\004\068\000\118\000\067\004\094\004\077\004\ -\078\004\095\004\102\000\102\000\102\000\068\000\061\004\112\000\ -\068\000\091\004\209\004\105\004\197\000\102\000\102\000\107\004\ -\112\000\068\000\112\000\102\000\108\004\234\000\110\004\112\000\ -\111\004\115\004\112\004\113\004\075\004\068\000\123\004\114\004\ -\127\004\143\004\145\004\205\004\146\004\018\002\068\000\147\004\ -\112\000\068\000\151\004\165\004\102\000\166\004\068\000\197\000\ -\167\004\168\004\169\004\170\004\185\004\171\004\172\004\173\004\ -\174\004\182\004\175\004\197\000\068\000\068\000\068\000\068\000\ -\068\000\068\000\068\000\068\000\197\000\176\004\188\004\194\004\ -\027\002\186\004\189\004\190\004\191\004\203\004\192\004\183\001\ -\183\001\183\001\193\004\197\004\183\001\207\004\198\004\199\004\ -\076\000\200\004\068\000\201\004\208\004\236\001\058\000\068\000\ -\183\001\039\000\076\000\053\002\053\000\053\002\098\000\068\000\ -\043\000\066\002\076\000\076\000\045\000\076\000\154\004\155\004\ -\156\004\157\004\158\004\068\000\213\000\041\002\024\002\213\000\ -\053\000\041\000\070\000\062\002\214\000\076\000\214\000\235\000\ -\235\000\031\000\074\000\116\000\222\000\175\003\059\002\112\000\ -\060\000\053\000\100\000\215\000\076\000\053\000\222\000\143\002\ -\183\001\114\000\235\000\037\001\050\002\072\000\222\000\112\000\ -\060\002\222\000\183\001\183\001\183\001\183\001\183\001\183\001\ -\112\000\106\000\185\001\007\000\076\000\062\000\183\001\108\000\ -\051\002\222\000\185\001\185\001\023\000\058\000\035\000\064\002\ -\136\003\056\002\112\000\110\003\137\002\183\001\122\004\057\002\ -\222\000\045\002\143\001\144\001\185\001\119\002\213\000\144\002\ -\222\001\067\002\068\002\069\002\070\002\071\002\111\003\066\003\ -\165\003\145\002\146\002\147\002\148\002\149\002\150\002\151\000\ -\222\000\024\002\076\000\076\000\076\000\151\002\234\000\079\003\ -\010\003\185\001\220\002\005\004\179\000\076\000\076\000\049\004\ -\234\000\234\000\234\000\076\000\234\000\082\003\161\002\189\002\ -\143\002\234\000\155\001\028\003\185\001\239\003\161\003\115\001\ -\186\002\124\003\007\001\234\000\014\004\013\004\185\001\185\001\ -\185\001\185\001\185\001\185\001\076\000\183\004\222\000\222\000\ -\222\000\241\003\185\001\090\002\234\000\234\000\133\001\224\002\ -\203\001\222\000\222\000\234\002\165\002\112\000\175\003\222\000\ -\084\003\185\001\145\000\056\002\214\000\214\000\234\000\038\002\ -\144\002\112\000\180\001\214\000\185\001\240\003\214\000\214\000\ -\005\000\082\002\145\000\146\002\147\002\148\002\149\002\150\002\ -\222\000\015\004\180\001\235\000\235\000\235\000\151\002\167\001\ -\196\004\208\003\142\002\235\000\206\002\006\000\000\000\000\000\ -\167\001\112\000\235\000\235\000\215\000\215\000\000\000\000\000\ -\000\000\235\000\000\000\143\002\000\000\000\000\214\000\214\000\ -\000\000\167\001\000\000\000\000\000\000\000\000\000\000\112\000\ -\000\000\000\000\000\000\007\000\008\000\000\000\000\000\000\000\ -\000\000\213\000\000\000\000\000\009\000\112\000\213\000\000\000\ -\000\000\000\000\000\000\235\000\010\000\011\000\167\001\000\000\ -\000\000\000\000\000\000\000\000\012\000\236\001\138\003\140\003\ -\142\003\144\003\215\000\144\002\000\000\013\000\000\000\014\000\ -\015\000\167\001\016\000\000\000\000\000\145\002\146\002\147\002\ -\148\002\149\002\150\002\167\001\167\001\167\001\167\001\167\001\ -\167\001\151\002\000\000\000\000\179\001\179\001\000\000\167\001\ -\000\000\017\000\000\000\138\003\140\003\142\003\144\003\215\000\ -\000\000\000\000\000\000\000\000\018\000\019\000\167\001\000\000\ -\179\001\112\000\112\000\112\000\112\000\112\000\234\000\000\000\ -\000\000\167\001\000\000\179\001\000\000\234\000\234\000\000\000\ -\020\000\000\000\000\000\000\000\234\000\179\001\046\001\000\000\ -\000\000\000\000\152\000\000\000\021\000\022\000\000\000\023\000\ -\000\000\000\000\214\000\000\000\000\000\214\000\000\000\000\000\ -\000\000\153\000\211\000\000\000\211\000\230\000\230\000\000\000\ -\154\000\000\000\000\000\000\000\155\000\000\000\000\000\000\000\ -\000\000\234\000\234\000\000\000\000\000\000\000\156\000\000\000\ -\230\000\157\000\000\000\119\003\000\000\000\000\000\000\000\000\ -\000\000\000\000\158\000\000\000\000\000\000\000\000\000\215\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\159\000\000\000\ -\000\000\234\000\000\000\000\000\000\000\000\000\000\000\160\000\ -\000\000\000\000\161\000\000\000\000\000\184\001\000\000\162\000\ -\000\000\000\000\000\000\000\000\214\000\224\002\184\001\000\000\ -\000\000\000\000\000\000\000\000\000\000\163\000\164\000\165\000\ -\166\000\167\000\168\000\169\000\170\000\000\000\000\000\184\001\ -\000\000\000\000\000\000\000\000\235\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\235\000\235\000\ -\235\000\000\000\235\000\171\000\000\000\000\000\000\000\235\000\ -\172\000\000\000\000\000\000\000\184\001\000\000\000\000\000\000\ -\173\000\235\000\000\000\000\000\236\001\000\000\236\001\000\000\ -\000\000\236\001\000\000\000\000\174\000\000\000\000\000\184\001\ -\000\000\179\001\235\000\235\000\000\000\227\002\000\000\000\000\ -\213\000\184\001\184\001\184\001\184\001\184\001\184\001\000\000\ -\000\000\000\000\211\000\211\000\235\000\184\001\000\000\000\000\ -\000\000\145\003\000\000\000\000\211\000\211\000\000\000\178\000\ -\000\000\000\000\220\003\236\001\184\001\000\000\224\002\000\000\ -\000\000\230\000\230\000\230\000\000\000\000\000\000\000\184\001\ -\000\000\230\000\000\000\000\000\006\000\000\000\000\000\000\000\ -\230\000\230\000\000\000\000\000\000\000\000\000\000\000\230\000\ -\000\000\000\000\000\000\000\000\211\000\211\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\247\003\000\000\000\000\000\000\ -\224\002\000\000\007\000\008\000\000\000\000\000\000\000\214\000\ -\000\000\000\000\000\000\009\000\214\000\000\000\000\000\000\000\ -\000\000\230\000\000\000\010\000\011\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\012\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\013\000\000\000\014\000\015\000\ -\000\000\016\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\234\000\234\000\234\000\234\000\234\000\ -\222\003\000\000\180\001\180\001\236\001\236\001\000\000\224\002\ -\017\000\000\000\000\000\224\002\000\000\000\000\000\000\000\000\ -\234\000\000\000\000\000\018\000\019\000\000\000\180\001\000\000\ -\000\000\000\000\041\004\045\004\235\000\000\000\000\000\000\000\ -\000\000\180\001\000\000\235\000\235\000\000\000\000\000\020\000\ -\000\000\000\000\235\000\180\001\236\001\000\000\236\001\000\000\ -\000\000\213\000\213\000\021\000\022\000\000\000\023\000\000\000\ -\211\000\000\000\000\000\211\000\188\001\000\000\000\000\000\000\ -\000\000\000\000\000\000\188\001\000\000\000\000\224\002\000\000\ -\000\000\188\001\224\002\000\000\000\000\000\000\000\000\235\000\ -\235\000\188\001\188\001\000\000\000\000\114\003\000\000\000\000\ -\098\004\000\000\045\004\000\000\000\000\154\001\000\000\000\000\ -\000\000\000\000\000\000\188\001\000\000\000\000\154\001\213\000\ -\000\000\000\000\222\003\000\000\000\000\236\001\000\000\235\000\ -\000\000\236\001\000\000\000\000\000\000\000\000\224\002\154\001\ -\000\000\045\004\000\000\225\000\225\000\000\000\000\000\234\000\ -\188\001\000\000\000\000\227\002\000\000\000\000\045\004\000\000\ -\000\000\000\000\000\000\000\000\213\000\000\000\225\000\000\000\ -\000\000\000\000\000\000\188\001\154\001\000\000\000\000\000\000\ -\000\000\000\000\000\000\224\002\000\000\188\001\188\001\188\001\ -\188\001\188\001\188\001\188\001\230\000\230\000\230\000\000\000\ -\230\000\188\001\000\000\000\000\236\001\230\000\000\000\000\000\ -\236\001\154\001\154\001\154\001\154\001\000\000\000\000\230\000\ -\188\001\000\000\000\000\000\000\000\000\154\001\000\000\000\000\ -\000\000\000\000\000\000\188\001\000\000\000\000\000\000\180\001\ -\230\000\230\000\000\000\000\000\154\001\000\000\214\000\000\000\ -\000\000\000\000\000\000\045\004\000\000\000\000\000\000\154\001\ -\000\000\000\000\230\000\000\000\236\001\000\000\000\000\000\000\ -\000\000\000\000\114\001\000\000\213\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\227\002\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\114\001\ -\151\001\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\236\001\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\227\002\000\000\ -\000\000\012\003\000\000\000\000\000\000\211\000\000\000\000\000\ -\000\000\000\000\211\000\000\000\000\000\000\000\000\000\225\000\ -\225\000\225\000\000\000\000\000\000\000\000\000\000\000\225\000\ -\000\000\000\000\000\000\000\000\220\001\221\001\225\000\225\000\ -\000\000\000\000\000\000\000\000\000\000\225\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\235\000\235\000\235\000\235\000\235\000\000\000\000\000\ -\000\000\000\000\068\000\000\000\000\000\227\002\000\000\140\000\ -\000\000\227\002\000\000\000\000\000\000\000\000\235\000\225\000\ -\000\000\068\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\068\000\048\004\230\000\000\000\068\000\000\000\000\000\000\000\ -\000\000\230\000\230\000\000\000\101\001\000\000\068\000\000\000\ -\000\000\068\000\000\000\186\001\000\000\000\000\000\000\214\000\ -\214\000\000\000\068\000\231\000\231\000\034\002\000\000\000\000\ -\186\001\000\000\000\000\000\000\000\000\000\000\068\000\000\000\ -\186\001\186\001\000\000\000\000\227\002\000\000\231\000\068\000\ -\227\002\000\000\068\000\000\000\000\000\230\000\230\000\068\000\ -\000\000\000\000\186\001\000\000\000\000\000\000\000\000\000\000\ -\048\004\000\000\000\000\000\000\078\002\068\000\068\000\068\000\ -\068\000\068\000\068\000\068\000\068\000\214\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\230\000\000\000\186\001\ -\000\000\188\001\000\000\000\000\227\002\000\000\000\000\048\004\ -\000\000\000\000\000\000\068\000\000\000\235\000\000\000\000\000\ -\068\000\188\001\186\001\000\000\048\004\000\000\000\000\000\000\ -\068\000\000\000\214\000\114\001\186\001\186\001\186\001\186\001\ -\186\001\186\001\186\001\000\000\068\000\000\000\000\000\000\000\ -\186\001\227\002\117\001\119\001\121\001\231\000\231\000\000\000\ -\000\000\000\000\129\001\131\001\231\000\000\000\231\000\186\001\ -\000\000\000\000\000\000\000\000\000\000\000\000\231\000\117\001\ -\000\000\000\000\186\001\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\232\000\243\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\211\000\000\000\000\000\000\000\ -\225\000\048\004\000\000\243\000\000\000\000\000\000\000\000\000\ -\185\002\000\000\225\000\225\000\225\000\201\001\225\000\231\000\ -\231\000\231\000\214\000\225\000\000\000\000\000\188\001\231\000\ -\000\000\000\000\000\000\000\000\000\000\225\000\231\000\231\000\ -\000\000\000\000\000\000\000\000\000\000\231\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\218\002\225\000\ -\000\000\000\000\000\000\233\000\233\000\073\002\000\000\000\000\ -\000\000\101\001\000\000\000\000\000\000\012\003\000\000\000\000\ -\225\000\000\000\000\000\000\000\000\000\000\000\233\000\231\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\246\002\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\122\001\123\001\000\000\000\000\000\000\130\001\ -\000\000\135\001\000\000\138\001\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\146\001\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\006\003\230\000\ -\230\000\230\000\230\000\230\000\000\000\000\000\000\000\188\001\ -\000\000\000\000\000\000\000\000\188\001\188\001\188\001\000\000\ -\188\001\188\001\000\000\000\000\230\000\000\000\000\000\000\000\ -\000\000\101\001\000\000\114\001\000\000\000\000\000\000\000\000\ -\000\000\201\001\000\000\000\000\243\000\243\000\243\000\000\000\ -\000\000\000\000\000\000\000\000\232\000\000\000\000\000\000\000\ -\089\002\000\000\000\000\243\000\243\000\211\000\211\000\094\002\ -\152\001\000\000\232\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\231\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\117\001\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\225\000\000\000\000\000\000\000\243\000\000\000\000\000\225\000\ -\225\000\000\000\000\000\000\000\000\000\000\000\225\000\233\000\ -\233\000\233\000\000\000\211\000\000\000\000\000\020\002\233\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\233\000\233\000\ -\000\000\000\000\000\000\000\000\000\000\233\000\000\000\000\000\ -\000\000\000\000\000\000\230\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\225\000\225\000\000\000\000\000\000\000\ -\211\000\000\000\000\000\000\000\000\000\231\000\000\000\000\000\ -\000\000\000\000\231\000\231\000\231\000\000\000\231\000\233\000\ -\000\000\000\000\000\000\231\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\225\000\000\000\231\000\000\000\000\000\ -\000\000\000\000\201\001\000\000\201\001\000\000\000\000\000\000\ -\000\000\000\000\078\002\000\000\000\000\000\000\231\000\231\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\114\001\000\000\000\000\099\002\ -\231\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\188\001\000\000\ -\211\000\188\001\188\001\188\001\188\001\188\001\188\001\188\001\ -\188\001\188\001\188\001\188\001\188\001\188\001\188\001\188\001\ -\188\001\188\001\073\002\101\001\000\000\188\001\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\117\001\000\000\ -\188\001\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\188\001\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\231\000\000\000\000\000\231\000\000\000\ -\000\000\000\000\231\000\117\001\000\000\000\000\000\000\243\000\ -\243\000\243\000\000\000\232\000\000\000\000\000\000\000\000\000\ -\243\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\243\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\078\002\ -\006\003\000\000\000\000\243\000\243\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\243\000\000\000\000\000\ -\231\000\201\001\101\001\201\001\000\000\000\000\000\000\231\000\ -\231\000\000\000\233\000\233\000\233\000\000\000\233\000\000\000\ -\000\000\000\000\000\000\233\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\233\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\225\000\225\000\225\000\ -\225\000\225\000\000\000\000\000\000\000\000\000\233\000\233\000\ -\000\000\000\000\000\000\231\000\231\000\000\000\000\000\000\000\ -\000\000\000\000\225\000\000\000\000\000\000\000\000\000\000\000\ -\233\000\000\000\000\000\000\000\000\000\043\004\000\000\000\000\ -\021\003\000\000\000\000\024\003\000\000\000\000\188\001\027\003\ -\000\000\041\000\131\001\231\000\000\000\000\000\000\000\000\000\ +let yytablesize = 21824 +let yytable = "\122\000\ +\045\001\002\001\115\001\238\001\231\000\016\001\107\002\150\000\ +\122\000\227\001\083\002\057\002\218\001\064\001\084\002\007\002\ +\216\000\035\002\220\000\030\000\030\000\212\002\151\000\208\001\ +\213\001\194\001\204\000\222\001\162\002\078\002\033\002\150\000\ +\150\000\150\000\058\002\150\000\150\000\150\000\205\000\122\000\ +\150\000\192\002\254\000\036\002\007\001\030\000\038\003\030\000\ +\008\001\009\001\010\001\136\002\080\001\068\004\255\000\140\002\ +\142\002\082\001\080\001\196\000\158\003\046\001\080\001\122\000\ +\122\000\122\000\185\002\061\001\056\002\122\000\065\003\066\003\ +\067\003\068\003\069\003\044\001\122\000\151\000\151\000\151\000\ +\154\003\047\000\003\001\190\000\055\002\150\000\150\000\128\001\ +\122\000\131\001\151\000\235\002\066\001\068\001\200\000\082\001\ +\110\000\206\000\154\000\073\001\132\001\055\000\207\000\191\003\ +\110\001\220\003\135\002\110\004\080\001\078\000\009\000\242\000\ +\242\000\155\003\242\000\122\000\122\000\145\000\201\000\199\003\ +\220\001\201\000\042\001\047\000\233\001\130\002\047\001\079\002\ +\068\000\097\002\072\001\158\001\080\001\048\001\220\001\143\001\ +\080\001\145\000\207\000\115\002\158\001\234\001\122\000\055\000\ +\207\000\150\000\168\003\079\002\251\000\250\001\207\000\080\002\ +\080\000\062\001\063\001\233\001\052\001\201\000\212\001\204\001\ +\219\001\080\001\009\000\055\001\208\000\229\003\235\001\111\001\ +\169\001\170\001\068\000\080\002\234\001\026\003\219\001\080\001\ +\068\000\078\000\209\000\251\000\056\002\186\002\150\000\058\001\ +\231\000\080\000\158\001\096\001\137\004\055\001\143\001\059\001\ +\048\003\012\003\110\000\122\000\055\002\202\002\232\001\157\002\ +\091\003\222\001\092\003\074\001\047\003\081\001\036\002\253\003\ +\060\001\168\001\149\000\045\001\080\001\210\000\016\001\158\001\ +\158\001\150\000\241\001\242\001\172\003\173\003\174\003\175\003\ +\176\003\243\002\242\000\245\002\081\002\039\002\020\004\215\002\ +\208\004\041\002\184\000\185\000\186\000\228\003\191\000\192\000\ +\193\000\034\002\158\001\198\000\237\002\084\003\145\001\093\000\ +\081\002\210\000\089\002\105\003\122\000\158\001\162\002\210\000\ +\222\001\143\001\041\002\184\002\008\003\210\000\196\002\090\000\ +\122\000\137\004\194\001\203\001\145\000\052\003\037\004\171\001\ +\244\002\129\003\039\004\023\002\074\004\075\004\076\004\077\004\ +\078\004\069\001\090\000\191\002\099\001\157\001\044\001\082\001\ +\049\001\050\001\150\000\241\002\012\001\002\001\157\001\044\001\ +\160\002\172\001\084\004\085\002\145\000\040\002\208\002\114\002\ +\150\000\122\000\122\000\150\000\078\000\244\002\131\002\047\000\ +\222\001\122\000\062\001\063\001\137\002\138\002\095\001\145\000\ +\126\003\096\000\122\000\222\001\013\001\096\001\056\002\122\000\ +\122\000\122\000\122\000\055\000\207\000\080\004\096\004\083\001\ +\043\002\083\004\096\000\097\001\157\001\122\001\092\002\070\001\ +\095\001\092\002\085\004\145\000\149\000\155\001\055\001\096\001\ +\208\000\088\002\156\002\115\002\129\001\084\001\068\000\078\000\ +\015\004\157\002\078\000\096\000\156\002\122\000\209\000\090\000\ +\122\000\157\001\002\001\157\002\086\004\068\000\003\001\158\002\ +\095\001\078\000\212\002\087\002\109\001\120\004\151\000\007\003\ +\035\002\186\000\061\002\062\002\063\002\064\002\065\002\066\002\ +\067\002\068\002\069\002\070\002\157\001\097\001\107\001\107\002\ +\100\003\050\002\051\002\047\000\155\001\152\004\129\002\157\001\ +\145\000\103\004\113\001\059\002\122\000\122\000\122\000\122\000\ +\153\004\128\002\163\004\122\000\202\001\098\001\099\001\055\000\ +\145\000\045\001\217\002\150\000\150\000\150\000\130\003\151\003\ +\153\003\124\002\150\000\093\000\016\001\064\001\196\000\232\001\ +\124\004\170\003\145\004\210\000\209\002\156\002\209\002\098\001\ +\099\001\038\001\068\000\003\001\190\003\223\002\016\001\053\003\ +\054\003\159\002\160\002\038\001\155\001\117\002\151\001\013\000\ +\015\000\151\001\158\002\159\002\160\002\194\002\171\003\155\001\ +\114\001\122\000\103\003\071\003\231\001\029\000\201\002\098\001\ +\099\001\184\002\145\000\038\001\023\000\201\000\082\003\231\001\ +\122\000\200\002\117\003\151\001\119\003\050\001\255\003\124\003\ +\087\003\000\004\011\000\012\000\044\001\150\000\040\002\085\001\ +\255\001\122\000\047\000\028\002\068\000\116\001\031\002\245\000\ +\236\002\011\000\053\001\054\001\100\004\116\000\080\001\211\002\ +\086\001\023\000\245\000\012\003\002\001\243\001\055\000\145\000\ +\112\003\115\002\115\002\115\002\115\002\115\002\136\003\040\002\ +\023\000\220\002\055\003\056\003\057\003\058\003\059\003\060\003\ +\061\003\062\003\063\003\064\003\159\002\160\002\011\000\244\001\ +\060\001\068\000\013\000\015\000\208\000\116\000\150\001\069\001\ +\038\001\122\004\101\002\219\000\219\000\011\000\219\000\011\002\ +\087\001\233\003\209\000\020\003\242\003\153\003\117\001\243\000\ +\023\003\013\001\088\001\089\001\090\001\091\001\092\001\093\001\ +\150\001\243\000\113\003\116\000\029\000\219\000\094\001\149\001\ +\129\001\243\000\243\000\229\001\243\000\230\001\014\000\016\000\ +\160\001\161\001\162\001\163\001\201\000\003\001\116\000\195\001\ +\196\001\138\000\086\001\140\001\243\000\098\003\087\001\079\003\ +\229\001\149\001\178\002\214\003\149\001\204\001\212\001\139\000\ +\140\000\087\001\111\003\243\000\092\001\093\001\132\002\133\002\ +\134\002\116\000\122\000\122\000\141\003\139\002\225\003\092\001\ +\093\001\108\004\148\002\230\003\144\001\157\001\164\001\094\001\ +\045\003\046\003\165\001\243\000\198\003\068\000\166\001\022\003\ +\153\002\154\002\087\001\167\001\116\000\197\001\003\001\003\001\ +\155\002\003\001\219\001\078\000\088\001\089\001\090\001\091\001\ +\092\001\093\001\040\004\122\000\201\001\122\000\173\001\201\000\ +\094\001\003\001\252\003\221\001\223\001\231\000\201\000\148\002\ +\145\000\077\003\228\001\078\003\237\001\236\000\236\000\072\003\ +\236\000\243\000\243\000\243\000\245\001\153\002\154\002\252\001\ +\050\001\182\003\122\000\122\000\243\000\243\000\160\001\251\001\ +\236\000\253\001\243\000\087\001\219\000\219\000\254\001\160\001\ +\089\003\090\003\186\001\219\000\122\000\000\002\219\000\219\000\ +\091\001\092\001\093\001\198\001\200\001\122\000\188\004\122\000\ +\160\001\094\001\209\001\243\000\122\000\001\002\237\000\237\000\ +\035\004\237\000\002\002\150\000\003\002\220\002\160\001\161\001\ +\162\001\163\001\122\000\005\002\009\002\122\000\011\002\014\002\ +\021\002\237\000\122\000\018\002\025\002\160\001\219\000\219\000\ +\131\003\206\004\116\000\220\002\026\002\042\004\154\001\044\002\ +\137\003\027\004\028\004\029\004\030\004\031\004\116\000\154\001\ +\095\002\102\002\143\003\104\002\233\000\233\000\210\004\233\000\ +\108\002\150\000\160\001\160\001\160\001\160\001\109\002\143\002\ +\154\001\235\000\235\000\209\002\235\000\126\002\012\004\233\000\ +\150\000\125\002\144\002\122\000\122\000\122\000\127\002\138\003\ +\145\002\138\003\164\002\166\002\235\000\160\001\022\004\116\000\ +\116\000\186\003\187\003\188\003\170\002\154\001\171\002\116\000\ +\160\001\172\002\173\002\133\003\134\003\163\003\174\002\201\000\ +\175\002\201\000\182\003\176\002\179\002\116\000\116\000\116\000\ +\116\000\177\002\234\000\234\000\122\000\234\000\147\002\177\003\ +\180\002\122\000\154\001\154\001\154\001\201\000\181\002\182\002\ +\068\001\187\002\197\002\198\002\122\000\234\000\199\002\206\003\ +\187\001\001\000\002\000\003\000\004\000\122\000\203\002\201\000\ +\187\001\187\001\007\004\207\002\150\000\154\001\160\001\161\001\ +\162\001\163\001\219\000\214\002\232\002\219\000\216\002\234\002\ +\154\001\225\002\187\001\204\001\239\002\240\002\148\002\246\002\ +\242\002\122\000\251\002\122\000\253\002\122\000\249\002\122\000\ +\149\002\150\002\151\002\152\002\153\002\154\002\252\002\238\003\ +\255\002\239\003\145\003\240\003\155\002\241\003\254\002\187\001\ +\066\000\001\003\116\000\116\000\116\000\116\000\231\000\231\000\ +\231\000\231\000\231\000\160\001\161\001\162\001\163\001\003\003\ +\004\003\150\000\187\001\005\003\122\000\150\000\031\002\002\001\ +\075\001\076\001\077\001\078\001\187\001\187\001\187\001\187\001\ +\187\001\187\001\011\004\094\004\219\000\150\000\122\000\122\000\ +\187\001\168\002\040\002\006\003\099\001\086\001\009\003\147\003\ +\248\003\072\003\072\003\093\000\023\004\024\004\017\003\187\001\ +\021\003\024\003\122\000\160\001\161\001\162\001\163\001\116\000\ +\027\003\147\002\187\001\033\003\028\002\066\001\122\000\029\003\ +\201\000\122\000\030\003\041\003\069\004\146\002\116\000\201\000\ +\039\003\034\003\201\000\179\003\040\003\136\004\002\001\043\004\ +\035\003\144\004\036\003\037\003\042\003\087\001\147\002\043\003\ +\049\003\217\000\044\003\217\000\238\000\238\000\122\000\161\001\ +\003\001\090\001\091\001\092\001\093\001\045\001\051\003\168\004\ +\161\001\148\002\050\003\094\001\085\003\122\000\122\000\238\000\ +\016\001\229\001\099\003\101\003\122\000\151\002\152\002\153\002\ +\154\002\161\001\032\000\072\004\073\004\093\003\185\004\155\002\ +\104\003\120\003\122\000\106\003\122\003\156\003\148\002\127\003\ +\160\003\150\003\059\004\060\004\061\004\062\004\063\004\157\003\ +\149\002\150\002\151\002\152\002\153\002\154\002\161\001\028\002\ +\162\003\164\003\067\001\122\000\155\002\167\003\148\004\003\001\ +\184\003\189\003\201\000\201\000\201\000\201\000\201\000\086\001\ +\185\003\113\004\160\002\192\003\148\002\193\003\194\003\195\003\ +\044\001\150\000\196\003\161\001\161\001\161\001\161\001\219\000\ +\197\003\152\002\153\002\154\002\219\000\201\003\202\003\161\001\ +\203\003\122\000\155\002\122\000\122\000\207\003\122\000\122\000\ +\122\000\122\000\122\000\208\003\209\003\210\003\161\001\154\004\ +\211\003\156\004\157\004\215\003\179\003\216\003\147\002\087\001\ +\149\000\161\001\201\000\217\003\121\004\218\003\125\004\222\003\ +\223\003\219\003\089\001\090\001\091\001\092\001\093\001\227\003\ +\149\000\231\003\232\003\139\004\234\003\094\001\236\003\122\000\ +\235\003\217\000\217\000\237\003\201\000\001\004\201\000\183\001\ +\217\000\254\003\002\004\217\000\217\000\191\004\003\004\004\004\ +\005\004\199\001\006\004\201\000\008\004\013\004\148\002\183\001\ +\238\000\238\000\238\000\010\004\068\000\014\004\025\004\016\004\ +\238\000\150\002\151\002\152\002\153\002\154\002\226\001\238\000\ +\238\000\026\004\021\004\068\000\155\002\032\004\238\000\156\001\ +\034\004\036\004\068\000\217\000\217\000\038\004\068\000\122\000\ +\156\001\246\003\055\004\056\004\057\004\058\004\064\004\070\004\ +\068\000\067\004\116\000\068\000\071\004\213\004\082\004\201\000\ +\081\004\156\001\095\004\116\000\068\000\116\000\099\004\098\004\ +\238\000\111\004\116\000\109\004\119\004\114\004\112\004\115\004\ +\068\000\116\004\117\004\150\004\118\004\147\004\209\004\149\004\ +\022\002\068\000\127\004\116\000\068\000\131\004\156\001\151\004\ +\155\004\068\000\201\000\169\004\142\003\144\003\146\003\148\003\ +\170\004\171\004\172\004\173\004\174\004\186\004\201\000\068\000\ +\068\000\068\000\068\000\068\000\068\000\068\000\068\000\201\000\ +\175\004\189\004\192\004\156\001\156\001\156\001\156\001\176\004\ +\198\004\177\004\178\004\179\004\050\001\190\004\180\004\156\001\ +\006\000\142\003\144\003\146\003\148\003\068\000\193\004\194\004\ +\240\001\195\004\068\000\196\004\197\004\207\004\156\001\201\004\ +\202\004\203\004\068\000\204\004\205\004\007\000\211\004\212\004\ +\039\000\156\001\058\000\053\000\031\000\055\002\068\000\217\000\ +\055\002\045\000\217\000\098\000\043\002\070\000\053\000\218\000\ +\068\002\218\000\239\000\239\000\043\000\041\000\076\000\074\000\ +\116\000\064\002\116\000\008\000\009\000\061\002\219\000\060\000\ +\076\000\053\000\053\000\100\000\010\000\239\000\041\001\114\000\ +\076\000\076\000\116\000\076\000\011\000\012\000\052\002\072\000\ +\062\002\062\000\007\000\116\000\013\000\169\001\108\000\106\000\ +\053\002\023\000\058\000\076\000\102\000\014\000\169\001\015\000\ +\016\000\035\000\017\000\066\002\058\002\112\000\102\000\141\002\ +\059\002\140\003\076\000\114\003\049\002\126\004\102\000\169\001\ +\147\001\217\000\123\002\148\001\071\002\072\002\073\002\074\002\ +\075\002\018\000\169\003\083\003\014\003\224\002\053\004\115\003\ +\070\003\102\000\076\000\009\004\019\000\020\000\155\000\034\000\ +\165\002\238\000\183\000\159\001\169\001\165\003\086\003\243\003\ +\102\000\032\003\119\001\238\000\238\000\238\000\128\003\238\000\ +\021\000\011\001\018\004\193\002\238\000\017\004\245\003\169\001\ +\094\002\190\002\187\004\137\001\022\000\023\000\238\000\024\000\ +\102\000\169\001\169\001\169\001\169\001\169\001\169\001\207\001\ +\076\000\076\000\076\000\060\002\042\002\169\001\244\003\238\000\ +\238\000\169\002\228\002\076\000\076\000\088\003\086\002\210\002\ +\116\000\076\000\019\004\212\003\169\001\200\004\000\000\218\000\ +\218\000\238\000\000\000\000\000\116\000\184\001\218\000\169\001\ +\000\000\218\000\218\000\000\000\000\000\000\000\102\000\102\000\ +\102\000\000\000\076\000\238\002\000\000\184\001\239\000\239\000\ +\239\000\102\000\102\000\000\000\000\000\000\000\239\000\102\000\ +\000\000\000\000\000\000\000\000\116\000\239\000\239\000\219\000\ +\219\000\000\000\000\000\159\001\239\000\000\000\000\000\000\000\ +\000\000\218\000\218\000\000\000\159\001\000\000\000\000\000\000\ +\102\000\000\000\116\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\217\000\159\001\000\000\224\000\ +\116\000\217\000\000\000\000\000\000\000\000\000\239\000\000\000\ +\000\000\224\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\240\001\224\000\000\000\000\000\224\000\219\000\000\000\000\000\ +\000\000\000\000\159\001\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\224\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\183\001\ +\183\001\000\000\000\000\224\000\000\000\000\000\000\000\159\001\ +\159\001\159\001\219\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\183\001\116\000\116\000\116\000\116\000\ +\116\000\238\000\000\000\224\000\000\000\000\000\183\001\000\000\ +\238\000\238\000\159\001\000\000\000\000\000\000\000\000\238\000\ +\183\001\000\000\000\000\000\000\000\000\159\001\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\218\000\000\000\000\000\ +\218\000\188\001\000\000\000\000\000\000\215\000\000\000\215\000\ +\234\000\234\000\000\000\000\000\000\000\045\000\188\001\000\000\ +\000\000\224\000\224\000\224\000\238\000\238\000\188\001\188\001\ +\000\000\000\000\000\000\234\000\224\000\224\000\123\003\000\000\ +\000\000\000\000\224\000\000\000\000\000\000\000\051\000\000\000\ +\188\001\000\000\219\000\000\000\145\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\238\000\000\000\059\000\060\000\ +\061\000\062\000\063\000\224\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\188\001\000\000\218\000\ +\228\002\000\000\000\000\000\000\000\000\000\000\154\001\074\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\188\001\000\000\000\000\000\000\000\000\000\000\000\000\239\000\ +\000\000\000\000\188\001\188\001\188\001\188\001\188\001\188\001\ +\188\001\239\000\239\000\239\000\000\000\239\000\188\001\000\000\ +\000\000\000\000\239\000\000\000\000\000\000\000\082\000\000\000\ +\000\000\000\000\083\000\000\000\239\000\188\001\084\000\240\001\ +\000\000\240\001\000\000\000\000\240\001\086\000\087\000\000\000\ +\188\001\088\000\000\000\000\000\183\001\239\000\239\000\000\000\ +\231\002\091\000\000\000\217\000\000\000\000\000\000\000\094\000\ +\095\000\096\000\097\000\000\000\000\000\215\000\215\000\239\000\ +\000\000\000\000\000\000\000\000\149\003\000\000\000\000\215\000\ +\215\000\000\000\006\000\000\000\000\000\224\003\240\001\000\000\ +\000\000\228\002\000\000\000\000\234\000\234\000\234\000\000\000\ +\000\000\000\000\000\000\000\000\234\000\000\000\000\000\007\000\ +\000\000\000\000\000\000\234\000\234\000\000\000\000\000\000\000\ +\000\000\000\000\234\000\000\000\000\000\000\000\000\000\215\000\ +\215\000\000\000\000\000\000\000\000\000\000\000\000\000\251\003\ +\000\000\000\000\000\000\228\002\000\000\008\000\009\000\000\000\ +\000\000\000\000\218\000\000\000\000\000\000\000\010\000\218\000\ +\000\000\000\000\000\000\000\000\234\000\000\000\011\000\012\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\013\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\014\000\ +\000\000\015\000\016\000\000\000\017\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\238\000\238\000\ +\238\000\238\000\238\000\226\003\000\000\184\001\184\001\240\001\ +\240\001\000\000\228\002\018\000\000\000\000\000\228\002\000\000\ +\000\000\000\000\000\000\238\000\000\000\000\000\019\000\020\000\ +\000\000\184\001\000\000\000\000\000\000\045\004\049\004\239\000\ +\000\000\000\000\000\000\000\000\184\001\000\000\239\000\239\000\ +\000\000\000\000\021\000\000\000\000\000\239\000\184\001\240\001\ +\000\000\240\001\000\000\000\000\217\000\217\000\022\000\023\000\ +\000\000\024\000\000\000\215\000\000\000\000\000\215\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\228\002\000\000\099\000\000\000\228\002\000\000\000\000\ +\000\000\000\000\239\000\239\000\000\000\000\000\000\000\000\000\ +\118\003\000\000\000\000\102\004\000\000\049\004\000\000\000\000\ +\186\001\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\186\001\217\000\000\000\000\000\226\003\000\000\000\000\ +\240\001\000\000\239\000\197\000\240\001\000\000\000\000\000\000\ +\000\000\228\002\186\001\000\000\049\004\000\000\229\000\229\000\ +\000\000\000\000\238\000\000\000\000\000\000\000\231\002\000\000\ +\000\000\049\004\000\000\000\000\000\000\000\000\000\000\217\000\ +\000\000\229\000\000\000\000\000\000\000\000\000\000\000\186\001\ +\000\000\000\000\000\000\000\000\000\000\000\000\228\002\000\000\ +\000\000\000\000\000\000\000\000\051\001\000\000\000\000\234\000\ +\234\000\234\000\186\001\234\000\000\000\000\000\000\000\240\001\ +\234\000\000\000\000\000\240\001\186\001\186\001\186\001\186\001\ +\186\001\186\001\234\000\000\000\185\001\185\001\185\001\071\001\ +\186\001\185\001\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\105\001\184\001\234\000\234\000\185\001\000\000\186\001\ +\000\000\218\000\000\000\000\000\000\000\000\000\049\004\000\000\ +\000\000\000\000\186\001\000\000\000\000\234\000\000\000\240\001\ +\000\000\000\000\000\000\000\000\000\000\118\001\000\000\217\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\231\002\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\118\001\155\001\000\000\185\001\000\000\000\000\ +\000\000\000\000\000\000\000\000\240\001\000\000\000\000\185\001\ +\185\001\185\001\185\001\185\001\185\001\000\000\000\000\000\000\ +\000\000\000\000\000\000\185\001\000\000\000\000\192\001\000\000\ +\000\000\231\002\000\000\000\000\016\003\000\000\000\000\000\000\ +\215\000\000\000\185\001\000\000\000\000\215\000\192\001\000\000\ +\000\000\000\000\229\000\229\000\229\000\224\001\000\000\000\000\ +\000\000\000\000\229\000\000\000\000\000\182\000\000\000\224\001\ +\225\001\229\000\229\000\000\000\000\000\000\000\000\000\000\000\ +\229\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\007\000\000\000\239\000\239\000\239\000\239\000\ +\239\000\000\000\000\000\000\000\004\002\000\000\000\000\000\000\ +\231\002\000\000\000\000\000\000\231\002\000\000\000\000\000\000\ +\000\000\239\000\229\000\000\000\000\000\000\000\000\000\000\000\ +\008\000\009\000\000\000\000\000\052\004\234\000\000\000\000\000\ +\000\000\010\000\000\000\000\000\234\000\234\000\000\000\000\000\ +\000\000\011\000\012\000\000\000\000\000\029\002\030\002\000\000\ +\000\000\013\000\218\000\218\000\000\000\037\002\000\000\000\000\ +\038\002\000\000\014\000\192\001\015\000\016\000\000\000\017\000\ +\000\000\000\000\000\000\052\002\053\002\054\002\055\002\231\002\ +\000\000\000\000\000\000\231\002\000\000\000\000\000\000\000\000\ +\234\000\234\000\000\000\000\000\000\000\000\000\018\000\000\000\ +\110\002\000\000\077\002\052\004\000\000\000\000\105\001\082\002\ +\000\000\019\000\020\000\000\000\000\000\000\000\000\000\000\000\ +\218\000\000\000\000\000\018\001\019\001\000\000\020\001\000\000\ +\234\000\021\001\000\000\000\000\000\000\021\000\000\000\231\002\ +\000\000\000\000\052\004\000\000\000\000\022\001\000\000\000\000\ +\239\000\022\000\023\000\000\000\024\000\000\000\000\000\052\004\ +\000\000\023\001\000\000\024\001\000\000\218\000\118\001\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\119\002\120\002\121\002\122\002\231\002\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\192\001\000\000\111\002\000\000\ +\000\000\192\001\192\001\192\001\000\000\192\001\192\001\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\105\001\026\001\ +\027\001\028\001\029\001\030\001\031\001\032\001\033\001\112\002\ +\035\001\036\001\000\000\037\001\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\215\000\ +\000\000\000\000\000\000\229\000\052\004\188\002\000\000\000\000\ +\000\000\000\000\000\000\189\002\000\000\229\000\229\000\229\000\ +\000\000\229\000\156\000\000\000\195\002\218\000\229\000\000\000\ +\000\000\000\000\113\002\000\000\000\000\000\000\000\000\000\000\ +\229\000\157\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\158\000\000\000\000\000\000\000\159\000\000\000\000\000\000\000\ +\000\000\222\002\229\000\000\000\000\000\000\000\160\000\000\000\ +\000\000\161\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\016\003\000\000\162\000\229\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\163\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\250\002\000\000\164\000\ +\000\000\000\000\165\000\000\000\000\000\000\000\000\000\166\000\ +\000\000\000\000\000\000\144\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\167\000\168\000\169\000\ +\170\000\171\000\172\000\173\000\174\000\000\000\000\000\000\000\ +\000\000\010\003\234\000\234\000\234\000\234\000\234\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\175\000\000\000\000\000\000\000\234\000\ +\176\000\000\000\235\000\235\000\000\000\000\000\118\001\000\000\ +\177\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\178\000\235\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\215\000\215\000\000\000\192\001\000\000\000\000\192\001\192\001\ +\192\001\192\001\192\001\192\001\192\001\192\001\192\001\192\001\ +\192\001\192\001\192\001\192\001\192\001\192\001\192\001\077\002\ +\105\001\000\000\192\001\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\229\000\000\000\192\001\237\000\237\000\ +\000\000\000\000\229\000\229\000\000\000\000\000\000\000\192\001\ +\000\000\229\000\000\000\000\000\000\000\000\000\215\000\000\000\ +\000\000\237\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\121\001\123\001\125\001\235\000\235\000\234\000\000\000\ +\102\003\133\001\135\001\235\000\000\000\235\000\229\000\229\000\ +\000\000\116\003\000\000\215\000\000\000\235\000\121\001\000\000\ +\125\003\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\132\003\000\000\000\000\236\000\247\000\229\000\000\000\ +\000\000\000\000\000\000\221\000\000\000\000\000\000\000\105\001\ +\000\000\000\000\000\000\045\000\000\000\082\002\000\000\247\000\ +\000\000\000\000\000\000\000\000\205\001\047\000\235\000\235\000\ +\235\000\000\000\000\000\000\000\000\000\000\000\235\000\118\001\ +\000\000\000\000\000\000\000\000\051\000\235\000\235\000\000\000\ +\000\000\222\000\145\000\156\001\235\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\059\000\060\000\061\000\062\000\ +\063\000\000\000\000\000\215\000\000\000\223\000\224\000\000\000\ +\000\000\225\000\000\000\000\000\068\000\000\000\226\000\000\000\ +\000\000\000\000\000\000\000\000\227\000\074\000\235\000\000\000\ +\000\000\000\000\078\000\192\001\000\000\000\000\000\000\000\000\ +\200\003\000\000\237\000\237\000\237\000\000\000\000\000\000\000\ +\000\000\000\000\237\000\000\000\000\000\000\000\126\001\127\001\ +\213\003\237\000\237\000\134\001\000\000\139\001\000\000\142\001\ +\237\000\221\003\000\000\000\000\082\000\000\000\000\000\150\001\ +\083\000\000\000\000\000\000\000\084\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\086\000\087\000\000\000\000\000\088\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\091\000\ +\000\000\093\000\237\000\141\001\105\001\094\000\095\000\096\000\ +\097\000\228\000\082\002\010\003\000\000\000\000\000\000\000\000\ +\205\001\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\247\000\247\000\247\000\000\000\000\000\000\000\000\000\093\002\ +\236\000\000\000\000\000\000\000\000\000\000\000\098\002\247\000\ +\247\000\000\000\000\000\000\000\192\001\000\000\236\000\000\000\ +\000\000\235\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\121\001\000\000\000\000\000\000\000\000\000\000\ \000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\229\000\229\000\229\000\229\000\229\000\000\000\033\004\000\000\ +\247\000\000\000\000\000\000\000\190\001\000\000\000\000\000\000\ +\000\000\000\000\041\004\190\001\000\000\229\000\000\000\000\000\ +\000\000\190\001\024\002\000\000\000\000\000\000\000\000\000\000\ +\047\004\190\001\190\001\000\000\000\000\000\000\000\000\000\000\ \000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\047\000\000\000\000\000\000\000\000\000\000\000\ -\141\000\231\000\000\000\000\000\117\001\000\000\000\000\000\000\ -\000\000\000\000\055\000\056\000\057\000\058\000\059\000\093\004\ -\000\000\000\000\097\004\000\000\043\004\000\000\100\004\000\000\ -\000\000\000\000\000\000\000\000\000\000\243\000\000\000\101\001\ -\000\000\000\000\150\001\070\000\243\000\243\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\119\004\000\000\000\000\000\000\000\000\ -\000\000\225\000\000\000\000\000\000\000\000\000\000\000\187\001\ -\119\004\000\000\000\000\000\000\142\004\000\000\187\001\188\001\ -\000\000\201\001\078\000\000\000\187\001\000\000\079\000\000\000\ -\243\000\232\000\080\000\000\000\187\001\187\001\000\000\000\000\ -\000\000\082\000\083\000\000\000\000\000\084\000\000\000\000\000\ -\233\000\000\000\000\000\000\000\000\000\087\000\187\001\233\000\ -\233\000\000\000\000\000\090\000\091\000\092\000\093\000\000\000\ -\243\000\000\000\000\000\000\000\100\004\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\117\001\000\000\000\000\187\001\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\119\004\162\003\000\000\ -\000\000\000\000\000\000\233\000\233\000\000\000\187\001\000\000\ -\201\001\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\187\001\187\001\187\001\187\001\187\001\187\001\187\001\000\000\ -\000\000\201\001\000\000\000\000\187\001\000\000\000\000\000\000\ -\000\000\000\000\000\000\233\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\187\001\000\000\231\000\231\000\231\000\ -\231\000\231\000\000\000\000\000\000\000\000\000\187\001\000\000\ -\000\000\027\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\231\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\040\004\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\131\001\131\001\131\001\000\000\ -\131\001\000\000\000\000\000\000\000\000\000\000\000\000\131\001\ -\000\000\131\001\000\000\131\001\000\000\000\000\131\001\000\000\ -\000\000\000\000\000\000\000\000\000\000\131\001\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\131\001\000\000\131\001\ -\131\001\000\000\000\000\131\001\131\001\000\000\000\000\000\000\ -\131\001\000\000\000\000\085\004\000\000\000\000\131\001\000\000\ -\131\001\131\001\131\001\131\001\131\001\131\001\000\000\231\000\ -\000\000\000\000\000\000\000\000\131\001\000\000\131\001\000\000\ -\000\000\000\000\000\000\000\000\131\001\131\001\000\000\000\000\ -\000\000\131\001\131\001\000\000\131\001\000\000\000\000\000\000\ -\000\000\000\000\000\000\131\001\000\000\131\001\131\001\131\001\ -\131\001\231\000\131\001\131\001\000\000\000\000\000\000\000\000\ -\000\000\000\000\232\000\232\000\232\000\232\000\232\000\000\000\ -\000\000\131\001\131\001\131\001\131\001\131\001\131\001\131\001\ -\131\001\131\001\000\000\000\000\000\000\000\000\000\000\243\000\ -\131\001\131\001\000\000\000\000\131\001\131\001\000\000\000\000\ -\131\001\131\001\131\001\131\001\131\001\000\000\000\000\131\001\ -\000\000\034\001\131\001\131\001\131\001\000\000\131\001\000\000\ -\131\001\000\000\000\000\000\000\131\001\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\131\001\131\001\000\000\131\001\ -\131\001\000\000\131\001\000\000\000\000\233\000\233\000\233\000\ -\233\000\233\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\233\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\103\004\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\027\000\027\000\027\000\027\000\243\000\027\000\ -\000\000\027\000\000\000\000\000\027\000\027\000\027\000\027\000\ +\000\000\000\000\065\004\190\001\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\235\000\000\000\000\000\000\000\ +\000\000\235\000\235\000\235\000\000\000\235\000\000\000\000\000\ +\079\004\000\000\235\000\000\000\000\000\000\000\000\000\000\000\ +\190\001\000\000\000\000\000\000\235\000\000\000\000\000\000\000\ +\000\000\205\001\097\004\205\001\000\000\101\004\000\000\047\004\ +\000\000\104\004\000\000\190\001\000\000\235\000\235\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\190\001\190\001\190\001\ +\190\001\190\001\190\001\190\001\000\000\000\000\000\000\235\000\ +\000\000\190\001\000\000\103\002\000\000\000\000\123\004\000\000\ +\000\000\000\000\000\000\000\000\229\000\237\000\237\000\237\000\ +\190\001\237\000\000\000\123\004\000\000\000\000\237\000\146\004\ +\000\000\000\000\000\000\190\001\000\000\000\000\000\000\000\000\ +\237\000\000\000\158\004\159\004\160\004\161\004\162\004\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\237\000\237\000\000\000\000\000\121\001\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\237\000\000\000\141\001\141\001\104\004\ +\000\000\141\001\235\000\000\000\000\000\235\000\000\000\000\000\ +\000\000\235\000\121\001\000\000\141\001\000\000\000\000\141\001\ +\000\000\000\000\000\000\247\000\247\000\247\000\141\001\236\000\ +\123\004\000\000\000\000\000\000\247\000\000\000\141\001\000\000\ +\000\000\000\000\000\000\141\001\000\000\141\001\247\000\000\000\ +\000\000\141\001\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\141\001\141\001\141\001\141\001\141\001\247\000\ +\247\000\000\000\000\000\000\000\000\000\141\001\000\000\141\001\ +\000\000\000\000\000\000\000\000\000\000\141\001\141\001\235\000\ +\205\001\247\000\205\001\141\001\000\000\141\001\235\000\235\000\ +\000\000\000\000\000\000\000\000\141\001\000\000\141\001\141\001\ +\141\001\141\001\000\000\141\001\141\001\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\141\001\141\001\141\001\141\001\141\001\141\001\ +\141\001\141\001\141\001\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\235\000\235\000\000\000\141\001\141\001\000\000\ +\000\000\000\000\000\000\189\001\000\000\000\000\000\000\000\000\ +\141\001\000\000\189\001\237\000\000\000\141\001\000\000\000\000\ +\189\001\141\001\237\000\237\000\025\003\141\001\000\000\028\003\ +\189\001\189\001\235\000\031\003\000\000\141\001\141\001\000\000\ +\141\001\141\001\000\000\141\001\000\000\000\000\000\000\000\000\ +\000\000\000\000\189\001\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\235\000\000\000\000\000\121\001\000\000\000\000\237\000\237\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\189\001\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\247\000\189\001\000\000\000\000\000\000\237\000\000\000\ +\247\000\247\000\000\000\000\000\189\001\189\001\189\001\189\001\ +\189\001\189\001\189\001\000\000\000\000\000\000\000\000\000\000\ +\189\001\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\189\001\ +\205\001\198\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\189\001\000\000\247\000\236\000\000\000\198\000\ +\000\000\027\000\000\000\000\000\000\000\000\000\198\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\198\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\198\000\000\000\000\000\ +\000\000\198\000\000\000\000\000\247\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\198\000\198\000\198\000\198\000\121\001\ +\000\000\000\000\000\000\000\000\000\000\198\000\000\000\198\000\ +\000\000\000\000\000\000\000\000\000\000\198\000\198\000\000\000\ +\000\000\000\000\166\003\198\000\000\000\198\000\000\000\205\001\ +\000\000\000\000\000\000\000\000\198\000\000\000\198\000\198\000\ +\198\000\198\000\000\000\198\000\198\000\000\000\000\000\000\000\ +\205\001\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\198\000\198\000\198\000\198\000\198\000\198\000\ +\198\000\198\000\198\000\000\000\235\000\235\000\235\000\235\000\ +\235\000\000\000\000\000\000\000\000\000\198\000\198\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\198\000\235\000\000\000\000\000\000\000\198\000\000\000\000\000\ +\000\000\198\000\000\000\044\004\000\000\198\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\198\000\198\000\000\000\ +\198\000\198\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\237\000\237\000\237\000\237\000\237\000\036\001\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\089\004\000\000\000\000\237\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\235\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\235\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\236\000\236\000\ +\236\000\236\000\236\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\027\000\027\000\027\000\027\000\000\000\027\000\ +\000\000\027\000\000\000\247\000\027\000\027\000\027\000\027\000\ \027\000\027\000\027\000\000\000\027\000\027\000\027\000\027\000\ \027\000\027\000\027\000\000\000\027\000\027\000\027\000\027\000\ \027\000\027\000\027\000\027\000\027\000\000\000\027\000\027\000\ -\000\000\027\000\027\000\027\000\000\000\027\000\000\000\027\000\ +\000\000\027\000\027\000\027\000\237\000\027\000\000\000\027\000\ \000\000\000\000\027\000\027\000\000\000\027\000\000\000\027\000\ \027\000\027\000\027\000\027\000\027\000\027\000\027\000\027\000\ \027\000\027\000\000\000\027\000\000\000\027\000\027\000\027\000\ -\000\000\233\000\027\000\027\000\027\000\027\000\027\000\027\000\ -\027\000\027\000\051\000\027\000\000\000\027\000\027\000\000\000\ -\000\000\027\000\027\000\027\000\027\000\027\000\027\000\027\000\ +\000\000\000\000\027\000\027\000\027\000\027\000\027\000\027\000\ +\027\000\027\000\000\000\027\000\000\000\027\000\027\000\051\000\ +\107\004\027\000\027\000\027\000\027\000\027\000\027\000\027\000\ \000\000\027\000\027\000\027\000\027\000\027\000\027\000\027\000\ \027\000\027\000\027\000\027\000\000\000\000\000\027\000\027\000\ \027\000\027\000\027\000\027\000\027\000\027\000\027\000\027\000\ -\027\000\000\000\000\000\000\000\000\000\027\000\027\000\027\000\ +\027\000\000\000\247\000\000\000\000\000\027\000\027\000\027\000\ \027\000\027\000\000\000\027\000\027\000\027\000\000\000\027\000\ \027\000\027\000\027\000\027\000\027\000\027\000\027\000\027\000\ \027\000\027\000\027\000\027\000\000\000\027\000\000\000\027\000\ \027\000\000\000\027\000\027\000\027\000\000\000\027\000\027\000\ \027\000\027\000\027\000\027\000\027\000\027\000\027\000\027\000\ -\027\000\027\000\034\001\034\001\034\001\034\001\000\000\034\001\ -\000\000\034\001\000\000\000\000\034\001\034\001\034\001\034\001\ -\034\001\034\001\034\001\000\000\034\001\034\001\034\001\034\001\ -\034\001\034\001\034\001\000\000\034\001\034\001\034\001\000\000\ -\000\000\000\000\000\000\034\001\034\001\000\000\034\001\034\001\ -\000\000\034\001\034\001\034\001\000\000\034\001\000\000\034\001\ -\000\000\000\000\034\001\034\001\000\000\034\001\000\000\034\001\ -\034\001\034\001\034\001\034\001\034\001\034\001\034\001\034\001\ -\034\001\034\001\000\000\034\001\000\000\034\001\034\001\034\001\ -\000\000\000\000\034\001\034\001\034\001\034\001\034\001\034\001\ -\034\001\034\001\050\001\034\001\000\000\034\001\034\001\000\000\ -\000\000\034\001\034\001\034\001\034\001\034\001\034\001\034\001\ -\000\000\034\001\034\001\034\001\034\001\034\001\034\001\034\001\ -\034\001\034\001\034\001\034\001\000\000\000\000\034\001\034\001\ -\034\001\034\001\034\001\034\001\034\001\034\001\034\001\034\001\ -\034\001\034\001\034\001\000\000\000\000\034\001\034\001\034\001\ -\034\001\034\001\000\000\034\001\034\001\034\001\000\000\034\001\ -\034\001\034\001\034\001\034\001\034\001\034\001\034\001\034\001\ -\034\001\034\001\034\001\034\001\000\000\034\001\000\000\034\001\ -\034\001\000\000\034\001\034\001\034\001\034\001\034\001\034\001\ -\034\001\034\001\034\001\034\001\034\001\034\001\034\001\034\001\ -\034\001\034\001\000\000\051\000\051\000\051\000\051\000\000\000\ -\051\000\000\000\051\000\000\000\000\000\051\000\051\000\051\000\ -\051\000\051\000\051\000\051\000\000\000\051\000\051\000\051\000\ -\051\000\051\000\051\000\051\000\000\000\051\000\051\000\051\000\ -\000\000\051\000\051\000\051\000\051\000\051\000\000\000\051\000\ -\051\000\000\000\051\000\051\000\051\000\000\000\051\000\000\000\ -\051\000\000\000\000\000\051\000\051\000\000\000\051\000\000\000\ -\051\000\051\000\051\000\051\000\051\000\051\000\051\000\051\000\ -\051\000\051\000\051\000\000\000\051\000\000\000\051\000\051\000\ -\051\000\000\000\000\000\051\000\051\000\051\000\051\000\051\000\ -\051\000\051\000\051\000\240\001\051\000\000\000\051\000\051\000\ +\027\000\027\000\000\000\000\000\000\000\000\000\036\001\036\001\ +\036\001\036\001\000\000\036\001\000\000\036\001\000\000\000\000\ +\036\001\036\001\036\001\036\001\036\001\036\001\036\001\000\000\ +\036\001\036\001\036\001\036\001\036\001\036\001\036\001\000\000\ +\036\001\036\001\036\001\000\000\000\000\000\000\000\000\036\001\ +\036\001\000\000\036\001\036\001\000\000\036\001\036\001\036\001\ +\000\000\036\001\000\000\036\001\000\000\000\000\036\001\036\001\ +\000\000\036\001\000\000\036\001\036\001\036\001\036\001\036\001\ +\036\001\036\001\036\001\036\001\036\001\036\001\000\000\036\001\ +\000\000\036\001\036\001\036\001\000\000\000\000\036\001\036\001\ +\036\001\036\001\036\001\036\001\036\001\036\001\000\000\036\001\ +\052\001\036\001\036\001\000\000\000\000\036\001\036\001\036\001\ +\036\001\036\001\036\001\036\001\000\000\036\001\036\001\036\001\ +\036\001\036\001\036\001\036\001\036\001\036\001\036\001\036\001\ +\000\000\000\000\036\001\036\001\036\001\036\001\036\001\036\001\ +\036\001\036\001\036\001\036\001\036\001\036\001\036\001\000\000\ +\000\000\036\001\036\001\036\001\036\001\036\001\000\000\036\001\ +\036\001\036\001\000\000\036\001\036\001\036\001\036\001\036\001\ +\036\001\036\001\036\001\036\001\036\001\036\001\036\001\036\001\ +\000\000\036\001\000\000\036\001\036\001\000\000\036\001\036\001\ +\036\001\036\001\036\001\036\001\036\001\036\001\036\001\036\001\ +\036\001\036\001\036\001\036\001\036\001\036\001\000\000\000\000\ +\051\000\051\000\051\000\051\000\000\000\051\000\000\000\051\000\ \000\000\000\000\051\000\051\000\051\000\051\000\051\000\051\000\ \051\000\000\000\051\000\051\000\051\000\051\000\051\000\051\000\ -\051\000\051\000\051\000\051\000\051\000\000\000\000\000\051\000\ +\051\000\000\000\051\000\051\000\051\000\000\000\051\000\051\000\ +\051\000\051\000\051\000\000\000\051\000\051\000\000\000\051\000\ +\051\000\051\000\000\000\051\000\000\000\051\000\000\000\000\000\ +\051\000\051\000\000\000\051\000\000\000\051\000\051\000\051\000\ \051\000\051\000\051\000\051\000\051\000\051\000\051\000\051\000\ -\051\000\051\000\000\000\000\000\000\000\000\000\051\000\051\000\ -\051\000\051\000\051\000\000\000\051\000\051\000\051\000\000\000\ +\000\000\051\000\000\000\051\000\051\000\051\000\000\000\000\000\ \051\000\051\000\051\000\051\000\051\000\051\000\051\000\051\000\ -\051\000\051\000\051\000\051\000\051\000\000\000\051\000\000\000\ -\051\000\051\000\000\000\051\000\051\000\051\000\000\000\051\000\ +\242\001\051\000\000\000\051\000\051\000\000\000\000\000\051\000\ +\051\000\051\000\051\000\051\000\051\000\051\000\000\000\051\000\ \051\000\051\000\051\000\051\000\051\000\051\000\051\000\051\000\ -\051\000\051\000\051\000\050\001\050\001\050\001\050\001\000\000\ -\050\001\000\000\050\001\000\000\000\000\000\000\050\001\050\001\ -\050\001\050\001\050\001\050\001\000\000\050\001\050\001\048\001\ -\050\001\050\001\050\001\050\001\000\000\050\001\050\001\000\000\ -\050\001\050\001\050\001\050\001\050\001\050\001\000\000\050\001\ -\050\001\000\000\050\001\050\001\050\001\000\000\050\001\000\000\ -\050\001\000\000\000\000\050\001\050\001\000\000\050\001\000\000\ -\050\001\050\001\050\001\050\001\050\001\050\001\050\001\050\001\ -\050\001\050\001\050\001\000\000\050\001\000\000\050\001\050\001\ -\000\000\000\000\000\000\050\001\050\001\050\001\000\000\050\001\ -\050\001\050\001\050\001\242\001\050\001\048\001\050\001\050\001\ -\000\000\000\000\050\001\050\001\050\001\050\001\050\001\050\001\ -\050\001\000\000\050\001\050\001\050\001\050\001\050\001\050\001\ -\050\001\050\001\050\001\050\001\050\001\000\000\000\000\050\001\ -\050\001\050\001\050\001\050\001\050\001\050\001\050\001\050\001\ -\050\001\050\001\000\000\000\000\000\000\000\000\050\001\050\001\ -\050\001\050\001\050\001\000\000\050\001\050\001\050\001\000\000\ -\050\001\050\001\050\001\050\001\050\001\050\001\050\001\050\001\ -\050\001\050\001\050\001\050\001\050\001\000\000\050\001\000\000\ -\050\001\050\001\000\000\050\001\050\001\000\000\050\001\050\001\ -\050\001\050\001\050\001\050\001\050\001\050\001\050\001\050\001\ -\050\001\000\000\050\001\000\000\240\001\240\001\240\001\240\001\ -\000\000\240\001\000\000\240\001\000\000\000\000\240\001\240\001\ -\240\001\240\001\240\001\240\001\240\001\000\000\240\001\240\001\ -\240\001\240\001\240\001\240\001\240\001\000\000\240\001\240\001\ -\240\001\000\000\000\000\000\000\000\000\000\000\240\001\000\000\ -\240\001\240\001\000\000\240\001\240\001\240\001\000\000\240\001\ -\000\000\240\001\000\000\000\000\240\001\240\001\000\000\240\001\ -\000\000\240\001\240\001\240\001\240\001\240\001\240\001\240\001\ -\240\001\240\001\240\001\240\001\000\000\240\001\000\000\240\001\ -\240\001\240\001\000\000\000\000\240\001\240\001\240\001\240\001\ -\240\001\240\001\240\001\240\001\242\001\240\001\000\000\240\001\ -\240\001\000\000\000\000\240\001\240\001\240\001\240\001\240\001\ -\240\001\240\001\000\000\240\001\240\001\240\001\240\001\240\001\ -\240\001\240\001\240\001\240\001\240\001\240\001\000\000\000\000\ -\240\001\240\001\240\001\240\001\240\001\240\001\240\001\240\001\ -\240\001\240\001\240\001\000\000\000\000\000\000\000\000\240\001\ -\240\001\240\001\240\001\240\001\000\000\240\001\240\001\240\001\ -\000\000\240\001\240\001\240\001\240\001\240\001\240\001\240\001\ -\240\001\240\001\240\001\240\001\240\001\240\001\000\000\240\001\ -\000\000\240\001\240\001\000\000\240\001\240\001\240\001\000\000\ -\240\001\240\001\240\001\240\001\240\001\240\001\240\001\240\001\ -\240\001\240\001\240\001\240\001\242\001\242\001\242\001\242\001\ -\000\000\242\001\000\000\242\001\000\000\000\000\000\000\242\001\ -\242\001\242\001\242\001\242\001\242\001\000\000\242\001\242\001\ -\049\001\242\001\242\001\242\001\242\001\000\000\242\001\242\001\ -\000\000\000\000\242\001\242\001\242\001\242\001\242\001\000\000\ -\242\001\242\001\000\000\242\001\242\001\242\001\000\000\242\001\ -\000\000\242\001\000\000\000\000\242\001\242\001\000\000\242\001\ -\000\000\242\001\242\001\242\001\242\001\242\001\242\001\242\001\ -\242\001\242\001\242\001\242\001\000\000\242\001\000\000\242\001\ -\242\001\000\000\000\000\000\000\242\001\242\001\242\001\000\000\ -\242\001\242\001\242\001\242\001\237\001\242\001\000\000\242\001\ +\051\000\051\000\000\000\000\000\051\000\051\000\051\000\051\000\ +\051\000\051\000\051\000\051\000\051\000\051\000\051\000\000\000\ +\000\000\000\000\000\000\051\000\051\000\051\000\051\000\051\000\ +\000\000\051\000\051\000\051\000\000\000\051\000\051\000\051\000\ +\051\000\051\000\051\000\051\000\051\000\051\000\051\000\051\000\ +\051\000\051\000\000\000\051\000\000\000\051\000\051\000\000\000\ +\051\000\051\000\051\000\000\000\051\000\051\000\051\000\051\000\ +\051\000\051\000\051\000\051\000\051\000\051\000\051\000\051\000\ +\000\000\052\001\052\001\052\001\052\001\000\000\052\001\000\000\ +\052\001\000\000\000\000\000\000\052\001\052\001\052\001\052\001\ +\052\001\052\001\000\000\052\001\052\001\050\001\052\001\052\001\ +\052\001\052\001\000\000\052\001\052\001\000\000\052\001\052\001\ +\052\001\052\001\052\001\052\001\000\000\052\001\052\001\000\000\ +\052\001\052\001\052\001\000\000\052\001\000\000\052\001\000\000\ +\000\000\052\001\052\001\000\000\052\001\000\000\052\001\052\001\ +\052\001\052\001\052\001\052\001\052\001\052\001\052\001\052\001\ +\052\001\000\000\052\001\000\000\052\001\052\001\000\000\000\000\ +\000\000\052\001\052\001\052\001\000\000\052\001\052\001\052\001\ +\052\001\244\001\052\001\050\001\052\001\052\001\000\000\000\000\ +\052\001\052\001\052\001\052\001\052\001\052\001\052\001\000\000\ +\052\001\052\001\052\001\052\001\052\001\052\001\052\001\052\001\ +\052\001\052\001\052\001\000\000\000\000\052\001\052\001\052\001\ +\052\001\052\001\052\001\052\001\052\001\052\001\052\001\052\001\ +\000\000\000\000\000\000\000\000\052\001\052\001\052\001\052\001\ +\052\001\000\000\052\001\052\001\052\001\000\000\052\001\052\001\ +\052\001\052\001\052\001\052\001\052\001\052\001\052\001\052\001\ +\052\001\052\001\052\001\000\000\052\001\000\000\052\001\052\001\ +\000\000\052\001\052\001\000\000\052\001\052\001\052\001\052\001\ +\052\001\052\001\052\001\052\001\052\001\052\001\052\001\000\000\ +\052\001\242\001\242\001\242\001\242\001\000\000\242\001\000\000\ \242\001\000\000\000\000\242\001\242\001\242\001\242\001\242\001\ \242\001\242\001\000\000\242\001\242\001\242\001\242\001\242\001\ -\242\001\242\001\242\001\242\001\242\001\242\001\000\000\000\000\ +\242\001\242\001\000\000\242\001\242\001\242\001\000\000\000\000\ +\000\000\000\000\000\000\242\001\000\000\242\001\242\001\000\000\ +\242\001\242\001\242\001\000\000\242\001\000\000\242\001\000\000\ +\000\000\242\001\242\001\000\000\242\001\000\000\242\001\242\001\ \242\001\242\001\242\001\242\001\242\001\242\001\242\001\242\001\ -\242\001\242\001\242\001\000\000\000\000\000\000\000\000\242\001\ -\242\001\242\001\242\001\242\001\000\000\242\001\242\001\242\001\ +\242\001\000\000\242\001\000\000\242\001\242\001\242\001\000\000\ \000\000\242\001\242\001\242\001\242\001\242\001\242\001\242\001\ -\242\001\242\001\242\001\242\001\242\001\242\001\000\000\242\001\ -\000\000\242\001\242\001\000\000\242\001\242\001\000\000\000\000\ -\242\001\242\001\242\001\242\001\242\001\242\001\242\001\242\001\ -\242\001\242\001\000\000\242\001\000\000\242\001\242\001\242\001\ -\242\001\000\000\242\001\000\000\242\001\000\000\000\000\000\000\ -\242\001\242\001\000\000\242\001\242\001\000\000\000\000\242\001\ -\242\001\049\001\242\001\242\001\242\001\242\001\000\000\242\001\ -\242\001\000\000\000\000\242\001\242\001\242\001\242\001\242\001\ -\000\000\242\001\242\001\000\000\242\001\242\001\242\001\000\000\ -\242\001\000\000\242\001\000\000\000\000\242\001\242\001\000\000\ -\242\001\000\000\242\001\242\001\242\001\242\001\242\001\242\001\ -\242\001\242\001\242\001\242\001\242\001\000\000\242\001\000\000\ -\242\001\242\001\000\000\000\000\000\000\242\001\242\001\242\001\ -\000\000\242\001\242\001\242\001\242\001\083\000\242\001\049\001\ -\242\001\242\001\000\000\000\000\242\001\242\001\242\001\242\001\ -\242\001\242\001\242\001\000\000\242\001\242\001\242\001\242\001\ +\242\001\244\001\242\001\000\000\242\001\242\001\000\000\000\000\ \242\001\242\001\242\001\242\001\242\001\242\001\242\001\000\000\ -\000\000\242\001\242\001\242\001\242\001\242\001\242\001\242\001\ -\242\001\242\001\242\001\242\001\000\000\000\000\000\000\000\000\ -\242\001\242\001\242\001\242\001\242\001\000\000\242\001\242\001\ -\242\001\000\000\242\001\242\001\242\001\242\001\242\001\242\001\ -\242\001\242\001\000\000\242\001\242\001\242\001\242\001\000\000\ -\242\001\000\000\242\001\242\001\000\000\242\001\242\001\000\000\ -\000\000\242\001\242\001\242\001\242\001\242\001\242\001\242\001\ -\242\001\242\001\242\001\000\000\242\001\237\001\237\001\237\001\ -\237\001\000\000\237\001\000\000\237\001\000\000\000\000\000\000\ -\237\001\237\001\237\001\237\001\237\001\237\001\000\000\237\001\ -\237\001\000\000\237\001\237\001\237\001\237\001\000\000\237\001\ -\237\001\000\000\000\000\000\000\000\000\000\000\000\000\237\001\ -\000\000\237\001\237\001\000\000\237\001\237\001\237\001\000\000\ -\237\001\000\000\237\001\000\000\000\000\237\001\237\001\000\000\ -\237\001\000\000\237\001\237\001\237\001\237\001\237\001\237\001\ -\237\001\237\001\237\001\237\001\237\001\000\000\237\001\000\000\ -\237\001\237\001\000\000\000\000\000\000\237\001\237\001\237\001\ -\000\000\237\001\237\001\237\001\237\001\084\000\237\001\000\000\ -\237\001\237\001\000\000\000\000\237\001\237\001\237\001\237\001\ -\237\001\237\001\237\001\000\000\237\001\237\001\237\001\237\001\ -\237\001\237\001\237\001\237\001\237\001\237\001\237\001\000\000\ -\000\000\237\001\237\001\237\001\237\001\237\001\237\001\237\001\ -\237\001\237\001\237\001\237\001\000\000\000\000\000\000\000\000\ -\237\001\237\001\237\001\237\001\237\001\000\000\237\001\237\001\ -\237\001\000\000\237\001\237\001\237\001\237\001\237\001\237\001\ -\237\001\237\001\237\001\237\001\237\001\237\001\237\001\000\000\ -\237\001\000\000\237\001\237\001\000\000\237\001\237\001\000\000\ -\000\000\237\001\237\001\237\001\237\001\237\001\237\001\237\001\ -\237\001\237\001\237\001\000\000\237\001\000\000\083\000\083\000\ +\242\001\242\001\242\001\242\001\242\001\242\001\242\001\242\001\ +\242\001\242\001\242\001\000\000\000\000\242\001\242\001\242\001\ +\242\001\242\001\242\001\242\001\242\001\242\001\242\001\242\001\ +\000\000\000\000\000\000\000\000\242\001\242\001\242\001\242\001\ +\242\001\000\000\242\001\242\001\242\001\000\000\242\001\242\001\ +\242\001\242\001\242\001\242\001\242\001\242\001\242\001\242\001\ +\242\001\242\001\242\001\000\000\242\001\000\000\242\001\242\001\ +\000\000\242\001\242\001\242\001\000\000\242\001\242\001\242\001\ +\242\001\242\001\242\001\242\001\242\001\242\001\242\001\242\001\ +\242\001\000\000\244\001\244\001\244\001\244\001\000\000\244\001\ +\000\000\244\001\000\000\000\000\000\000\244\001\244\001\244\001\ +\244\001\244\001\244\001\000\000\244\001\244\001\051\001\244\001\ +\244\001\244\001\244\001\000\000\244\001\244\001\000\000\000\000\ +\244\001\244\001\244\001\244\001\244\001\000\000\244\001\244\001\ +\000\000\244\001\244\001\244\001\000\000\244\001\000\000\244\001\ +\000\000\000\000\244\001\244\001\000\000\244\001\000\000\244\001\ +\244\001\244\001\244\001\244\001\244\001\244\001\244\001\244\001\ +\244\001\244\001\000\000\244\001\000\000\244\001\244\001\000\000\ +\000\000\000\000\244\001\244\001\244\001\000\000\244\001\244\001\ +\244\001\244\001\239\001\244\001\000\000\244\001\244\001\000\000\ +\000\000\244\001\244\001\244\001\244\001\244\001\244\001\244\001\ +\000\000\244\001\244\001\244\001\244\001\244\001\244\001\244\001\ +\244\001\244\001\244\001\244\001\000\000\000\000\244\001\244\001\ +\244\001\244\001\244\001\244\001\244\001\244\001\244\001\244\001\ +\244\001\000\000\000\000\000\000\000\000\244\001\244\001\244\001\ +\244\001\244\001\000\000\244\001\244\001\244\001\000\000\244\001\ +\244\001\244\001\244\001\244\001\244\001\244\001\244\001\244\001\ +\244\001\244\001\244\001\244\001\000\000\244\001\000\000\244\001\ +\244\001\000\000\244\001\244\001\000\000\000\000\244\001\244\001\ +\244\001\244\001\244\001\244\001\244\001\244\001\244\001\244\001\ +\000\000\244\001\244\001\244\001\244\001\244\001\000\000\244\001\ +\000\000\244\001\000\000\000\000\000\000\244\001\244\001\000\000\ +\244\001\244\001\000\000\000\000\244\001\244\001\051\001\244\001\ +\244\001\244\001\244\001\000\000\244\001\244\001\000\000\000\000\ +\244\001\244\001\244\001\244\001\244\001\000\000\244\001\244\001\ +\000\000\244\001\244\001\244\001\000\000\244\001\000\000\244\001\ +\000\000\000\000\244\001\244\001\000\000\244\001\000\000\244\001\ +\244\001\244\001\244\001\244\001\244\001\244\001\244\001\244\001\ +\244\001\244\001\000\000\244\001\000\000\244\001\244\001\000\000\ +\000\000\000\000\244\001\244\001\244\001\000\000\244\001\244\001\ +\244\001\244\001\083\000\244\001\051\001\244\001\244\001\000\000\ +\000\000\244\001\244\001\244\001\244\001\244\001\244\001\244\001\ +\000\000\244\001\244\001\244\001\244\001\244\001\244\001\244\001\ +\244\001\244\001\244\001\244\001\000\000\000\000\244\001\244\001\ +\244\001\244\001\244\001\244\001\244\001\244\001\244\001\244\001\ +\244\001\000\000\000\000\000\000\000\000\244\001\244\001\244\001\ +\244\001\244\001\000\000\244\001\244\001\244\001\000\000\244\001\ +\244\001\244\001\244\001\244\001\244\001\244\001\244\001\000\000\ +\244\001\244\001\244\001\244\001\000\000\244\001\000\000\244\001\ +\244\001\000\000\244\001\244\001\000\000\000\000\244\001\244\001\ +\244\001\244\001\244\001\244\001\244\001\244\001\244\001\244\001\ +\000\000\244\001\000\000\239\001\239\001\239\001\239\001\000\000\ +\239\001\000\000\239\001\000\000\000\000\000\000\239\001\239\001\ +\239\001\239\001\239\001\239\001\000\000\239\001\239\001\000\000\ +\239\001\239\001\239\001\239\001\000\000\239\001\239\001\000\000\ +\000\000\000\000\000\000\000\000\000\000\239\001\000\000\239\001\ +\239\001\000\000\239\001\239\001\239\001\000\000\239\001\000\000\ +\239\001\000\000\000\000\239\001\239\001\000\000\239\001\000\000\ +\239\001\239\001\239\001\239\001\239\001\239\001\239\001\239\001\ +\239\001\239\001\239\001\000\000\239\001\000\000\239\001\239\001\ +\000\000\000\000\000\000\239\001\239\001\239\001\000\000\239\001\ +\239\001\239\001\239\001\084\000\239\001\000\000\239\001\239\001\ +\000\000\000\000\239\001\239\001\239\001\239\001\239\001\239\001\ +\239\001\000\000\239\001\239\001\239\001\239\001\239\001\239\001\ +\239\001\239\001\239\001\239\001\239\001\000\000\000\000\239\001\ +\239\001\239\001\239\001\239\001\239\001\239\001\239\001\239\001\ +\239\001\239\001\000\000\000\000\000\000\000\000\239\001\239\001\ +\239\001\239\001\239\001\000\000\239\001\239\001\239\001\000\000\ +\239\001\239\001\239\001\239\001\239\001\239\001\239\001\239\001\ +\239\001\239\001\239\001\239\001\239\001\000\000\239\001\000\000\ +\239\001\239\001\000\000\239\001\239\001\000\000\000\000\239\001\ +\239\001\239\001\239\001\239\001\239\001\239\001\239\001\239\001\ +\239\001\000\000\239\001\083\000\083\000\083\000\083\000\000\000\ +\083\000\000\000\083\000\000\000\000\000\000\000\083\000\083\000\ +\083\000\083\000\083\000\083\000\000\000\083\000\083\000\000\000\ +\083\000\083\000\083\000\083\000\000\000\083\000\083\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\083\000\000\000\083\000\ +\083\000\000\000\083\000\083\000\083\000\000\000\083\000\000\000\ +\083\000\000\000\000\000\083\000\083\000\000\000\083\000\000\000\ +\083\000\083\000\083\000\083\000\083\000\083\000\083\000\083\000\ +\083\000\083\000\083\000\000\000\083\000\000\000\083\000\083\000\ +\000\000\000\000\000\000\083\000\083\000\083\000\000\000\083\000\ +\083\000\083\000\083\000\085\000\083\000\000\000\083\000\083\000\ +\000\000\000\000\083\000\083\000\083\000\083\000\083\000\083\000\ +\083\000\000\000\083\000\083\000\083\000\083\000\083\000\083\000\ +\083\000\083\000\083\000\083\000\083\000\000\000\000\000\083\000\ +\083\000\083\000\083\000\083\000\083\000\083\000\083\000\083\000\ +\083\000\083\000\000\000\000\000\000\000\000\000\083\000\083\000\ +\083\000\083\000\083\000\000\000\083\000\083\000\083\000\000\000\ +\083\000\083\000\083\000\083\000\083\000\083\000\083\000\083\000\ +\083\000\083\000\083\000\083\000\083\000\000\000\083\000\000\000\ +\083\000\083\000\000\000\083\000\083\000\000\000\000\000\083\000\ +\083\000\083\000\083\000\083\000\083\000\083\000\083\000\083\000\ +\083\000\000\000\083\000\000\000\084\000\084\000\084\000\084\000\ +\000\000\084\000\000\000\084\000\000\000\000\000\000\000\084\000\ +\084\000\084\000\084\000\084\000\084\000\000\000\084\000\084\000\ +\000\000\084\000\084\000\084\000\084\000\000\000\084\000\084\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\084\000\000\000\ +\084\000\084\000\000\000\084\000\084\000\084\000\000\000\084\000\ +\000\000\084\000\000\000\000\000\084\000\084\000\000\000\084\000\ +\000\000\084\000\084\000\084\000\084\000\084\000\084\000\084\000\ +\084\000\084\000\084\000\084\000\000\000\084\000\000\000\084\000\ +\084\000\000\000\000\000\000\000\084\000\084\000\084\000\000\000\ +\084\000\084\000\084\000\084\000\082\000\084\000\000\000\084\000\ +\084\000\000\000\000\000\084\000\084\000\084\000\084\000\084\000\ +\084\000\084\000\000\000\084\000\084\000\084\000\084\000\084\000\ +\084\000\084\000\084\000\084\000\084\000\084\000\000\000\000\000\ +\084\000\084\000\084\000\084\000\084\000\084\000\084\000\084\000\ +\084\000\084\000\084\000\000\000\000\000\000\000\000\000\084\000\ +\084\000\084\000\084\000\084\000\000\000\084\000\084\000\084\000\ +\000\000\084\000\084\000\084\000\084\000\084\000\084\000\084\000\ +\084\000\084\000\084\000\084\000\084\000\084\000\000\000\084\000\ +\000\000\084\000\084\000\000\000\084\000\084\000\000\000\000\000\ +\084\000\084\000\084\000\084\000\084\000\084\000\084\000\084\000\ +\084\000\084\000\000\000\084\000\085\000\085\000\085\000\085\000\ +\000\000\085\000\000\000\085\000\000\000\000\000\000\000\085\000\ +\085\000\085\000\085\000\085\000\085\000\000\000\085\000\085\000\ +\000\000\085\000\085\000\085\000\085\000\000\000\085\000\085\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\085\000\000\000\ +\085\000\085\000\000\000\085\000\085\000\085\000\000\000\085\000\ +\000\000\085\000\000\000\000\000\085\000\085\000\000\000\085\000\ +\000\000\085\000\085\000\085\000\085\000\085\000\085\000\085\000\ +\085\000\085\000\085\000\085\000\000\000\085\000\000\000\085\000\ +\085\000\000\000\000\000\000\000\085\000\085\000\085\000\000\000\ +\085\000\085\000\085\000\085\000\239\001\085\000\000\000\085\000\ +\085\000\000\000\000\000\085\000\085\000\085\000\085\000\085\000\ +\085\000\085\000\000\000\085\000\085\000\085\000\085\000\085\000\ +\085\000\085\000\085\000\085\000\085\000\085\000\000\000\000\000\ +\085\000\085\000\085\000\085\000\085\000\085\000\085\000\085\000\ +\085\000\085\000\085\000\000\000\000\000\000\000\000\000\085\000\ +\085\000\085\000\085\000\085\000\000\000\085\000\085\000\085\000\ +\000\000\085\000\085\000\085\000\085\000\085\000\085\000\085\000\ +\085\000\085\000\085\000\085\000\085\000\085\000\000\000\085\000\ +\000\000\085\000\085\000\000\000\085\000\085\000\000\000\000\000\ +\085\000\085\000\085\000\085\000\085\000\085\000\085\000\085\000\ +\085\000\085\000\000\000\085\000\000\000\082\000\082\000\082\000\ +\082\000\000\000\082\000\000\000\082\000\000\000\000\000\000\000\ +\082\000\082\000\082\000\082\000\082\000\082\000\000\000\082\000\ +\082\000\000\000\082\000\082\000\082\000\082\000\000\000\082\000\ +\082\000\000\000\000\000\000\000\000\000\000\000\000\000\082\000\ +\000\000\082\000\082\000\000\000\082\000\082\000\082\000\000\000\ +\082\000\000\000\082\000\000\000\000\000\082\000\082\000\000\000\ +\082\000\000\000\082\000\082\000\082\000\082\000\082\000\082\000\ +\082\000\082\000\082\000\082\000\082\000\000\000\082\000\000\000\ +\082\000\082\000\000\000\000\000\000\000\082\000\082\000\082\000\ +\000\000\082\000\082\000\082\000\082\000\083\000\082\000\000\000\ +\082\000\082\000\000\000\000\000\082\000\082\000\082\000\082\000\ +\082\000\082\000\082\000\000\000\082\000\082\000\082\000\082\000\ +\082\000\082\000\082\000\082\000\082\000\082\000\082\000\000\000\ +\000\000\082\000\082\000\082\000\082\000\082\000\082\000\082\000\ +\082\000\082\000\082\000\082\000\000\000\000\000\000\000\000\000\ +\082\000\082\000\082\000\082\000\082\000\000\000\082\000\082\000\ +\082\000\000\000\082\000\082\000\082\000\082\000\082\000\082\000\ +\082\000\082\000\082\000\082\000\082\000\082\000\082\000\000\000\ +\082\000\000\000\082\000\082\000\000\000\082\000\082\000\000\000\ +\000\000\082\000\082\000\082\000\082\000\082\000\082\000\082\000\ +\082\000\082\000\082\000\000\000\082\000\239\001\239\001\239\001\ +\239\001\000\000\239\001\000\000\239\001\000\000\000\000\000\000\ +\239\001\239\001\000\000\239\001\239\001\000\000\000\000\239\001\ +\239\001\000\000\239\001\239\001\239\001\239\001\000\000\239\001\ +\239\001\000\000\000\000\000\000\000\000\000\000\000\000\239\001\ +\000\000\239\001\239\001\000\000\239\001\239\001\239\001\000\000\ +\239\001\000\000\239\001\000\000\000\000\239\001\239\001\000\000\ +\239\001\000\000\239\001\239\001\239\001\239\001\239\001\239\001\ +\239\001\239\001\239\001\239\001\239\001\000\000\239\001\000\000\ +\239\001\239\001\000\000\000\000\000\000\239\001\239\001\239\001\ +\000\000\239\001\239\001\239\001\239\001\084\000\239\001\000\000\ +\239\001\239\001\000\000\000\000\239\001\239\001\239\001\239\001\ +\239\001\239\001\239\001\000\000\239\001\239\001\239\001\239\001\ +\239\001\239\001\239\001\239\001\239\001\239\001\239\001\000\000\ +\000\000\239\001\239\001\239\001\239\001\239\001\239\001\239\001\ +\239\001\239\001\239\001\239\001\000\000\000\000\000\000\000\000\ +\239\001\239\001\239\001\239\001\239\001\000\000\239\001\239\001\ +\239\001\000\000\239\001\239\001\239\001\239\001\239\001\239\001\ +\239\001\239\001\000\000\239\001\239\001\239\001\239\001\000\000\ +\239\001\000\000\239\001\239\001\000\000\239\001\239\001\000\000\ +\000\000\239\001\239\001\239\001\239\001\239\001\239\001\239\001\ +\239\001\239\001\239\001\000\000\239\001\000\000\083\000\083\000\ \083\000\083\000\000\000\083\000\000\000\083\000\000\000\000\000\ -\000\000\083\000\083\000\083\000\083\000\083\000\083\000\000\000\ +\000\000\083\000\083\000\000\000\083\000\083\000\000\000\000\000\ \083\000\083\000\000\000\083\000\083\000\083\000\083\000\000\000\ \083\000\083\000\000\000\000\000\000\000\000\000\000\000\000\000\ \083\000\000\000\083\000\083\000\000\000\083\000\083\000\083\000\ @@ -1706,12 +1836,12 @@ let yytable = "\118\000\ \083\000\083\000\083\000\083\000\083\000\000\000\000\000\000\000\ \000\000\083\000\083\000\083\000\083\000\083\000\000\000\083\000\ \083\000\083\000\000\000\083\000\083\000\083\000\083\000\083\000\ -\083\000\083\000\083\000\083\000\083\000\083\000\083\000\083\000\ +\083\000\083\000\083\000\000\000\083\000\083\000\083\000\083\000\ \000\000\083\000\000\000\083\000\083\000\000\000\083\000\083\000\ \000\000\000\000\083\000\083\000\083\000\083\000\083\000\083\000\ \083\000\083\000\083\000\083\000\000\000\083\000\084\000\084\000\ \084\000\084\000\000\000\084\000\000\000\084\000\000\000\000\000\ -\000\000\084\000\084\000\084\000\084\000\084\000\084\000\000\000\ +\000\000\084\000\084\000\000\000\084\000\084\000\000\000\000\000\ \084\000\084\000\000\000\084\000\084\000\084\000\084\000\000\000\ \084\000\084\000\000\000\000\000\000\000\000\000\000\000\000\000\ \084\000\000\000\084\000\084\000\000\000\084\000\084\000\084\000\ @@ -1727,12 +1857,12 @@ let yytable = "\118\000\ \084\000\084\000\084\000\084\000\084\000\000\000\000\000\000\000\ \000\000\084\000\084\000\084\000\084\000\084\000\000\000\084\000\ \084\000\084\000\000\000\084\000\084\000\084\000\084\000\084\000\ -\084\000\084\000\084\000\084\000\084\000\084\000\084\000\084\000\ +\084\000\084\000\084\000\000\000\084\000\084\000\084\000\084\000\ \000\000\084\000\000\000\084\000\084\000\000\000\084\000\084\000\ \000\000\000\000\084\000\084\000\084\000\084\000\084\000\084\000\ \084\000\084\000\084\000\084\000\000\000\084\000\000\000\085\000\ \085\000\085\000\085\000\000\000\085\000\000\000\085\000\000\000\ -\000\000\000\000\085\000\085\000\085\000\085\000\085\000\085\000\ +\000\000\000\000\085\000\085\000\000\000\085\000\085\000\000\000\ \000\000\085\000\085\000\000\000\085\000\085\000\085\000\085\000\ \000\000\085\000\085\000\000\000\000\000\000\000\000\000\000\000\ \000\000\085\000\000\000\085\000\085\000\000\000\085\000\085\000\ @@ -1740,7 +1870,7 @@ let yytable = "\118\000\ \085\000\000\000\085\000\000\000\085\000\085\000\085\000\085\000\ \085\000\085\000\085\000\085\000\085\000\085\000\085\000\000\000\ \085\000\000\000\085\000\085\000\000\000\000\000\000\000\085\000\ -\085\000\085\000\000\000\085\000\085\000\085\000\085\000\237\001\ +\085\000\085\000\000\000\085\000\085\000\085\000\085\000\056\000\ \085\000\000\000\085\000\085\000\000\000\000\000\085\000\085\000\ \085\000\085\000\085\000\085\000\085\000\000\000\085\000\085\000\ \085\000\085\000\085\000\085\000\085\000\085\000\085\000\085\000\ @@ -1748,12 +1878,12 @@ let yytable = "\118\000\ \085\000\085\000\085\000\085\000\085\000\085\000\000\000\000\000\ \000\000\000\000\085\000\085\000\085\000\085\000\085\000\000\000\ \085\000\085\000\085\000\000\000\085\000\085\000\085\000\085\000\ -\085\000\085\000\085\000\085\000\085\000\085\000\085\000\085\000\ +\085\000\085\000\085\000\085\000\000\000\085\000\085\000\085\000\ \085\000\000\000\085\000\000\000\085\000\085\000\000\000\085\000\ \085\000\000\000\000\000\085\000\085\000\085\000\085\000\085\000\ \085\000\085\000\085\000\085\000\085\000\000\000\085\000\082\000\ \082\000\082\000\082\000\000\000\082\000\000\000\082\000\000\000\ -\000\000\000\000\082\000\082\000\082\000\082\000\082\000\082\000\ +\000\000\000\000\082\000\082\000\000\000\082\000\082\000\000\000\ \000\000\082\000\082\000\000\000\082\000\082\000\082\000\082\000\ \000\000\082\000\082\000\000\000\000\000\000\000\000\000\000\000\ \000\000\082\000\000\000\082\000\082\000\000\000\082\000\082\000\ @@ -1761,7 +1891,7 @@ let yytable = "\118\000\ \082\000\000\000\082\000\000\000\082\000\082\000\082\000\082\000\ \082\000\082\000\082\000\082\000\082\000\082\000\082\000\000\000\ \082\000\000\000\082\000\082\000\000\000\000\000\000\000\082\000\ -\082\000\082\000\000\000\082\000\082\000\082\000\082\000\083\000\ +\082\000\082\000\000\000\082\000\082\000\082\000\082\000\047\000\ \082\000\000\000\082\000\082\000\000\000\000\000\082\000\082\000\ \082\000\082\000\082\000\082\000\082\000\000\000\082\000\082\000\ \082\000\082\000\082\000\082\000\082\000\082\000\082\000\082\000\ @@ -1769,566 +1899,499 @@ let yytable = "\118\000\ \082\000\082\000\082\000\082\000\082\000\082\000\000\000\000\000\ \000\000\000\000\082\000\082\000\082\000\082\000\082\000\000\000\ \082\000\082\000\082\000\000\000\082\000\082\000\082\000\082\000\ -\082\000\082\000\082\000\082\000\082\000\082\000\082\000\082\000\ +\082\000\082\000\082\000\082\000\000\000\082\000\082\000\082\000\ \082\000\000\000\082\000\000\000\082\000\082\000\000\000\082\000\ \082\000\000\000\000\000\082\000\082\000\082\000\082\000\082\000\ \082\000\082\000\082\000\082\000\082\000\000\000\082\000\000\000\ -\237\001\237\001\237\001\237\001\000\000\237\001\000\000\237\001\ -\000\000\000\000\000\000\237\001\237\001\000\000\237\001\237\001\ -\000\000\000\000\237\001\237\001\000\000\237\001\237\001\237\001\ -\237\001\000\000\237\001\237\001\000\000\000\000\000\000\000\000\ -\000\000\000\000\237\001\000\000\237\001\237\001\000\000\237\001\ -\237\001\237\001\000\000\237\001\000\000\237\001\000\000\000\000\ -\237\001\237\001\000\000\237\001\000\000\237\001\237\001\237\001\ -\237\001\237\001\237\001\237\001\237\001\237\001\237\001\237\001\ -\000\000\237\001\000\000\237\001\237\001\000\000\000\000\000\000\ -\237\001\237\001\237\001\000\000\237\001\237\001\237\001\237\001\ -\084\000\237\001\000\000\237\001\237\001\000\000\000\000\237\001\ -\237\001\237\001\237\001\237\001\237\001\237\001\000\000\237\001\ -\237\001\237\001\237\001\237\001\237\001\237\001\237\001\237\001\ -\237\001\237\001\000\000\000\000\237\001\237\001\237\001\237\001\ -\237\001\237\001\237\001\237\001\237\001\237\001\237\001\000\000\ -\000\000\000\000\000\000\237\001\237\001\237\001\237\001\237\001\ -\000\000\237\001\237\001\237\001\000\000\237\001\237\001\237\001\ -\237\001\237\001\237\001\237\001\237\001\000\000\237\001\237\001\ -\237\001\237\001\000\000\237\001\000\000\237\001\237\001\000\000\ -\237\001\237\001\000\000\000\000\237\001\237\001\237\001\237\001\ -\237\001\237\001\237\001\237\001\237\001\237\001\000\000\237\001\ -\083\000\083\000\083\000\083\000\000\000\083\000\000\000\083\000\ -\000\000\000\000\000\000\083\000\083\000\000\000\083\000\083\000\ -\000\000\000\000\083\000\083\000\000\000\083\000\083\000\083\000\ -\083\000\000\000\083\000\083\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\083\000\000\000\083\000\083\000\000\000\083\000\ -\083\000\083\000\000\000\083\000\000\000\083\000\000\000\000\000\ -\083\000\083\000\000\000\083\000\000\000\083\000\083\000\083\000\ -\083\000\083\000\083\000\083\000\083\000\083\000\083\000\083\000\ -\000\000\083\000\000\000\083\000\083\000\000\000\000\000\000\000\ -\083\000\083\000\083\000\000\000\083\000\083\000\083\000\083\000\ -\085\000\083\000\000\000\083\000\083\000\000\000\000\000\083\000\ -\083\000\083\000\083\000\083\000\083\000\083\000\000\000\083\000\ -\083\000\083\000\083\000\083\000\083\000\083\000\083\000\083\000\ -\083\000\083\000\000\000\000\000\083\000\083\000\083\000\083\000\ -\083\000\083\000\083\000\083\000\083\000\083\000\083\000\000\000\ -\000\000\000\000\000\000\083\000\083\000\083\000\083\000\083\000\ -\000\000\083\000\083\000\083\000\000\000\083\000\083\000\083\000\ -\083\000\083\000\083\000\083\000\083\000\000\000\083\000\083\000\ -\083\000\083\000\000\000\083\000\000\000\083\000\083\000\000\000\ -\083\000\083\000\000\000\000\000\083\000\083\000\083\000\083\000\ -\083\000\083\000\083\000\083\000\083\000\083\000\000\000\083\000\ -\000\000\084\000\084\000\084\000\084\000\000\000\084\000\000\000\ -\084\000\000\000\000\000\000\000\084\000\084\000\000\000\084\000\ -\084\000\000\000\000\000\084\000\084\000\000\000\084\000\084\000\ -\084\000\084\000\000\000\084\000\084\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\084\000\000\000\084\000\084\000\000\000\ -\084\000\084\000\084\000\000\000\084\000\000\000\084\000\000\000\ -\000\000\084\000\084\000\000\000\084\000\000\000\084\000\084\000\ -\084\000\084\000\084\000\084\000\084\000\084\000\084\000\084\000\ -\084\000\000\000\084\000\000\000\084\000\084\000\000\000\000\000\ -\000\000\084\000\084\000\084\000\000\000\084\000\084\000\084\000\ -\084\000\082\000\084\000\000\000\084\000\084\000\000\000\000\000\ -\084\000\084\000\084\000\084\000\084\000\084\000\084\000\000\000\ -\084\000\084\000\084\000\084\000\084\000\084\000\084\000\084\000\ -\084\000\084\000\084\000\000\000\000\000\084\000\084\000\084\000\ -\084\000\084\000\084\000\084\000\084\000\084\000\084\000\084\000\ -\000\000\000\000\000\000\000\000\084\000\084\000\084\000\084\000\ -\084\000\000\000\084\000\084\000\084\000\000\000\084\000\084\000\ -\084\000\084\000\084\000\084\000\084\000\084\000\000\000\084\000\ -\084\000\084\000\084\000\000\000\084\000\000\000\084\000\084\000\ -\000\000\084\000\084\000\000\000\000\000\084\000\084\000\084\000\ -\084\000\084\000\084\000\084\000\084\000\084\000\084\000\000\000\ -\084\000\085\000\085\000\085\000\085\000\000\000\085\000\000\000\ -\085\000\000\000\000\000\000\000\085\000\085\000\000\000\085\000\ -\085\000\000\000\000\000\085\000\085\000\000\000\085\000\085\000\ -\085\000\085\000\000\000\085\000\085\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\085\000\000\000\085\000\085\000\000\000\ -\085\000\085\000\085\000\000\000\085\000\000\000\085\000\000\000\ -\000\000\085\000\085\000\000\000\085\000\000\000\085\000\085\000\ -\085\000\085\000\085\000\085\000\085\000\085\000\085\000\085\000\ -\085\000\000\000\085\000\000\000\085\000\085\000\000\000\000\000\ -\000\000\085\000\085\000\085\000\000\000\085\000\085\000\085\000\ -\085\000\056\000\085\000\000\000\085\000\085\000\000\000\000\000\ -\085\000\085\000\085\000\085\000\085\000\085\000\085\000\000\000\ -\085\000\085\000\085\000\085\000\085\000\085\000\085\000\085\000\ -\085\000\085\000\085\000\000\000\000\000\085\000\085\000\085\000\ -\085\000\085\000\085\000\085\000\085\000\085\000\085\000\085\000\ -\000\000\000\000\000\000\000\000\085\000\085\000\085\000\085\000\ -\085\000\000\000\085\000\085\000\085\000\000\000\085\000\085\000\ -\085\000\085\000\085\000\085\000\085\000\085\000\000\000\085\000\ -\085\000\085\000\085\000\000\000\085\000\000\000\085\000\085\000\ -\000\000\085\000\085\000\000\000\000\000\085\000\085\000\085\000\ -\085\000\085\000\085\000\085\000\085\000\085\000\085\000\000\000\ -\085\000\000\000\082\000\082\000\082\000\082\000\000\000\082\000\ -\000\000\082\000\000\000\000\000\000\000\082\000\082\000\000\000\ -\082\000\082\000\000\000\000\000\082\000\082\000\000\000\082\000\ -\082\000\082\000\082\000\000\000\082\000\082\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\082\000\000\000\082\000\082\000\ -\000\000\082\000\082\000\082\000\000\000\082\000\000\000\082\000\ -\000\000\000\000\082\000\082\000\000\000\082\000\000\000\082\000\ -\082\000\082\000\082\000\082\000\082\000\082\000\082\000\082\000\ -\082\000\082\000\000\000\082\000\000\000\082\000\082\000\000\000\ -\000\000\000\000\082\000\082\000\082\000\000\000\082\000\082\000\ -\082\000\082\000\047\000\082\000\000\000\082\000\082\000\000\000\ -\000\000\082\000\082\000\082\000\082\000\082\000\082\000\082\000\ -\000\000\082\000\082\000\082\000\082\000\082\000\082\000\082\000\ -\082\000\082\000\082\000\082\000\000\000\000\000\082\000\082\000\ -\082\000\082\000\082\000\082\000\082\000\082\000\082\000\082\000\ -\082\000\000\000\000\000\000\000\000\000\082\000\082\000\082\000\ -\082\000\082\000\000\000\082\000\082\000\082\000\000\000\082\000\ -\082\000\082\000\082\000\082\000\082\000\082\000\082\000\000\000\ -\082\000\082\000\082\000\082\000\000\000\082\000\000\000\082\000\ -\082\000\000\000\082\000\082\000\000\000\000\000\082\000\082\000\ -\082\000\082\000\082\000\082\000\082\000\082\000\082\000\082\000\ -\000\000\082\000\056\000\056\000\056\000\056\000\000\000\056\000\ -\000\000\056\000\000\000\000\000\000\000\000\000\056\000\056\000\ -\056\000\000\000\056\000\000\000\000\000\056\000\000\000\056\000\ -\056\000\056\000\056\000\000\000\056\000\056\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\056\000\000\000\056\000\056\000\ -\000\000\056\000\056\000\056\000\000\000\000\000\000\000\056\000\ -\000\000\000\000\000\000\000\000\000\000\056\000\000\000\056\000\ -\056\000\056\000\056\000\056\000\056\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\056\000\000\000\056\000\056\000\000\000\ -\000\000\000\000\000\000\056\000\056\000\000\000\000\000\000\000\ -\056\000\056\000\007\000\056\000\000\000\000\000\000\000\000\000\ -\000\000\056\000\056\000\000\000\056\000\056\000\056\000\056\000\ -\000\000\056\000\056\000\056\000\056\000\056\000\056\000\056\000\ -\056\000\056\000\056\000\000\000\000\000\000\000\000\000\056\000\ +\056\000\056\000\056\000\056\000\000\000\056\000\000\000\056\000\ +\000\000\000\000\000\000\000\000\056\000\056\000\056\000\000\000\ +\056\000\000\000\000\000\056\000\000\000\056\000\056\000\056\000\ +\056\000\000\000\056\000\056\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\056\000\000\000\056\000\056\000\000\000\056\000\ +\056\000\056\000\000\000\000\000\000\000\056\000\000\000\000\000\ +\000\000\000\000\000\000\056\000\000\000\056\000\056\000\056\000\ +\056\000\056\000\056\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\056\000\000\000\056\000\056\000\000\000\000\000\000\000\ +\000\000\056\000\056\000\000\000\000\000\000\000\056\000\056\000\ +\007\000\056\000\000\000\000\000\000\000\000\000\000\000\056\000\ +\056\000\000\000\056\000\056\000\056\000\056\000\000\000\056\000\ \056\000\056\000\056\000\056\000\056\000\056\000\056\000\056\000\ -\056\000\000\000\000\000\000\000\000\000\000\000\056\000\056\000\ -\056\000\000\000\000\000\056\000\056\000\000\000\000\000\056\000\ -\056\000\056\000\056\000\056\000\000\000\000\000\056\000\056\000\ -\000\000\056\000\056\000\056\000\000\000\056\000\000\000\056\000\ -\000\000\000\000\000\000\056\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\056\000\056\000\000\000\056\000\056\000\ -\000\000\056\000\000\000\047\000\047\000\047\000\047\000\000\000\ -\047\000\000\000\047\000\000\000\000\000\000\000\000\000\047\000\ -\047\000\047\000\000\000\047\000\000\000\000\000\047\000\000\000\ -\047\000\047\000\047\000\047\000\000\000\047\000\047\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\047\000\000\000\047\000\ -\047\000\000\000\047\000\047\000\047\000\000\000\000\000\000\000\ -\047\000\000\000\000\000\000\000\000\000\000\000\047\000\000\000\ -\047\000\047\000\047\000\047\000\047\000\047\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\047\000\000\000\047\000\000\000\ -\000\000\000\000\000\000\000\000\047\000\047\000\000\000\000\000\ -\000\000\047\000\047\000\197\001\047\000\000\000\000\000\000\000\ -\000\000\000\000\047\000\047\000\000\000\047\000\047\000\047\000\ -\047\000\000\000\047\000\047\000\047\000\047\000\047\000\047\000\ -\047\000\047\000\047\000\047\000\000\000\000\000\000\000\000\000\ +\056\000\000\000\000\000\000\000\000\000\056\000\056\000\056\000\ +\056\000\056\000\056\000\056\000\056\000\056\000\056\000\000\000\ +\000\000\000\000\000\000\000\000\056\000\056\000\056\000\000\000\ +\000\000\056\000\056\000\000\000\000\000\056\000\056\000\056\000\ +\056\000\056\000\000\000\000\000\056\000\056\000\000\000\056\000\ +\056\000\056\000\000\000\056\000\000\000\056\000\000\000\000\000\ +\000\000\056\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\056\000\056\000\000\000\056\000\056\000\000\000\056\000\ +\047\000\047\000\047\000\047\000\000\000\047\000\000\000\047\000\ +\000\000\000\000\000\000\000\000\047\000\047\000\047\000\000\000\ +\047\000\000\000\000\000\047\000\000\000\047\000\047\000\047\000\ +\047\000\000\000\047\000\047\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\047\000\000\000\047\000\047\000\000\000\047\000\ +\047\000\047\000\000\000\000\000\000\000\047\000\000\000\000\000\ +\000\000\000\000\000\000\047\000\000\000\047\000\047\000\047\000\ +\047\000\047\000\047\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\047\000\000\000\047\000\000\000\000\000\000\000\000\000\ +\000\000\047\000\047\000\000\000\000\000\000\000\047\000\047\000\ +\199\001\047\000\000\000\000\000\000\000\000\000\000\000\047\000\ +\047\000\000\000\047\000\047\000\047\000\047\000\000\000\047\000\ \047\000\047\000\047\000\047\000\047\000\047\000\047\000\047\000\ -\047\000\047\000\000\000\000\000\000\000\000\000\000\000\047\000\ -\047\000\047\000\000\000\000\000\047\000\047\000\000\000\000\000\ -\047\000\047\000\047\000\047\000\047\000\000\000\000\000\047\000\ -\047\000\000\000\047\000\047\000\047\000\000\000\047\000\000\000\ -\047\000\000\000\000\000\000\000\047\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\047\000\047\000\000\000\047\000\ -\047\000\000\000\047\000\007\000\007\000\007\000\007\000\000\000\ -\007\000\000\000\007\000\000\000\000\000\000\000\000\000\007\000\ -\007\000\007\000\000\000\007\000\000\000\000\000\007\000\000\000\ -\007\000\007\000\007\000\007\000\000\000\007\000\007\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\007\000\000\000\007\000\ -\007\000\000\000\007\000\007\000\007\000\000\000\000\000\000\000\ -\007\000\000\000\000\000\000\000\000\000\000\000\007\000\000\000\ -\007\000\007\000\007\000\007\000\007\000\007\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\007\000\000\000\007\000\000\000\ -\000\000\000\000\000\000\000\000\007\000\007\000\000\000\000\000\ -\000\000\007\000\007\000\196\001\007\000\000\000\000\000\000\000\ -\000\000\000\000\007\000\007\000\000\000\007\000\007\000\007\000\ -\007\000\000\000\007\000\007\000\007\000\007\000\007\000\007\000\ +\047\000\000\000\000\000\000\000\000\000\047\000\047\000\047\000\ +\047\000\047\000\047\000\047\000\047\000\047\000\047\000\000\000\ +\000\000\000\000\000\000\000\000\047\000\047\000\047\000\000\000\ +\000\000\047\000\047\000\000\000\000\000\047\000\047\000\047\000\ +\047\000\047\000\000\000\000\000\047\000\047\000\000\000\047\000\ +\047\000\047\000\000\000\047\000\000\000\047\000\000\000\000\000\ +\000\000\047\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\047\000\047\000\000\000\047\000\047\000\000\000\047\000\ +\000\000\007\000\007\000\007\000\007\000\000\000\007\000\000\000\ +\007\000\000\000\000\000\000\000\000\000\007\000\007\000\007\000\ +\000\000\007\000\000\000\000\000\007\000\000\000\007\000\007\000\ +\007\000\007\000\000\000\007\000\007\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\007\000\000\000\007\000\007\000\000\000\ +\007\000\007\000\007\000\000\000\000\000\000\000\007\000\000\000\ +\000\000\000\000\000\000\000\000\007\000\000\000\007\000\007\000\ \007\000\007\000\007\000\007\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\007\000\000\000\007\000\000\000\000\000\000\000\ +\000\000\000\000\007\000\007\000\000\000\000\000\000\000\007\000\ +\007\000\198\001\007\000\000\000\000\000\000\000\000\000\000\000\ +\007\000\007\000\000\000\007\000\007\000\007\000\007\000\000\000\ +\007\000\007\000\007\000\007\000\007\000\007\000\007\000\007\000\ +\007\000\007\000\000\000\000\000\000\000\000\000\007\000\007\000\ \007\000\007\000\007\000\007\000\007\000\007\000\007\000\007\000\ -\007\000\007\000\000\000\000\000\000\000\000\000\000\000\007\000\ -\007\000\007\000\000\000\000\000\007\000\007\000\000\000\000\000\ -\007\000\007\000\007\000\007\000\007\000\000\000\000\000\007\000\ -\007\000\000\000\007\000\007\000\007\000\000\000\007\000\000\000\ -\007\000\000\000\000\000\000\000\007\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\007\000\007\000\000\000\007\000\ -\007\000\000\000\007\000\000\000\197\001\197\001\197\001\197\001\ -\000\000\197\001\000\000\197\001\000\000\000\000\000\000\000\000\ -\197\001\197\001\197\001\000\000\197\001\000\000\000\000\197\001\ -\000\000\197\001\197\001\197\001\197\001\000\000\197\001\197\001\ -\000\000\000\000\000\000\000\000\000\000\000\000\197\001\000\000\ -\197\001\197\001\000\000\197\001\197\001\197\001\000\000\000\000\ -\000\000\197\001\000\000\000\000\000\000\000\000\000\000\197\001\ -\000\000\197\001\197\001\197\001\197\001\197\001\197\001\000\000\ +\000\000\000\000\000\000\000\000\000\000\007\000\007\000\007\000\ +\000\000\000\000\007\000\007\000\000\000\000\000\007\000\007\000\ +\007\000\007\000\007\000\000\000\000\000\007\000\007\000\000\000\ +\007\000\007\000\007\000\000\000\007\000\000\000\007\000\000\000\ +\000\000\000\000\007\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\007\000\007\000\000\000\007\000\007\000\000\000\ +\007\000\199\001\199\001\199\001\199\001\000\000\199\001\000\000\ +\199\001\000\000\000\000\000\000\000\000\199\001\199\001\199\001\ +\000\000\199\001\000\000\000\000\199\001\000\000\199\001\199\001\ +\199\001\199\001\000\000\199\001\199\001\000\000\000\000\000\000\ +\000\000\000\000\000\000\199\001\000\000\199\001\199\001\000\000\ +\199\001\199\001\199\001\000\000\000\000\000\000\199\001\000\000\ +\000\000\000\000\000\000\000\000\199\001\000\000\199\001\199\001\ +\199\001\199\001\199\001\199\001\000\000\000\000\000\000\000\000\ +\000\000\000\000\199\001\000\000\199\001\000\000\000\000\000\000\ +\000\000\000\000\199\001\199\001\000\000\000\000\000\000\199\001\ +\199\001\197\001\199\001\000\000\000\000\000\000\000\000\000\000\ +\199\001\199\001\000\000\199\001\199\001\199\001\199\001\000\000\ +\199\001\199\001\199\001\199\001\199\001\199\001\199\001\199\001\ +\199\001\000\000\000\000\000\000\000\000\000\000\199\001\199\001\ +\199\001\199\001\199\001\199\001\199\001\199\001\199\001\199\001\ +\000\000\000\000\000\000\000\000\000\000\199\001\199\001\199\001\ +\000\000\000\000\199\001\199\001\000\000\000\000\199\001\199\001\ +\199\001\199\001\199\001\000\000\000\000\199\001\199\001\000\000\ +\199\001\199\001\199\001\000\000\199\001\000\000\199\001\000\000\ +\000\000\000\000\199\001\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\199\001\199\001\000\000\199\001\199\001\000\000\ +\199\001\000\000\198\001\198\001\198\001\198\001\000\000\198\001\ +\000\000\198\001\000\000\000\000\000\000\000\000\198\001\198\001\ +\198\001\000\000\198\001\000\000\000\000\198\001\000\000\198\001\ +\198\001\198\001\198\001\000\000\198\001\198\001\000\000\000\000\ +\000\000\000\000\000\000\000\000\198\001\000\000\198\001\198\001\ +\000\000\198\001\198\001\198\001\000\000\000\000\000\000\198\001\ +\000\000\000\000\000\000\000\000\000\000\198\001\000\000\198\001\ +\198\001\198\001\198\001\198\001\198\001\000\000\000\000\000\000\ +\000\000\000\000\000\000\198\001\000\000\198\001\000\000\000\000\ +\000\000\000\000\000\000\198\001\198\001\000\000\000\000\000\000\ +\198\001\198\001\196\001\198\001\000\000\000\000\000\000\000\000\ +\000\000\198\001\198\001\000\000\198\001\198\001\198\001\198\001\ +\000\000\198\001\198\001\198\001\198\001\198\001\198\001\198\001\ +\198\001\198\001\000\000\000\000\000\000\000\000\000\000\198\001\ +\198\001\198\001\198\001\198\001\198\001\198\001\198\001\198\001\ +\198\001\000\000\000\000\000\000\000\000\000\000\198\001\198\001\ +\198\001\000\000\000\000\198\001\198\001\000\000\000\000\198\001\ +\198\001\198\001\198\001\198\001\000\000\000\000\198\001\198\001\ +\000\000\198\001\198\001\198\001\000\000\198\001\000\000\198\001\ +\000\000\000\000\000\000\198\001\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\198\001\198\001\000\000\198\001\198\001\ +\000\000\198\001\197\001\197\001\197\001\197\001\000\000\197\001\ +\000\000\000\000\000\000\000\000\000\000\000\000\197\001\197\001\ +\197\001\000\000\197\001\000\000\000\000\197\001\000\000\197\001\ +\197\001\197\001\197\001\000\000\197\001\197\001\000\000\000\000\ +\000\000\000\000\000\000\000\000\197\001\000\000\197\001\197\001\ +\000\000\197\001\197\001\197\001\000\000\000\000\000\000\197\001\ \000\000\000\000\000\000\000\000\000\000\197\001\000\000\197\001\ -\000\000\000\000\000\000\000\000\000\000\197\001\197\001\000\000\ -\000\000\000\000\197\001\197\001\195\001\197\001\000\000\000\000\ -\000\000\000\000\000\000\197\001\197\001\000\000\197\001\197\001\ -\197\001\197\001\000\000\197\001\197\001\197\001\197\001\197\001\ -\197\001\197\001\197\001\197\001\000\000\000\000\000\000\000\000\ +\197\001\197\001\197\001\197\001\197\001\000\000\000\000\000\000\ +\000\000\000\000\000\000\197\001\000\000\197\001\000\000\000\000\ +\000\000\000\000\000\000\197\001\197\001\000\000\000\000\000\000\ +\197\001\197\001\185\001\197\001\000\000\000\000\000\000\000\000\ +\000\000\197\001\197\001\000\000\197\001\197\001\197\001\197\001\ \000\000\197\001\197\001\197\001\197\001\197\001\197\001\197\001\ -\197\001\197\001\197\001\000\000\000\000\000\000\000\000\000\000\ -\197\001\197\001\197\001\000\000\000\000\197\001\197\001\000\000\ -\000\000\197\001\197\001\197\001\197\001\197\001\000\000\000\000\ -\197\001\197\001\000\000\197\001\197\001\197\001\000\000\197\001\ -\000\000\197\001\000\000\000\000\000\000\197\001\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\197\001\197\001\000\000\ -\197\001\197\001\000\000\197\001\196\001\196\001\196\001\196\001\ -\000\000\196\001\000\000\196\001\000\000\000\000\000\000\000\000\ -\196\001\196\001\196\001\000\000\196\001\000\000\000\000\196\001\ -\000\000\196\001\196\001\196\001\196\001\000\000\196\001\196\001\ -\000\000\000\000\000\000\000\000\000\000\000\000\196\001\000\000\ -\196\001\196\001\000\000\196\001\196\001\196\001\000\000\000\000\ -\000\000\196\001\000\000\000\000\000\000\000\000\000\000\196\001\ -\000\000\196\001\196\001\196\001\196\001\196\001\196\001\000\000\ +\197\001\197\001\000\000\000\000\000\000\000\000\000\000\197\001\ +\197\001\197\001\197\001\197\001\197\001\197\001\197\001\197\001\ +\197\001\000\000\000\000\000\000\000\000\000\000\197\001\197\001\ +\197\001\000\000\000\000\197\001\197\001\000\000\000\000\197\001\ +\197\001\197\001\197\001\197\001\000\000\000\000\197\001\197\001\ +\000\000\197\001\197\001\197\001\000\000\197\001\000\000\197\001\ +\000\000\000\000\000\000\197\001\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\197\001\197\001\000\000\197\001\197\001\ +\000\000\197\001\000\000\000\000\196\001\196\001\196\001\000\000\ +\196\001\000\000\000\000\000\000\000\000\000\000\000\000\196\001\ +\196\001\196\001\000\000\196\001\000\000\000\000\196\001\000\000\ +\196\001\196\001\196\001\196\001\000\000\196\001\196\001\000\000\ \000\000\000\000\000\000\000\000\000\000\196\001\000\000\196\001\ -\000\000\000\000\000\000\000\000\000\000\196\001\196\001\000\000\ -\000\000\000\000\196\001\196\001\194\001\196\001\000\000\000\000\ -\000\000\000\000\000\000\196\001\196\001\000\000\196\001\196\001\ -\196\001\196\001\000\000\196\001\196\001\196\001\196\001\196\001\ -\196\001\196\001\196\001\196\001\000\000\000\000\000\000\000\000\ -\000\000\196\001\196\001\196\001\196\001\196\001\196\001\196\001\ -\196\001\196\001\196\001\000\000\000\000\000\000\000\000\000\000\ -\196\001\196\001\196\001\000\000\000\000\196\001\196\001\000\000\ -\000\000\196\001\196\001\196\001\196\001\196\001\000\000\000\000\ -\196\001\196\001\000\000\196\001\196\001\196\001\000\000\196\001\ -\000\000\196\001\000\000\000\000\000\000\196\001\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\196\001\196\001\000\000\ -\196\001\196\001\000\000\196\001\000\000\195\001\195\001\195\001\ -\195\001\000\000\195\001\000\000\000\000\000\000\000\000\000\000\ -\000\000\195\001\195\001\195\001\000\000\195\001\000\000\000\000\ -\195\001\000\000\195\001\195\001\195\001\195\001\000\000\195\001\ -\195\001\000\000\000\000\000\000\000\000\000\000\000\000\195\001\ -\000\000\195\001\195\001\000\000\195\001\195\001\195\001\000\000\ -\000\000\000\000\195\001\000\000\000\000\000\000\000\000\000\000\ -\195\001\000\000\195\001\195\001\195\001\195\001\195\001\195\001\ +\196\001\000\000\196\001\196\001\196\001\000\000\000\000\000\000\ +\196\001\000\000\000\000\000\000\000\000\000\000\196\001\000\000\ +\196\001\196\001\196\001\196\001\196\001\196\001\000\000\000\000\ +\000\000\000\000\000\000\000\000\196\001\000\000\196\001\000\000\ +\000\000\000\000\000\000\000\000\196\001\196\001\000\000\000\000\ +\000\000\196\001\196\001\195\001\196\001\000\000\000\000\000\000\ +\000\000\000\000\196\001\196\001\000\000\196\001\196\001\196\001\ +\196\001\000\000\196\001\196\001\196\001\196\001\196\001\196\001\ +\196\001\196\001\000\000\000\000\000\000\000\000\000\000\000\000\ +\196\001\196\001\196\001\196\001\196\001\196\001\196\001\196\001\ +\196\001\196\001\000\000\000\000\000\000\000\000\000\000\196\001\ +\196\001\196\001\000\000\000\000\196\001\196\001\000\000\000\000\ +\196\001\196\001\196\001\196\001\196\001\000\000\000\000\196\001\ +\196\001\000\000\196\001\196\001\196\001\000\000\196\001\000\000\ +\196\001\000\000\000\000\000\000\196\001\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\196\001\196\001\000\000\196\001\ +\196\001\000\000\196\001\000\000\185\001\185\001\185\001\000\000\ +\185\001\000\000\000\000\000\000\000\000\000\000\000\000\185\001\ +\185\001\185\001\000\000\185\001\000\000\000\000\185\001\000\000\ +\000\000\185\001\185\001\185\001\000\000\185\001\185\001\000\000\ +\000\000\000\000\000\000\000\000\000\000\185\001\000\000\185\001\ +\185\001\000\000\185\001\185\001\185\001\000\000\000\000\000\000\ +\185\001\000\000\000\000\000\000\000\000\000\000\185\001\000\000\ +\185\001\185\001\185\001\185\001\185\001\185\001\000\000\000\000\ +\000\000\000\000\000\000\000\000\185\001\000\000\185\001\000\000\ +\000\000\000\000\000\000\000\000\185\001\185\001\000\000\000\000\ +\000\000\185\001\185\001\177\001\185\001\000\000\000\000\000\000\ +\000\000\000\000\185\001\185\001\000\000\185\001\185\001\185\001\ +\185\001\000\000\185\001\185\001\185\001\185\001\185\001\185\001\ +\185\001\185\001\000\000\000\000\000\000\000\000\000\000\000\000\ +\185\001\185\001\185\001\185\001\185\001\185\001\185\001\185\001\ +\185\001\185\001\000\000\000\000\000\000\000\000\000\000\185\001\ +\185\001\185\001\000\000\000\000\185\001\185\001\000\000\000\000\ +\185\001\185\001\185\001\185\001\185\001\000\000\000\000\185\001\ +\185\001\000\000\185\001\185\001\185\001\000\000\185\001\000\000\ +\185\001\000\000\000\000\000\000\185\001\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\185\001\185\001\000\000\185\001\ +\185\001\000\000\185\001\000\000\000\000\195\001\195\001\195\001\ +\000\000\195\001\000\000\000\000\000\000\000\000\000\000\000\000\ +\195\001\195\001\195\001\000\000\195\001\000\000\000\000\195\001\ +\000\000\000\000\195\001\195\001\195\001\000\000\195\001\195\001\ \000\000\000\000\000\000\000\000\000\000\000\000\195\001\000\000\ -\195\001\000\000\000\000\000\000\000\000\000\000\195\001\195\001\ -\000\000\000\000\000\000\195\001\195\001\183\001\195\001\000\000\ -\000\000\000\000\000\000\000\000\195\001\195\001\000\000\195\001\ -\195\001\195\001\195\001\000\000\195\001\195\001\195\001\195\001\ -\195\001\195\001\195\001\195\001\195\001\000\000\000\000\000\000\ -\000\000\000\000\195\001\195\001\195\001\195\001\195\001\195\001\ -\195\001\195\001\195\001\195\001\000\000\000\000\000\000\000\000\ -\000\000\195\001\195\001\195\001\000\000\000\000\195\001\195\001\ -\000\000\000\000\195\001\195\001\195\001\195\001\195\001\000\000\ -\000\000\195\001\195\001\000\000\195\001\195\001\195\001\000\000\ -\195\001\000\000\195\001\000\000\000\000\000\000\195\001\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\195\001\195\001\ -\000\000\195\001\195\001\000\000\195\001\000\000\194\001\194\001\ -\194\001\000\000\194\001\000\000\000\000\000\000\000\000\000\000\ -\000\000\194\001\194\001\194\001\000\000\194\001\000\000\000\000\ -\194\001\000\000\194\001\194\001\194\001\194\001\000\000\194\001\ -\194\001\000\000\000\000\000\000\000\000\000\000\000\000\194\001\ -\000\000\194\001\194\001\000\000\194\001\194\001\194\001\000\000\ -\000\000\000\000\194\001\000\000\000\000\000\000\000\000\000\000\ -\194\001\000\000\194\001\194\001\194\001\194\001\194\001\194\001\ -\000\000\000\000\000\000\000\000\000\000\000\000\194\001\000\000\ -\194\001\000\000\000\000\000\000\000\000\000\000\194\001\194\001\ -\000\000\000\000\000\000\194\001\194\001\193\001\194\001\000\000\ -\000\000\000\000\000\000\000\000\194\001\194\001\000\000\194\001\ -\194\001\194\001\194\001\000\000\194\001\194\001\194\001\194\001\ -\194\001\194\001\194\001\194\001\000\000\000\000\000\000\000\000\ -\000\000\000\000\194\001\194\001\194\001\194\001\194\001\194\001\ -\194\001\194\001\194\001\194\001\000\000\000\000\000\000\000\000\ -\000\000\194\001\194\001\194\001\000\000\000\000\194\001\194\001\ -\000\000\000\000\194\001\194\001\194\001\194\001\194\001\000\000\ -\000\000\194\001\194\001\000\000\194\001\194\001\194\001\000\000\ -\194\001\000\000\194\001\000\000\000\000\000\000\194\001\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\194\001\194\001\ -\000\000\194\001\194\001\000\000\194\001\000\000\000\000\183\001\ -\183\001\183\001\000\000\183\001\000\000\000\000\000\000\000\000\ -\000\000\000\000\183\001\183\001\183\001\000\000\183\001\000\000\ -\000\000\183\001\000\000\000\000\183\001\183\001\183\001\000\000\ -\183\001\183\001\000\000\000\000\000\000\000\000\000\000\000\000\ -\183\001\000\000\183\001\183\001\000\000\183\001\183\001\183\001\ -\000\000\000\000\000\000\183\001\000\000\000\000\000\000\000\000\ -\000\000\183\001\000\000\183\001\183\001\183\001\183\001\183\001\ -\183\001\000\000\000\000\000\000\000\000\000\000\000\000\183\001\ -\000\000\183\001\000\000\000\000\000\000\000\000\000\000\183\001\ -\183\001\000\000\000\000\000\000\183\001\183\001\175\001\183\001\ -\000\000\000\000\000\000\000\000\000\000\183\001\183\001\000\000\ -\183\001\183\001\183\001\183\001\000\000\183\001\183\001\183\001\ -\183\001\183\001\183\001\183\001\183\001\000\000\000\000\000\000\ -\000\000\000\000\000\000\183\001\183\001\183\001\183\001\183\001\ -\183\001\183\001\183\001\183\001\183\001\000\000\000\000\000\000\ -\000\000\000\000\183\001\183\001\183\001\000\000\000\000\183\001\ -\183\001\000\000\000\000\183\001\183\001\183\001\183\001\183\001\ -\000\000\000\000\183\001\183\001\000\000\183\001\183\001\183\001\ -\000\000\183\001\000\000\183\001\000\000\000\000\000\000\183\001\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\183\001\ -\183\001\000\000\183\001\183\001\000\000\183\001\000\000\193\001\ -\193\001\193\001\000\000\193\001\000\000\000\000\000\000\000\000\ -\000\000\000\000\193\001\193\001\193\001\000\000\193\001\000\000\ -\000\000\193\001\000\000\000\000\193\001\193\001\193\001\000\000\ -\193\001\193\001\000\000\000\000\000\000\000\000\000\000\000\000\ -\193\001\000\000\193\001\193\001\000\000\193\001\193\001\193\001\ -\000\000\000\000\000\000\193\001\000\000\000\000\000\000\000\000\ -\000\000\193\001\000\000\193\001\193\001\193\001\193\001\193\001\ -\193\001\000\000\000\000\000\000\000\000\000\000\000\000\193\001\ -\000\000\193\001\000\000\000\000\000\000\000\000\000\000\193\001\ -\193\001\000\000\000\000\000\000\193\001\193\001\170\001\193\001\ -\000\000\000\000\000\000\000\000\000\000\193\001\193\001\000\000\ -\193\001\193\001\193\001\193\001\000\000\193\001\193\001\193\001\ -\193\001\193\001\193\001\193\001\193\001\000\000\000\000\000\000\ -\000\000\000\000\000\000\193\001\193\001\193\001\193\001\193\001\ -\193\001\193\001\193\001\193\001\193\001\000\000\000\000\000\000\ -\000\000\000\000\193\001\193\001\193\001\000\000\000\000\193\001\ -\193\001\000\000\000\000\193\001\193\001\193\001\193\001\193\001\ -\000\000\000\000\193\001\193\001\000\000\193\001\193\001\193\001\ -\000\000\193\001\000\000\193\001\000\000\000\000\000\000\193\001\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\193\001\ -\193\001\000\000\193\001\193\001\000\000\193\001\000\000\000\000\ -\175\001\175\001\175\001\000\000\175\001\000\000\000\000\000\000\ -\000\000\000\000\000\000\175\001\175\001\175\001\000\000\175\001\ -\000\000\000\000\175\001\000\000\000\000\175\001\175\001\175\001\ -\000\000\175\001\175\001\000\000\000\000\000\000\000\000\000\000\ -\000\000\175\001\000\000\175\001\175\001\000\000\175\001\175\001\ -\175\001\000\000\000\000\000\000\175\001\000\000\000\000\000\000\ -\000\000\000\000\175\001\000\000\175\001\175\001\175\001\175\001\ -\175\001\175\001\000\000\000\000\000\000\000\000\000\000\000\000\ -\175\001\000\000\175\001\000\000\000\000\000\000\000\000\000\000\ -\175\001\175\001\000\000\000\000\000\000\175\001\175\001\174\001\ -\175\001\000\000\000\000\000\000\000\000\000\000\000\000\175\001\ -\000\000\175\001\175\001\175\001\175\001\000\000\175\001\175\001\ -\175\001\175\001\175\001\175\001\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\175\001\175\001\175\001\175\001\ -\175\001\175\001\175\001\175\001\175\001\175\001\000\000\000\000\ -\000\000\000\000\000\000\175\001\175\001\175\001\000\000\000\000\ -\175\001\175\001\000\000\000\000\175\001\175\001\175\001\175\001\ -\175\001\000\000\000\000\175\001\175\001\000\000\175\001\175\001\ -\175\001\000\000\175\001\000\000\175\001\000\000\000\000\000\000\ -\175\001\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\175\001\175\001\000\000\175\001\175\001\000\000\175\001\000\000\ -\170\001\170\001\170\001\000\000\170\001\000\000\000\000\000\000\ -\000\000\000\000\000\000\170\001\170\001\170\001\000\000\170\001\ -\000\000\000\000\170\001\000\000\000\000\170\001\170\001\170\001\ -\000\000\170\001\170\001\000\000\000\000\000\000\000\000\000\000\ -\000\000\170\001\000\000\170\001\170\001\000\000\170\001\170\001\ -\170\001\000\000\000\000\000\000\170\001\000\000\000\000\000\000\ -\000\000\000\000\170\001\000\000\170\001\170\001\170\001\170\001\ +\195\001\195\001\000\000\195\001\195\001\195\001\000\000\000\000\ +\000\000\195\001\000\000\000\000\000\000\000\000\000\000\195\001\ +\000\000\195\001\195\001\195\001\195\001\195\001\195\001\000\000\ +\000\000\000\000\000\000\000\000\000\000\195\001\000\000\195\001\ +\000\000\000\000\000\000\000\000\000\000\195\001\195\001\000\000\ +\000\000\000\000\195\001\195\001\172\001\195\001\000\000\000\000\ +\000\000\000\000\000\000\195\001\195\001\000\000\195\001\195\001\ +\195\001\195\001\000\000\195\001\195\001\195\001\195\001\195\001\ +\195\001\195\001\195\001\000\000\000\000\000\000\000\000\000\000\ +\000\000\195\001\195\001\195\001\195\001\195\001\195\001\195\001\ +\195\001\195\001\195\001\000\000\000\000\000\000\000\000\000\000\ +\195\001\195\001\195\001\000\000\000\000\195\001\195\001\000\000\ +\000\000\195\001\195\001\195\001\195\001\195\001\000\000\000\000\ +\195\001\195\001\000\000\195\001\195\001\195\001\000\000\195\001\ +\000\000\195\001\000\000\000\000\000\000\195\001\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\195\001\195\001\000\000\ +\195\001\195\001\000\000\195\001\000\000\177\001\177\001\177\001\ +\000\000\177\001\000\000\000\000\000\000\000\000\000\000\000\000\ +\177\001\177\001\177\001\000\000\177\001\000\000\000\000\177\001\ +\000\000\000\000\177\001\177\001\177\001\000\000\177\001\177\001\ +\000\000\000\000\000\000\000\000\000\000\000\000\177\001\000\000\ +\177\001\177\001\000\000\177\001\177\001\177\001\000\000\000\000\ +\000\000\177\001\000\000\000\000\000\000\000\000\000\000\177\001\ +\000\000\177\001\177\001\177\001\177\001\177\001\177\001\000\000\ +\000\000\000\000\000\000\000\000\000\000\177\001\000\000\177\001\ +\000\000\000\000\000\000\000\000\000\000\177\001\177\001\000\000\ +\000\000\000\000\177\001\177\001\176\001\177\001\000\000\000\000\ +\000\000\000\000\000\000\000\000\177\001\000\000\177\001\177\001\ +\177\001\177\001\000\000\177\001\177\001\177\001\177\001\177\001\ +\177\001\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\177\001\177\001\177\001\177\001\177\001\177\001\177\001\ +\177\001\177\001\177\001\000\000\000\000\000\000\000\000\000\000\ +\177\001\177\001\177\001\000\000\000\000\177\001\177\001\000\000\ +\000\000\177\001\177\001\177\001\177\001\177\001\000\000\000\000\ +\177\001\177\001\000\000\177\001\177\001\177\001\000\000\177\001\ +\000\000\177\001\000\000\000\000\000\000\177\001\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\177\001\177\001\000\000\ +\177\001\177\001\000\000\177\001\000\000\000\000\172\001\172\001\ +\172\001\000\000\172\001\000\000\000\000\000\000\000\000\000\000\ +\000\000\172\001\172\001\172\001\000\000\172\001\000\000\000\000\ +\172\001\000\000\000\000\172\001\172\001\172\001\000\000\172\001\ +\172\001\000\000\000\000\000\000\000\000\000\000\000\000\172\001\ +\000\000\172\001\172\001\000\000\172\001\172\001\172\001\000\000\ +\000\000\000\000\172\001\000\000\000\000\000\000\000\000\000\000\ +\172\001\000\000\172\001\172\001\172\001\172\001\172\001\172\001\ +\000\000\000\000\000\000\000\000\000\000\000\000\172\001\000\000\ +\172\001\000\000\000\000\000\000\000\000\000\000\172\001\172\001\ +\000\000\000\000\000\000\172\001\172\001\170\001\172\001\000\000\ +\000\000\000\000\000\000\000\000\000\000\172\001\000\000\172\001\ +\172\001\172\001\172\001\000\000\172\001\172\001\172\001\172\001\ +\172\001\172\001\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\172\001\172\001\172\001\172\001\172\001\172\001\ +\172\001\172\001\172\001\172\001\000\000\000\000\000\000\000\000\ +\000\000\172\001\172\001\172\001\000\000\000\000\172\001\172\001\ +\000\000\000\000\172\001\172\001\172\001\172\001\172\001\000\000\ +\000\000\172\001\172\001\000\000\172\001\172\001\172\001\000\000\ +\172\001\000\000\172\001\000\000\000\000\000\000\172\001\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\172\001\172\001\ +\000\000\172\001\172\001\000\000\172\001\000\000\176\001\176\001\ +\176\001\000\000\176\001\000\000\000\000\000\000\000\000\000\000\ +\000\000\176\001\176\001\176\001\000\000\176\001\000\000\000\000\ +\176\001\000\000\000\000\176\001\176\001\176\001\000\000\176\001\ +\176\001\000\000\000\000\000\000\000\000\000\000\000\000\176\001\ +\000\000\176\001\176\001\000\000\176\001\176\001\176\001\000\000\ +\000\000\000\000\176\001\000\000\000\000\000\000\000\000\000\000\ +\176\001\000\000\176\001\176\001\176\001\176\001\176\001\176\001\ +\000\000\000\000\000\000\000\000\000\000\000\000\176\001\000\000\ +\176\001\000\000\000\000\000\000\000\000\000\000\176\001\176\001\ +\000\000\000\000\000\000\176\001\176\001\175\001\176\001\000\000\ +\000\000\000\000\000\000\000\000\000\000\176\001\000\000\176\001\ +\176\001\176\001\176\001\000\000\176\001\176\001\176\001\176\001\ +\176\001\176\001\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\176\001\176\001\176\001\176\001\176\001\ +\176\001\176\001\176\001\176\001\000\000\000\000\000\000\000\000\ +\000\000\176\001\176\001\176\001\000\000\000\000\176\001\176\001\ +\000\000\000\000\176\001\176\001\176\001\176\001\176\001\000\000\ +\000\000\176\001\176\001\000\000\176\001\176\001\176\001\000\000\ +\176\001\000\000\176\001\000\000\000\000\000\000\176\001\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\176\001\176\001\ +\000\000\176\001\176\001\000\000\176\001\000\000\000\000\170\001\ +\170\001\170\001\000\000\170\001\000\000\000\000\000\000\000\000\ +\000\000\000\000\170\001\170\001\170\001\000\000\170\001\000\000\ +\000\000\170\001\000\000\000\000\170\001\170\001\170\001\000\000\ \170\001\170\001\000\000\000\000\000\000\000\000\000\000\000\000\ -\170\001\000\000\170\001\000\000\000\000\000\000\000\000\000\000\ -\170\001\170\001\000\000\000\000\000\000\170\001\170\001\168\001\ +\170\001\000\000\170\001\170\001\000\000\170\001\170\001\170\001\ +\000\000\000\000\000\000\170\001\000\000\000\000\000\000\000\000\ +\000\000\170\001\000\000\170\001\170\001\170\001\170\001\170\001\ \170\001\000\000\000\000\000\000\000\000\000\000\000\000\170\001\ -\000\000\170\001\170\001\170\001\170\001\000\000\170\001\170\001\ -\170\001\170\001\170\001\170\001\000\000\000\000\000\000\000\000\ +\000\000\170\001\000\000\000\000\000\000\000\000\000\000\170\001\ +\170\001\000\000\000\000\000\000\170\001\170\001\174\001\170\001\ +\000\000\000\000\000\000\000\000\000\000\000\000\170\001\000\000\ +\170\001\170\001\170\001\170\001\000\000\170\001\170\001\170\001\ +\170\001\170\001\000\000\000\000\000\000\000\000\000\000\000\000\ \000\000\000\000\000\000\000\000\170\001\170\001\170\001\170\001\ -\170\001\170\001\170\001\170\001\170\001\170\001\000\000\000\000\ -\000\000\000\000\000\000\170\001\170\001\170\001\000\000\000\000\ -\170\001\170\001\000\000\000\000\170\001\170\001\170\001\170\001\ -\170\001\000\000\000\000\170\001\170\001\000\000\170\001\170\001\ -\170\001\000\000\170\001\000\000\170\001\000\000\000\000\000\000\ -\170\001\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\170\001\170\001\000\000\170\001\170\001\000\000\170\001\000\000\ -\000\000\174\001\174\001\174\001\000\000\174\001\000\000\000\000\ -\000\000\000\000\000\000\000\000\174\001\174\001\174\001\000\000\ -\174\001\000\000\000\000\174\001\000\000\000\000\174\001\174\001\ -\174\001\000\000\174\001\174\001\000\000\000\000\000\000\000\000\ -\000\000\000\000\174\001\000\000\174\001\174\001\000\000\174\001\ -\174\001\174\001\000\000\000\000\000\000\174\001\000\000\000\000\ -\000\000\000\000\000\000\174\001\000\000\174\001\174\001\174\001\ -\174\001\174\001\174\001\000\000\000\000\000\000\000\000\000\000\ -\000\000\174\001\000\000\174\001\000\000\000\000\000\000\000\000\ -\000\000\174\001\174\001\000\000\000\000\000\000\174\001\174\001\ -\173\001\174\001\000\000\000\000\000\000\000\000\000\000\000\000\ -\174\001\000\000\174\001\174\001\174\001\174\001\000\000\174\001\ -\174\001\174\001\174\001\174\001\174\001\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\174\001\174\001\ -\174\001\174\001\174\001\174\001\174\001\174\001\174\001\000\000\ -\000\000\000\000\000\000\000\000\174\001\174\001\174\001\000\000\ -\000\000\174\001\174\001\000\000\000\000\174\001\174\001\174\001\ -\174\001\174\001\000\000\000\000\174\001\174\001\000\000\174\001\ -\174\001\174\001\000\000\174\001\000\000\174\001\000\000\000\000\ -\000\000\174\001\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\174\001\174\001\000\000\174\001\174\001\000\000\174\001\ -\000\000\168\001\168\001\168\001\000\000\168\001\000\000\000\000\ -\000\000\000\000\000\000\000\000\168\001\168\001\168\001\000\000\ -\168\001\000\000\000\000\168\001\000\000\000\000\168\001\168\001\ -\168\001\000\000\168\001\168\001\000\000\000\000\000\000\000\000\ -\000\000\000\000\168\001\000\000\168\001\168\001\000\000\168\001\ -\168\001\168\001\000\000\000\000\000\000\168\001\000\000\000\000\ -\000\000\000\000\000\000\168\001\000\000\168\001\168\001\168\001\ -\168\001\168\001\168\001\000\000\000\000\000\000\000\000\000\000\ -\000\000\168\001\000\000\168\001\000\000\000\000\000\000\000\000\ -\000\000\168\001\168\001\000\000\000\000\000\000\168\001\168\001\ -\172\001\168\001\000\000\000\000\000\000\000\000\000\000\000\000\ -\168\001\000\000\168\001\168\001\168\001\168\001\000\000\168\001\ -\168\001\168\001\168\001\168\001\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\168\001\168\001\ -\168\001\168\001\168\001\168\001\168\001\168\001\168\001\000\000\ -\000\000\000\000\000\000\000\000\168\001\168\001\168\001\000\000\ -\000\000\168\001\168\001\000\000\000\000\168\001\168\001\168\001\ -\168\001\168\001\000\000\000\000\168\001\168\001\000\000\168\001\ -\168\001\168\001\000\000\168\001\000\000\168\001\000\000\000\000\ -\000\000\168\001\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\168\001\168\001\000\000\168\001\168\001\000\000\168\001\ -\000\000\000\000\173\001\173\001\173\001\000\000\173\001\000\000\ -\000\000\000\000\000\000\000\000\000\000\173\001\173\001\173\001\ -\000\000\173\001\000\000\000\000\173\001\000\000\000\000\173\001\ -\173\001\173\001\000\000\173\001\173\001\000\000\000\000\000\000\ -\000\000\000\000\000\000\173\001\000\000\173\001\173\001\000\000\ -\173\001\173\001\173\001\000\000\000\000\000\000\173\001\000\000\ -\000\000\000\000\000\000\000\000\173\001\000\000\173\001\173\001\ -\173\001\173\001\173\001\173\001\000\000\000\000\000\000\000\000\ -\000\000\000\000\173\001\000\000\173\001\000\000\000\000\000\000\ -\000\000\000\000\173\001\173\001\000\000\000\000\000\000\173\001\ -\173\001\171\001\173\001\000\000\000\000\000\000\000\000\000\000\ -\000\000\173\001\000\000\173\001\173\001\173\001\173\001\000\000\ -\173\001\173\001\173\001\173\001\173\001\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\173\001\ -\173\001\173\001\173\001\173\001\173\001\173\001\173\001\173\001\ +\170\001\170\001\170\001\170\001\170\001\000\000\000\000\000\000\ +\000\000\000\000\170\001\170\001\170\001\000\000\000\000\170\001\ +\170\001\000\000\000\000\170\001\170\001\170\001\170\001\170\001\ +\000\000\000\000\170\001\170\001\000\000\170\001\170\001\170\001\ +\000\000\170\001\000\000\170\001\000\000\000\000\000\000\170\001\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\170\001\ +\170\001\000\000\170\001\170\001\000\000\170\001\000\000\175\001\ +\175\001\175\001\000\000\175\001\000\000\000\000\000\000\000\000\ +\000\000\000\000\175\001\175\001\175\001\000\000\175\001\000\000\ +\000\000\175\001\000\000\000\000\175\001\175\001\175\001\000\000\ +\175\001\175\001\000\000\000\000\000\000\000\000\000\000\000\000\ +\175\001\000\000\175\001\175\001\000\000\175\001\175\001\175\001\ +\000\000\000\000\000\000\175\001\000\000\000\000\000\000\000\000\ +\000\000\175\001\000\000\175\001\175\001\175\001\175\001\175\001\ +\175\001\000\000\000\000\000\000\000\000\000\000\000\000\175\001\ +\000\000\175\001\000\000\000\000\000\000\000\000\000\000\175\001\ +\175\001\000\000\000\000\000\000\175\001\175\001\173\001\175\001\ +\000\000\000\000\000\000\000\000\000\000\000\000\175\001\000\000\ +\175\001\175\001\175\001\175\001\000\000\175\001\175\001\175\001\ +\175\001\175\001\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\175\001\175\001\175\001\175\001\ +\175\001\175\001\175\001\175\001\175\001\000\000\000\000\000\000\ +\000\000\000\000\175\001\175\001\175\001\000\000\000\000\175\001\ +\175\001\000\000\000\000\175\001\175\001\175\001\175\001\175\001\ +\000\000\000\000\175\001\175\001\000\000\175\001\175\001\175\001\ +\000\000\175\001\000\000\175\001\000\000\000\000\000\000\175\001\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\175\001\ +\175\001\000\000\175\001\175\001\000\000\175\001\000\000\000\000\ +\174\001\174\001\174\001\000\000\174\001\000\000\000\000\000\000\ +\000\000\000\000\000\000\174\001\174\001\174\001\000\000\174\001\ +\000\000\000\000\174\001\000\000\000\000\174\001\174\001\174\001\ +\000\000\174\001\174\001\000\000\000\000\000\000\000\000\000\000\ +\000\000\174\001\000\000\174\001\174\001\000\000\000\000\174\001\ +\174\001\000\000\000\000\000\000\174\001\000\000\000\000\000\000\ +\000\000\000\000\174\001\000\000\174\001\174\001\174\001\174\001\ +\174\001\174\001\000\000\000\000\000\000\000\000\000\000\000\000\ +\174\001\000\000\174\001\000\000\000\000\000\000\000\000\000\000\ +\174\001\174\001\000\000\000\000\000\000\174\001\174\001\223\001\ +\174\001\000\000\000\000\000\000\000\000\000\000\000\000\174\001\ +\000\000\174\001\174\001\174\001\174\001\000\000\174\001\174\001\ +\174\001\174\001\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\174\001\174\001\174\001\ +\174\001\174\001\174\001\174\001\174\001\174\001\000\000\000\000\ +\000\000\000\000\000\000\174\001\174\001\174\001\000\000\000\000\ +\174\001\174\001\000\000\000\000\174\001\174\001\174\001\174\001\ +\174\001\000\000\000\000\174\001\174\001\000\000\174\001\174\001\ +\174\001\000\000\174\001\000\000\174\001\000\000\000\000\000\000\ +\174\001\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\174\001\174\001\000\000\174\001\174\001\000\000\174\001\000\000\ +\173\001\173\001\173\001\000\000\173\001\000\000\000\000\000\000\ +\000\000\000\000\000\000\173\001\173\001\173\001\000\000\173\001\ +\000\000\000\000\173\001\000\000\000\000\173\001\173\001\173\001\ +\000\000\173\001\173\001\000\000\000\000\000\000\000\000\000\000\ +\000\000\173\001\000\000\173\001\173\001\000\000\000\000\173\001\ +\173\001\000\000\000\000\000\000\173\001\000\000\000\000\000\000\ +\000\000\000\000\173\001\000\000\173\001\173\001\173\001\173\001\ +\173\001\173\001\000\000\000\000\000\000\000\000\000\000\000\000\ +\173\001\000\000\173\001\000\000\000\000\000\000\000\000\000\000\ +\173\001\173\001\000\000\000\000\000\000\173\001\173\001\171\001\ +\173\001\000\000\000\000\000\000\000\000\000\000\000\000\173\001\ +\000\000\173\001\173\001\173\001\173\001\000\000\173\001\173\001\ +\173\001\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ \000\000\000\000\000\000\000\000\000\000\173\001\173\001\173\001\ -\000\000\000\000\173\001\173\001\000\000\000\000\173\001\173\001\ -\173\001\173\001\173\001\000\000\000\000\173\001\173\001\000\000\ -\173\001\173\001\173\001\000\000\173\001\000\000\173\001\000\000\ -\000\000\000\000\173\001\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\173\001\173\001\000\000\173\001\173\001\000\000\ -\173\001\000\000\172\001\172\001\172\001\000\000\172\001\000\000\ -\000\000\000\000\000\000\000\000\000\000\172\001\172\001\172\001\ -\000\000\172\001\000\000\000\000\172\001\000\000\000\000\172\001\ -\172\001\172\001\000\000\172\001\172\001\000\000\000\000\000\000\ -\000\000\000\000\000\000\172\001\000\000\172\001\172\001\000\000\ -\000\000\172\001\172\001\000\000\000\000\000\000\172\001\000\000\ -\000\000\000\000\000\000\000\000\172\001\000\000\172\001\172\001\ -\172\001\172\001\172\001\172\001\000\000\000\000\000\000\000\000\ -\000\000\000\000\172\001\000\000\172\001\000\000\000\000\000\000\ -\000\000\000\000\172\001\172\001\000\000\000\000\000\000\172\001\ -\172\001\221\001\172\001\000\000\000\000\000\000\000\000\000\000\ -\000\000\172\001\000\000\172\001\172\001\172\001\172\001\000\000\ -\172\001\172\001\172\001\172\001\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\172\001\ -\172\001\172\001\172\001\172\001\172\001\172\001\172\001\172\001\ -\000\000\000\000\000\000\000\000\000\000\172\001\172\001\172\001\ -\000\000\000\000\172\001\172\001\000\000\000\000\172\001\172\001\ -\172\001\172\001\172\001\000\000\000\000\172\001\172\001\000\000\ -\172\001\172\001\172\001\000\000\172\001\000\000\172\001\000\000\ -\000\000\000\000\172\001\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\172\001\172\001\000\000\172\001\172\001\000\000\ -\172\001\000\000\000\000\171\001\171\001\171\001\000\000\171\001\ +\173\001\173\001\173\001\173\001\173\001\173\001\000\000\000\000\ +\000\000\000\000\000\000\173\001\173\001\173\001\000\000\000\000\ +\173\001\173\001\000\000\000\000\173\001\173\001\173\001\173\001\ +\173\001\000\000\000\000\173\001\173\001\000\000\173\001\173\001\ +\173\001\000\000\173\001\000\000\173\001\000\000\000\000\000\000\ +\173\001\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\173\001\173\001\000\000\173\001\173\001\000\000\173\001\000\000\ +\000\000\223\001\223\001\223\001\000\000\223\001\000\000\000\000\ +\000\000\000\000\000\000\000\000\223\001\223\001\223\001\000\000\ +\223\001\000\000\000\000\223\001\000\000\000\000\000\000\223\001\ +\223\001\000\000\223\001\223\001\000\000\000\000\000\000\000\000\ +\000\000\000\000\223\001\000\000\223\001\223\001\000\000\000\000\ +\223\001\223\001\000\000\000\000\000\000\223\001\000\000\000\000\ +\000\000\000\000\000\000\223\001\000\000\223\001\223\001\223\001\ +\223\001\223\001\223\001\000\000\000\000\000\000\000\000\000\000\ +\000\000\223\001\000\000\223\001\000\000\000\000\000\000\000\000\ +\000\000\223\001\223\001\000\000\000\000\000\000\223\001\223\001\ +\104\000\223\001\000\000\000\000\000\000\000\000\000\000\000\000\ +\223\001\000\000\223\001\223\001\223\001\223\001\000\000\223\001\ +\223\001\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\223\001\223\001\ +\223\001\223\001\223\001\223\001\223\001\223\001\223\001\000\000\ +\000\000\000\000\000\000\000\000\223\001\223\001\223\001\000\000\ +\000\000\223\001\223\001\000\000\000\000\223\001\223\001\223\001\ +\223\001\223\001\000\000\000\000\223\001\223\001\000\000\223\001\ +\223\001\223\001\000\000\223\001\000\000\223\001\000\000\000\000\ +\000\000\223\001\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\223\001\223\001\000\000\223\001\223\001\000\000\223\001\ +\000\000\171\001\171\001\171\001\000\000\171\001\000\000\000\000\ +\000\000\000\000\000\000\000\000\171\001\171\001\171\001\000\000\ +\171\001\000\000\000\000\171\001\000\000\000\000\000\000\171\001\ +\171\001\000\000\171\001\171\001\000\000\000\000\000\000\000\000\ +\000\000\000\000\171\001\000\000\171\001\171\001\000\000\000\000\ +\171\001\171\001\000\000\000\000\000\000\171\001\000\000\000\000\ +\000\000\000\000\000\000\171\001\000\000\171\001\171\001\171\001\ +\171\001\171\001\171\001\000\000\000\000\000\000\000\000\000\000\ +\000\000\171\001\000\000\171\001\000\000\000\000\000\000\000\000\ +\000\000\171\001\171\001\000\000\000\000\000\000\171\001\171\001\ +\150\001\171\001\000\000\000\000\000\000\000\000\000\000\000\000\ +\171\001\000\000\171\001\171\001\171\001\171\001\000\000\171\001\ +\171\001\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ \000\000\000\000\000\000\000\000\000\000\000\000\171\001\171\001\ -\171\001\000\000\171\001\000\000\000\000\171\001\000\000\000\000\ -\171\001\171\001\171\001\000\000\171\001\171\001\000\000\000\000\ -\000\000\000\000\000\000\000\000\171\001\000\000\171\001\171\001\ -\000\000\000\000\171\001\171\001\000\000\000\000\000\000\171\001\ -\000\000\000\000\000\000\000\000\000\000\171\001\000\000\171\001\ -\171\001\171\001\171\001\171\001\171\001\000\000\000\000\000\000\ -\000\000\000\000\000\000\171\001\000\000\171\001\000\000\000\000\ -\000\000\000\000\000\000\171\001\171\001\000\000\000\000\000\000\ -\171\001\171\001\169\001\171\001\000\000\000\000\000\000\000\000\ -\000\000\000\000\171\001\000\000\171\001\171\001\171\001\171\001\ -\000\000\171\001\171\001\171\001\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\171\001\171\001\171\001\171\001\171\001\171\001\171\001\171\001\ -\171\001\000\000\000\000\000\000\000\000\000\000\171\001\171\001\ -\171\001\000\000\000\000\171\001\171\001\000\000\000\000\171\001\ -\171\001\171\001\171\001\171\001\000\000\000\000\171\001\171\001\ -\000\000\171\001\171\001\171\001\000\000\171\001\000\000\171\001\ -\000\000\000\000\000\000\171\001\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\171\001\171\001\000\000\171\001\171\001\ -\000\000\171\001\000\000\221\001\221\001\221\001\000\000\221\001\ -\000\000\000\000\000\000\000\000\000\000\000\000\221\001\221\001\ -\221\001\000\000\221\001\000\000\000\000\221\001\000\000\000\000\ -\000\000\221\001\221\001\000\000\221\001\221\001\000\000\000\000\ -\000\000\000\000\000\000\000\000\221\001\000\000\221\001\221\001\ -\000\000\000\000\221\001\221\001\000\000\000\000\000\000\221\001\ -\000\000\000\000\000\000\000\000\000\000\221\001\000\000\221\001\ -\221\001\221\001\221\001\221\001\221\001\000\000\000\000\000\000\ -\000\000\000\000\000\000\221\001\000\000\221\001\000\000\000\000\ -\000\000\000\000\000\000\221\001\221\001\000\000\000\000\000\000\ -\221\001\221\001\104\000\221\001\000\000\000\000\000\000\000\000\ -\000\000\000\000\221\001\000\000\221\001\221\001\221\001\221\001\ -\000\000\221\001\221\001\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\221\001\221\001\221\001\221\001\221\001\221\001\221\001\221\001\ -\221\001\000\000\000\000\000\000\000\000\000\000\221\001\221\001\ -\221\001\000\000\000\000\221\001\221\001\000\000\000\000\221\001\ -\221\001\221\001\221\001\221\001\000\000\000\000\221\001\221\001\ -\000\000\221\001\221\001\221\001\000\000\221\001\000\000\221\001\ -\000\000\000\000\000\000\221\001\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\221\001\221\001\000\000\221\001\221\001\ -\000\000\221\001\000\000\000\000\169\001\169\001\169\001\000\000\ -\169\001\000\000\000\000\000\000\000\000\000\000\000\000\169\001\ -\169\001\169\001\000\000\169\001\000\000\000\000\169\001\000\000\ -\000\000\000\000\169\001\169\001\000\000\169\001\169\001\000\000\ -\000\000\000\000\000\000\000\000\000\000\169\001\000\000\169\001\ -\169\001\000\000\000\000\169\001\169\001\000\000\000\000\000\000\ -\169\001\000\000\000\000\000\000\000\000\000\000\169\001\000\000\ -\169\001\169\001\169\001\169\001\169\001\169\001\000\000\000\000\ -\000\000\000\000\000\000\000\000\169\001\000\000\169\001\000\000\ -\000\000\000\000\000\000\000\000\169\001\169\001\000\000\000\000\ -\000\000\169\001\169\001\148\001\169\001\000\000\000\000\000\000\ -\000\000\000\000\000\000\169\001\000\000\169\001\169\001\169\001\ -\169\001\000\000\169\001\169\001\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\169\001\169\001\169\001\169\001\169\001\169\001\169\001\ -\169\001\169\001\000\000\000\000\000\000\000\000\000\000\169\001\ -\169\001\169\001\000\000\000\000\169\001\169\001\000\000\000\000\ -\169\001\169\001\169\001\169\001\169\001\000\000\000\000\169\001\ -\169\001\000\000\169\001\169\001\169\001\000\000\169\001\000\000\ -\169\001\000\000\000\000\000\000\169\001\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\169\001\169\001\000\000\169\001\ -\169\001\000\000\169\001\000\000\104\000\104\000\104\000\000\000\ -\104\000\000\000\000\000\000\000\000\000\000\000\000\000\104\000\ -\104\000\104\000\000\000\104\000\000\000\000\000\104\000\000\000\ +\171\001\171\001\171\001\171\001\171\001\171\001\171\001\000\000\ +\000\000\000\000\000\000\000\000\171\001\171\001\171\001\000\000\ +\000\000\171\001\171\001\000\000\000\000\171\001\171\001\171\001\ +\171\001\171\001\000\000\000\000\171\001\171\001\000\000\171\001\ +\171\001\171\001\000\000\171\001\000\000\171\001\000\000\000\000\ +\000\000\171\001\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\171\001\171\001\000\000\171\001\171\001\000\000\171\001\ +\000\000\000\000\104\000\104\000\104\000\000\000\104\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\104\000\104\000\104\000\ +\000\000\104\000\000\000\000\000\104\000\000\000\000\000\000\000\ +\000\000\104\000\000\000\104\000\104\000\000\000\000\000\000\000\ \000\000\000\000\000\000\104\000\000\000\104\000\104\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\104\000\000\000\104\000\ -\104\000\000\000\000\000\104\000\104\000\000\000\000\000\000\000\ -\104\000\000\000\000\000\000\000\000\000\000\000\104\000\000\000\ -\104\000\104\000\104\000\104\000\104\000\104\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\104\000\000\000\104\000\000\000\ -\000\000\000\000\000\000\000\000\104\000\104\000\000\000\000\000\ -\000\000\104\000\104\000\147\001\104\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\104\000\000\000\104\000\104\000\104\000\ -\104\000\000\000\104\000\104\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\104\000\104\000\104\000\104\000\104\000\104\000\104\000\ -\104\000\104\000\000\000\000\000\000\000\000\000\000\000\104\000\ -\104\000\104\000\000\000\000\000\104\000\104\000\000\000\000\000\ -\104\000\104\000\104\000\104\000\104\000\000\000\000\000\104\000\ -\104\000\000\000\104\000\104\000\104\000\000\000\104\000\000\000\ -\104\000\000\000\000\000\000\000\104\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\104\000\104\000\000\000\104\000\ -\104\000\000\000\104\000\000\000\000\000\148\001\148\001\148\001\ -\000\000\148\001\000\000\000\000\000\000\000\000\000\000\000\000\ -\148\001\148\001\148\001\000\000\148\001\000\000\000\000\148\001\ -\000\000\000\000\000\000\000\000\000\000\000\000\148\001\148\001\ -\000\000\000\000\000\000\000\000\000\000\000\000\148\001\000\000\ -\148\001\148\001\000\000\000\000\148\001\148\001\000\000\000\000\ -\000\000\148\001\000\000\000\000\000\000\000\000\000\000\148\001\ -\000\000\148\001\148\001\148\001\148\001\148\001\148\001\000\000\ -\000\000\000\000\000\000\000\000\000\000\148\001\000\000\148\001\ -\000\000\000\000\000\000\000\000\000\000\148\001\148\001\000\000\ -\000\000\000\000\148\001\148\001\123\001\148\001\000\000\000\000\ -\000\000\000\000\000\000\000\000\148\001\000\000\148\001\148\001\ -\148\001\148\001\000\000\148\001\148\001\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\148\001\148\001\148\001\148\001\148\001\148\001\ -\148\001\148\001\148\001\000\000\000\000\000\000\000\000\000\000\ -\148\001\148\001\148\001\000\000\000\000\148\001\148\001\000\000\ -\000\000\148\001\148\001\148\001\148\001\148\001\000\000\000\000\ -\148\001\148\001\000\000\148\001\148\001\148\001\000\000\148\001\ -\000\000\148\001\000\000\000\000\000\000\148\001\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\148\001\148\001\000\000\ -\148\001\148\001\000\000\148\001\000\000\147\001\147\001\147\001\ -\000\000\147\001\000\000\000\000\000\000\000\000\000\000\000\000\ -\147\001\147\001\147\001\000\000\147\001\000\000\000\000\147\001\ -\000\000\000\000\000\000\000\000\000\000\000\000\147\001\147\001\ -\000\000\000\000\000\000\000\000\000\000\000\000\147\001\000\000\ -\147\001\147\001\000\000\000\000\147\001\147\001\000\000\000\000\ -\000\000\147\001\000\000\000\000\000\000\000\000\000\000\147\001\ -\000\000\147\001\147\001\147\001\147\001\147\001\147\001\000\000\ -\000\000\000\000\000\000\000\000\000\000\147\001\121\001\147\001\ -\000\000\000\000\000\000\000\000\000\000\147\001\147\001\000\000\ -\000\000\000\000\147\001\147\001\000\000\147\001\000\000\000\000\ -\000\000\000\000\000\000\000\000\147\001\000\000\147\001\147\001\ -\147\001\147\001\000\000\147\001\147\001\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\147\001\147\001\147\001\147\001\147\001\147\001\ -\147\001\147\001\147\001\000\000\000\000\000\000\000\000\000\000\ -\147\001\147\001\147\001\000\000\000\000\147\001\147\001\000\000\ -\000\000\147\001\147\001\147\001\147\001\147\001\000\000\000\000\ -\147\001\147\001\000\000\147\001\147\001\147\001\000\000\147\001\ -\000\000\147\001\000\000\000\000\000\000\147\001\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\147\001\147\001\000\000\ -\147\001\147\001\000\000\147\001\000\000\000\000\123\001\123\001\ -\123\001\000\000\123\001\000\000\000\000\000\000\000\000\000\000\ -\000\000\123\001\000\000\123\001\000\000\123\001\000\000\000\000\ +\000\000\104\000\104\000\000\000\000\000\000\000\104\000\000\000\ +\000\000\000\000\000\000\000\000\104\000\000\000\104\000\104\000\ +\104\000\104\000\104\000\104\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\104\000\000\000\104\000\000\000\000\000\000\000\ +\000\000\000\000\104\000\104\000\000\000\000\000\000\000\104\000\ +\104\000\149\001\104\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\104\000\000\000\104\000\104\000\104\000\104\000\000\000\ +\104\000\104\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\104\000\ +\104\000\104\000\104\000\104\000\104\000\104\000\104\000\104\000\ +\000\000\000\000\000\000\000\000\000\000\104\000\104\000\104\000\ +\000\000\000\000\104\000\104\000\000\000\000\000\104\000\104\000\ +\104\000\104\000\104\000\000\000\000\000\104\000\104\000\000\000\ +\104\000\104\000\104\000\000\000\104\000\000\000\104\000\000\000\ +\000\000\000\000\104\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\104\000\104\000\000\000\104\000\104\000\000\000\ +\104\000\000\000\150\001\150\001\150\001\000\000\150\001\000\000\ +\000\000\000\000\000\000\000\000\000\000\150\001\150\001\150\001\ +\000\000\150\001\000\000\000\000\150\001\000\000\000\000\000\000\ +\000\000\000\000\000\000\150\001\150\001\000\000\000\000\000\000\ +\000\000\000\000\000\000\150\001\000\000\150\001\150\001\000\000\ +\000\000\150\001\150\001\000\000\000\000\000\000\150\001\000\000\ +\000\000\000\000\000\000\000\000\150\001\000\000\150\001\150\001\ +\150\001\150\001\150\001\150\001\000\000\000\000\000\000\000\000\ +\000\000\000\000\150\001\133\001\150\001\000\000\000\000\000\000\ +\000\000\000\000\150\001\150\001\000\000\000\000\000\000\150\001\ +\150\001\000\000\150\001\000\000\000\000\000\000\000\000\000\000\ +\000\000\150\001\000\000\150\001\150\001\150\001\150\001\000\000\ +\150\001\150\001\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\150\001\ +\150\001\150\001\150\001\150\001\150\001\150\001\150\001\150\001\ +\000\000\000\000\000\000\000\000\000\000\150\001\150\001\150\001\ +\000\000\000\000\150\001\150\001\000\000\000\000\150\001\150\001\ +\150\001\150\001\150\001\000\000\000\000\150\001\150\001\000\000\ +\150\001\150\001\150\001\000\000\150\001\000\000\150\001\000\000\ +\000\000\000\000\150\001\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\150\001\150\001\000\000\150\001\150\001\000\000\ +\150\001\000\000\000\000\149\001\149\001\149\001\000\000\149\001\ +\000\000\000\000\000\000\000\000\000\000\000\000\149\001\149\001\ +\149\001\000\000\149\001\000\000\000\000\149\001\000\000\000\000\ +\000\000\000\000\000\000\000\000\149\001\149\001\000\000\000\000\ +\000\000\000\000\000\000\000\000\149\001\000\000\149\001\149\001\ +\000\000\000\000\149\001\149\001\000\000\000\000\000\000\149\001\ +\000\000\000\000\000\000\000\000\000\000\149\001\000\000\149\001\ +\149\001\149\001\149\001\149\001\149\001\000\000\000\000\000\000\ +\000\000\000\000\125\001\149\001\000\000\149\001\000\000\000\000\ +\000\000\000\000\000\000\149\001\149\001\000\000\000\000\000\000\ +\149\001\149\001\000\000\149\001\000\000\000\000\000\000\000\000\ +\000\000\000\000\149\001\000\000\149\001\149\001\149\001\149\001\ +\000\000\149\001\149\001\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\149\001\149\001\149\001\149\001\149\001\149\001\149\001\149\001\ +\149\001\000\000\000\000\000\000\000\000\000\000\149\001\149\001\ +\149\001\000\000\000\000\149\001\149\001\000\000\000\000\149\001\ +\149\001\149\001\149\001\149\001\000\000\000\000\149\001\149\001\ +\000\000\149\001\149\001\149\001\000\000\149\001\000\000\149\001\ +\000\000\000\000\000\000\149\001\000\000\133\001\133\001\133\001\ +\000\000\133\001\000\000\149\001\149\001\000\000\149\001\149\001\ +\133\001\149\001\133\001\000\000\133\001\000\000\000\000\133\001\ +\000\000\000\000\000\000\000\000\000\000\000\000\133\001\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\133\001\000\000\ +\133\001\133\001\000\000\000\000\133\001\133\001\000\000\000\000\ +\000\000\133\001\000\000\000\000\000\000\000\000\000\000\133\001\ +\000\000\133\001\133\001\133\001\133\001\133\001\133\001\000\000\ +\000\000\000\000\000\000\000\000\123\001\133\001\000\000\133\001\ +\000\000\000\000\000\000\000\000\000\000\133\001\133\001\000\000\ +\000\000\000\000\133\001\133\001\000\000\133\001\000\000\000\000\ +\000\000\000\000\000\000\000\000\133\001\000\000\133\001\133\001\ +\133\001\133\001\000\000\133\001\133\001\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\133\001\133\001\133\001\133\001\133\001\133\001\ +\133\001\133\001\133\001\000\000\000\000\000\000\000\000\000\000\ +\000\000\133\001\133\001\000\000\000\000\133\001\133\001\000\000\ +\000\000\133\001\133\001\133\001\133\001\133\001\000\000\000\000\ +\133\001\000\000\000\000\133\001\133\001\133\001\000\000\133\001\ +\000\000\133\001\000\000\000\000\000\000\133\001\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\133\001\133\001\000\000\ +\133\001\133\001\000\000\133\001\125\001\125\001\125\001\000\000\ +\125\001\000\000\000\000\000\000\000\000\000\000\000\000\125\001\ +\000\000\125\001\000\000\125\001\000\000\000\000\125\001\000\000\ +\000\000\000\000\000\000\000\000\000\000\125\001\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\125\001\000\000\125\001\ +\125\001\000\000\000\000\125\001\125\001\000\000\000\000\000\000\ +\125\001\000\000\000\000\000\000\000\000\000\000\125\001\000\000\ +\000\000\125\001\125\001\125\001\125\001\125\001\000\000\000\000\ +\000\000\000\000\000\000\116\001\125\001\000\000\125\001\000\000\ +\000\000\000\000\000\000\000\000\125\001\125\001\000\000\000\000\ +\000\000\125\001\125\001\000\000\125\001\000\000\000\000\000\000\ +\000\000\000\000\000\000\125\001\000\000\125\001\125\001\125\001\ +\125\001\000\000\125\001\125\001\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\125\001\125\001\125\001\125\001\125\001\125\001\125\001\ +\125\001\125\001\000\000\000\000\000\000\000\000\000\000\000\000\ +\125\001\125\001\000\000\000\000\125\001\125\001\000\000\000\000\ +\125\001\125\001\125\001\125\001\125\001\000\000\000\000\125\001\ +\000\000\000\000\125\001\125\001\125\001\000\000\125\001\000\000\ +\125\001\000\000\000\000\000\000\125\001\000\000\123\001\123\001\ +\123\001\000\000\123\001\000\000\125\001\125\001\000\000\125\001\ +\125\001\123\001\125\001\123\001\000\000\123\001\000\000\000\000\ \123\001\000\000\000\000\000\000\000\000\000\000\000\000\123\001\ \000\000\000\000\000\000\000\000\000\000\000\000\000\000\123\001\ \000\000\123\001\123\001\000\000\000\000\123\001\123\001\000\000\ \000\000\000\000\123\001\000\000\000\000\000\000\000\000\000\000\ -\123\001\000\000\000\000\123\001\123\001\123\001\123\001\123\001\ -\000\000\000\000\000\000\000\000\000\000\114\001\123\001\000\000\ +\000\000\000\000\000\000\123\001\123\001\123\001\123\001\123\001\ +\000\000\000\000\000\000\000\000\000\000\058\000\123\001\000\000\ \123\001\000\000\000\000\000\000\000\000\000\000\123\001\123\001\ \000\000\000\000\000\000\123\001\123\001\000\000\123\001\000\000\ \000\000\000\000\000\000\000\000\000\000\123\001\000\000\123\001\ @@ -2340,206 +2403,206 @@ let yytable = "\118\000\ \000\000\000\000\123\001\123\001\123\001\123\001\123\001\000\000\ \000\000\123\001\000\000\000\000\123\001\123\001\123\001\000\000\ \123\001\000\000\123\001\000\000\000\000\000\000\123\001\000\000\ -\121\001\121\001\121\001\000\000\121\001\000\000\123\001\123\001\ -\000\000\123\001\123\001\121\001\123\001\121\001\000\000\121\001\ -\000\000\000\000\121\001\000\000\000\000\000\000\000\000\000\000\ -\000\000\121\001\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\121\001\000\000\121\001\121\001\000\000\000\000\121\001\ -\121\001\000\000\000\000\000\000\121\001\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\121\001\121\001\121\001\ -\121\001\121\001\000\000\000\000\000\000\000\000\000\000\058\000\ -\121\001\000\000\121\001\000\000\000\000\000\000\000\000\000\000\ -\121\001\121\001\000\000\000\000\000\000\121\001\121\001\000\000\ -\121\001\000\000\000\000\000\000\000\000\000\000\000\000\121\001\ -\000\000\121\001\121\001\121\001\121\001\000\000\121\001\121\001\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\121\001\121\001\121\001\ -\121\001\121\001\121\001\121\001\121\001\121\001\000\000\000\000\ -\000\000\000\000\000\000\000\000\121\001\121\001\000\000\000\000\ -\121\001\121\001\000\000\000\000\121\001\121\001\121\001\121\001\ -\121\001\000\000\000\000\121\001\000\000\000\000\121\001\121\001\ -\121\001\000\000\121\001\000\000\121\001\000\000\000\000\000\000\ -\121\001\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\121\001\121\001\000\000\121\001\121\001\000\000\121\001\114\001\ -\114\001\114\001\000\000\114\001\000\000\000\000\000\000\000\000\ -\000\000\000\000\114\001\000\000\114\001\000\000\000\000\000\000\ -\000\000\114\001\000\000\000\000\000\000\000\000\000\000\000\000\ -\114\001\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\114\001\000\000\114\001\114\001\000\000\000\000\000\000\114\001\ -\000\000\000\000\000\000\114\001\000\000\000\000\000\000\000\000\ -\000\000\005\000\000\000\000\000\114\001\114\001\114\001\114\001\ -\114\001\000\000\000\000\000\000\000\000\000\000\000\000\114\001\ -\000\000\114\001\000\000\000\000\000\000\000\000\000\000\114\001\ -\114\001\000\000\000\000\000\000\114\001\114\001\000\000\114\001\ -\000\000\000\000\000\000\000\000\000\000\000\000\114\001\000\000\ -\114\001\114\001\114\001\114\001\000\000\114\001\114\001\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\114\001\114\001\114\001\114\001\ -\114\001\114\001\114\001\114\001\114\001\000\000\000\000\000\000\ -\000\000\000\000\000\000\114\001\114\001\000\000\000\000\114\001\ -\114\001\000\000\000\000\114\001\114\001\114\001\114\001\114\001\ -\000\000\000\000\114\001\000\000\000\000\114\001\000\000\114\001\ -\000\000\114\001\000\000\114\001\000\000\000\000\000\000\114\001\ -\000\000\058\000\058\000\058\000\000\000\058\000\000\000\114\001\ -\114\001\000\000\114\001\114\001\058\000\114\001\058\000\000\000\ -\000\000\000\000\000\000\058\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\058\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\058\000\000\000\058\000\058\000\000\000\000\000\ -\000\000\058\000\000\000\000\000\000\000\058\000\000\000\000\000\ -\000\000\000\000\000\000\078\000\000\000\000\000\058\000\058\000\ -\058\000\058\000\058\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\058\000\000\000\058\000\000\000\000\000\000\000\000\000\ -\000\000\058\000\058\000\000\000\000\000\000\000\058\000\058\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\123\001\123\001\ +\000\000\123\001\123\001\000\000\123\001\116\001\116\001\116\001\ +\000\000\116\001\000\000\000\000\000\000\000\000\000\000\000\000\ +\116\001\000\000\116\001\000\000\000\000\000\000\000\000\116\001\ +\000\000\000\000\000\000\000\000\000\000\000\000\116\001\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\116\001\000\000\ +\116\001\116\001\000\000\000\000\000\000\116\001\000\000\000\000\ +\000\000\116\001\000\000\000\000\000\000\000\000\000\000\005\000\ +\000\000\000\000\116\001\116\001\116\001\116\001\116\001\000\000\ +\000\000\000\000\000\000\000\000\000\000\116\001\000\000\116\001\ +\000\000\000\000\000\000\000\000\000\000\116\001\116\001\000\000\ +\000\000\000\000\116\001\116\001\000\000\116\001\000\000\000\000\ +\000\000\000\000\000\000\000\000\116\001\000\000\116\001\116\001\ +\116\001\116\001\000\000\116\001\116\001\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\116\001\116\001\116\001\116\001\116\001\116\001\ +\116\001\116\001\116\001\000\000\000\000\000\000\000\000\000\000\ +\000\000\116\001\116\001\000\000\000\000\116\001\116\001\000\000\ +\000\000\116\001\116\001\116\001\116\001\116\001\000\000\000\000\ +\116\001\000\000\000\000\116\001\000\000\116\001\000\000\116\001\ +\000\000\116\001\000\000\000\000\000\000\116\001\000\000\058\000\ +\058\000\058\000\000\000\058\000\000\000\116\001\116\001\000\000\ +\116\001\116\001\058\000\116\001\058\000\000\000\000\000\000\000\ \000\000\058\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\058\000\000\000\058\000\058\000\058\000\058\000\000\000\058\000\ \058\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\058\000\058\000\ -\058\000\058\000\058\000\058\000\058\000\058\000\058\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\058\000\058\000\000\000\ -\000\000\058\000\058\000\000\000\000\000\058\000\058\000\058\000\ -\058\000\058\000\000\000\000\000\058\000\000\000\000\000\058\000\ -\000\000\058\000\000\000\058\000\000\000\058\000\000\000\000\000\ -\000\000\058\000\000\000\005\000\005\000\005\000\000\000\005\000\ -\000\000\058\000\058\000\000\000\058\000\058\000\005\000\058\000\ -\005\000\000\000\000\000\000\000\000\000\005\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\005\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\005\000\000\000\005\000\005\000\ -\000\000\000\000\000\000\005\000\000\000\000\000\000\000\005\000\ -\000\000\000\000\000\000\000\000\000\000\019\000\000\000\000\000\ -\005\000\005\000\005\000\005\000\005\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\005\000\000\000\005\000\000\000\000\000\ -\000\000\000\000\000\000\005\000\005\000\000\000\000\000\000\000\ -\005\000\005\000\000\000\005\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\005\000\000\000\005\000\005\000\005\000\005\000\ -\000\000\005\000\005\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\005\000\005\000\005\000\005\000\005\000\005\000\005\000\005\000\ -\005\000\000\000\000\000\000\000\000\000\000\000\000\000\005\000\ -\005\000\000\000\000\000\005\000\005\000\000\000\000\000\005\000\ -\005\000\005\000\005\000\005\000\000\000\000\000\005\000\000\000\ -\000\000\005\000\000\000\005\000\000\000\005\000\000\000\005\000\ -\000\000\000\000\000\000\005\000\000\000\078\000\078\000\078\000\ -\000\000\078\000\000\000\005\000\005\000\000\000\005\000\005\000\ -\078\000\005\000\078\000\000\000\000\000\000\000\000\000\078\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\078\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\078\000\000\000\ -\078\000\078\000\000\000\000\000\000\000\078\000\000\000\000\000\ -\000\000\078\000\000\000\000\000\000\000\000\000\000\000\017\000\ -\000\000\000\000\078\000\078\000\078\000\078\000\078\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\078\000\000\000\078\000\ -\000\000\000\000\000\000\000\000\000\000\078\000\078\000\000\000\ -\000\000\000\000\078\000\078\000\000\000\078\000\000\000\000\000\ +\058\000\000\000\058\000\058\000\000\000\000\000\000\000\058\000\ +\000\000\000\000\000\000\058\000\000\000\000\000\000\000\000\000\ +\000\000\078\000\000\000\000\000\058\000\058\000\058\000\058\000\ +\058\000\000\000\000\000\000\000\000\000\000\000\000\000\058\000\ +\000\000\058\000\000\000\000\000\000\000\000\000\000\000\058\000\ +\058\000\000\000\000\000\000\000\058\000\058\000\000\000\058\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\058\000\000\000\ +\058\000\058\000\058\000\058\000\000\000\058\000\058\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\058\000\058\000\058\000\058\000\ +\058\000\058\000\058\000\058\000\058\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\058\000\058\000\000\000\000\000\058\000\ +\058\000\000\000\000\000\058\000\058\000\058\000\058\000\058\000\ +\000\000\000\000\058\000\000\000\000\000\058\000\000\000\058\000\ +\000\000\058\000\000\000\058\000\000\000\000\000\000\000\058\000\ +\000\000\005\000\005\000\005\000\000\000\005\000\000\000\058\000\ +\058\000\000\000\058\000\058\000\005\000\058\000\005\000\000\000\ +\000\000\000\000\000\000\005\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\005\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\005\000\000\000\005\000\005\000\000\000\000\000\ +\000\000\005\000\000\000\000\000\000\000\005\000\000\000\000\000\ +\000\000\000\000\000\000\019\000\000\000\000\000\005\000\005\000\ +\005\000\005\000\005\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\005\000\000\000\005\000\000\000\000\000\000\000\000\000\ +\000\000\005\000\005\000\000\000\000\000\000\000\005\000\005\000\ +\000\000\005\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\005\000\000\000\005\000\005\000\005\000\005\000\000\000\005\000\ +\005\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\005\000\005\000\ +\005\000\005\000\005\000\005\000\005\000\005\000\005\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\005\000\005\000\000\000\ +\000\000\005\000\005\000\000\000\000\000\005\000\005\000\005\000\ +\005\000\005\000\000\000\000\000\005\000\000\000\000\000\005\000\ +\000\000\005\000\000\000\005\000\000\000\005\000\000\000\000\000\ +\000\000\005\000\000\000\078\000\078\000\078\000\000\000\078\000\ +\000\000\005\000\005\000\000\000\005\000\005\000\078\000\005\000\ +\078\000\000\000\000\000\000\000\000\000\078\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\078\000\000\000\000\000\000\000\ \000\000\000\000\000\000\000\000\078\000\000\000\078\000\078\000\ -\078\000\078\000\000\000\078\000\078\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\078\000\078\000\078\000\078\000\078\000\078\000\ -\078\000\078\000\078\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\078\000\078\000\000\000\000\000\078\000\078\000\000\000\ -\000\000\078\000\078\000\078\000\078\000\078\000\000\000\000\000\ -\078\000\000\000\000\000\078\000\000\000\078\000\000\000\078\000\ -\000\000\078\000\000\000\000\000\000\000\078\000\000\000\019\000\ -\019\000\019\000\000\000\019\000\000\000\078\000\078\000\000\000\ -\078\000\078\000\019\000\078\000\019\000\000\000\000\000\000\000\ -\000\000\019\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\019\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\019\000\000\000\019\000\019\000\000\000\000\000\000\000\019\000\ -\000\000\000\000\000\000\019\000\000\000\000\000\000\000\000\000\ -\000\000\068\002\000\000\000\000\019\000\019\000\019\000\019\000\ -\019\000\000\000\000\000\000\000\000\000\000\000\000\000\019\000\ -\000\000\019\000\000\000\000\000\000\000\000\000\000\000\019\000\ -\019\000\000\000\000\000\000\000\019\000\019\000\000\000\019\000\ +\000\000\000\000\000\000\078\000\000\000\000\000\000\000\078\000\ +\000\000\000\000\000\000\000\000\000\000\017\000\000\000\000\000\ +\078\000\078\000\078\000\078\000\078\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\078\000\000\000\078\000\000\000\000\000\ +\000\000\000\000\000\000\078\000\078\000\000\000\000\000\000\000\ +\078\000\078\000\000\000\078\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\078\000\000\000\078\000\078\000\078\000\078\000\ +\000\000\078\000\078\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\078\000\078\000\078\000\078\000\078\000\078\000\078\000\078\000\ +\078\000\000\000\000\000\000\000\000\000\000\000\000\000\078\000\ +\078\000\000\000\000\000\078\000\078\000\000\000\000\000\078\000\ +\078\000\078\000\078\000\078\000\000\000\000\000\078\000\000\000\ +\000\000\078\000\000\000\078\000\000\000\078\000\000\000\078\000\ +\000\000\000\000\000\000\078\000\000\000\019\000\019\000\019\000\ +\000\000\019\000\000\000\078\000\078\000\000\000\078\000\078\000\ +\019\000\078\000\019\000\000\000\000\000\000\000\000\000\019\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\019\000\000\000\ \000\000\000\000\000\000\000\000\000\000\000\000\019\000\000\000\ -\019\000\019\000\019\000\019\000\000\000\019\000\019\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\019\000\019\000\019\000\019\000\ -\019\000\019\000\019\000\019\000\019\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\019\000\019\000\000\000\000\000\019\000\ -\019\000\000\000\000\000\019\000\019\000\019\000\019\000\019\000\ -\000\000\000\000\019\000\000\000\000\000\019\000\000\000\019\000\ -\000\000\019\000\000\000\019\000\000\000\000\000\000\000\019\000\ -\000\000\017\000\017\000\017\000\000\000\017\000\000\000\019\000\ -\019\000\000\000\019\000\019\000\017\000\019\000\017\000\000\000\ -\000\000\000\000\000\000\017\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\017\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\017\000\000\000\017\000\017\000\000\000\000\000\ -\000\000\017\000\000\000\000\000\000\000\017\000\086\001\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\017\000\017\000\ -\017\000\017\000\017\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\017\000\000\000\017\000\000\000\000\000\000\000\000\000\ -\000\000\017\000\017\000\000\000\000\000\000\000\017\000\017\000\ +\019\000\019\000\000\000\000\000\000\000\019\000\000\000\000\000\ +\000\000\019\000\000\000\000\000\000\000\000\000\000\000\070\002\ +\000\000\000\000\019\000\019\000\019\000\019\000\019\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\019\000\000\000\019\000\ +\000\000\000\000\000\000\000\000\000\000\019\000\019\000\000\000\ +\000\000\000\000\019\000\019\000\000\000\019\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\019\000\000\000\019\000\019\000\ +\019\000\019\000\000\000\019\000\019\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\019\000\019\000\019\000\019\000\019\000\019\000\ +\019\000\019\000\019\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\019\000\019\000\000\000\000\000\019\000\019\000\000\000\ +\000\000\019\000\019\000\019\000\019\000\019\000\000\000\000\000\ +\019\000\000\000\000\000\019\000\000\000\019\000\000\000\019\000\ +\000\000\019\000\000\000\000\000\000\000\019\000\000\000\017\000\ +\017\000\017\000\000\000\017\000\000\000\019\000\019\000\000\000\ +\019\000\019\000\017\000\019\000\017\000\000\000\000\000\000\000\ \000\000\017\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\017\000\000\000\017\000\017\000\017\000\017\000\000\000\017\000\ \017\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\017\000\017\000\ -\017\000\017\000\017\000\017\000\017\000\017\000\017\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\017\000\017\000\000\000\ -\000\000\017\000\017\000\000\000\000\000\017\000\017\000\017\000\ -\017\000\017\000\000\000\000\000\017\000\000\000\000\000\017\000\ -\000\000\017\000\000\000\017\000\000\000\017\000\000\000\000\000\ -\000\000\017\000\000\000\068\002\068\002\068\002\000\000\068\002\ -\000\000\017\000\017\000\000\000\017\000\017\000\068\002\017\000\ -\068\002\000\000\000\000\000\000\000\000\068\002\000\000\000\000\ -\000\000\000\000\000\000\000\000\068\002\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\068\002\000\000\068\002\068\002\ -\000\000\000\000\000\000\068\002\000\000\000\000\000\000\068\002\ -\080\001\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\068\002\068\002\068\002\068\002\068\002\000\000\000\000\000\000\ -\000\000\000\000\000\000\068\002\000\000\068\002\000\000\000\000\ -\000\000\000\000\000\000\068\002\068\002\000\000\000\000\000\000\ -\068\002\068\002\000\000\068\002\000\000\000\000\000\000\000\000\ -\000\000\000\000\068\002\000\000\068\002\068\002\068\002\068\002\ -\000\000\068\002\068\002\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\068\002\068\002\068\002\068\002\068\002\068\002\068\002\068\002\ -\068\002\000\000\000\000\000\000\000\000\000\000\000\000\068\002\ -\068\002\000\000\000\000\068\002\068\002\000\000\000\000\068\002\ -\068\002\068\002\068\002\068\002\000\000\000\000\068\002\000\000\ -\000\000\068\002\000\000\068\002\000\000\068\002\000\000\068\002\ -\086\001\086\001\086\001\068\002\086\001\000\000\000\000\000\000\ -\000\000\000\000\000\000\068\002\068\002\086\001\068\002\068\002\ -\000\000\068\002\086\001\000\000\000\000\000\000\000\000\000\000\ -\000\000\086\001\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\086\001\000\000\086\001\086\001\000\000\000\000\000\000\ -\086\001\000\000\000\000\000\000\086\001\081\001\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\086\001\086\001\086\001\ -\086\001\086\001\000\000\000\000\000\000\000\000\000\000\000\000\ -\086\001\000\000\086\001\000\000\000\000\000\000\000\000\000\000\ -\086\001\086\001\000\000\000\000\000\000\086\001\086\001\000\000\ -\086\001\000\000\000\000\000\000\000\000\000\000\000\000\086\001\ -\000\000\086\001\086\001\086\001\086\001\000\000\086\001\086\001\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\086\001\086\001\086\001\ -\086\001\086\001\086\001\086\001\086\001\086\001\000\000\000\000\ -\000\000\000\000\000\000\000\000\086\001\086\001\000\000\000\000\ -\086\001\086\001\000\000\000\000\086\001\086\001\086\001\086\001\ -\086\001\000\000\000\000\086\001\000\000\000\000\086\001\000\000\ -\086\001\000\000\086\001\000\000\086\001\000\000\000\000\000\000\ -\086\001\000\000\080\001\080\001\080\001\000\000\080\001\000\000\ -\086\001\086\001\000\000\086\001\086\001\000\000\086\001\080\001\ -\000\000\000\000\000\000\000\000\080\001\000\000\000\000\000\000\ -\000\000\000\000\000\000\080\001\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\080\001\000\000\080\001\080\001\000\000\ -\000\000\000\000\080\001\000\000\000\000\000\000\080\001\079\001\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\080\001\ -\080\001\080\001\080\001\080\001\000\000\000\000\000\000\000\000\ -\000\000\000\000\080\001\000\000\080\001\000\000\000\000\000\000\ -\000\000\000\000\080\001\080\001\000\000\000\000\000\000\080\001\ -\080\001\000\000\080\001\000\000\000\000\000\000\000\000\000\000\ -\000\000\080\001\000\000\080\001\080\001\080\001\080\001\000\000\ -\080\001\080\001\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\080\001\ -\080\001\080\001\080\001\080\001\080\001\080\001\080\001\080\001\ -\000\000\000\000\000\000\000\000\000\000\000\000\080\001\080\001\ -\000\000\000\000\080\001\080\001\000\000\000\000\080\001\080\001\ -\080\001\080\001\080\001\000\000\000\000\080\001\000\000\000\000\ -\080\001\000\000\080\001\000\000\080\001\000\000\080\001\081\001\ -\081\001\081\001\080\001\081\001\000\000\000\000\000\000\000\000\ -\000\000\000\000\080\001\080\001\081\001\080\001\080\001\000\000\ -\080\001\081\001\000\000\000\000\000\000\000\000\000\000\000\000\ +\017\000\000\000\017\000\017\000\000\000\000\000\000\000\017\000\ +\000\000\000\000\000\000\017\000\088\001\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\017\000\017\000\017\000\017\000\ +\017\000\000\000\000\000\000\000\000\000\000\000\000\000\017\000\ +\000\000\017\000\000\000\000\000\000\000\000\000\000\000\017\000\ +\017\000\000\000\000\000\000\000\017\000\017\000\000\000\017\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\017\000\000\000\ +\017\000\017\000\017\000\017\000\000\000\017\000\017\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\017\000\017\000\017\000\017\000\ +\017\000\017\000\017\000\017\000\017\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\017\000\017\000\000\000\000\000\017\000\ +\017\000\000\000\000\000\017\000\017\000\017\000\017\000\017\000\ +\000\000\000\000\017\000\000\000\000\000\017\000\000\000\017\000\ +\000\000\017\000\000\000\017\000\000\000\000\000\000\000\017\000\ +\000\000\070\002\070\002\070\002\000\000\070\002\000\000\017\000\ +\017\000\000\000\017\000\017\000\070\002\017\000\070\002\000\000\ +\000\000\000\000\000\000\070\002\000\000\000\000\000\000\000\000\ +\000\000\000\000\070\002\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\070\002\000\000\070\002\070\002\000\000\000\000\ +\000\000\070\002\000\000\000\000\000\000\070\002\082\001\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\070\002\070\002\ +\070\002\070\002\070\002\000\000\000\000\000\000\000\000\000\000\ +\000\000\070\002\000\000\070\002\000\000\000\000\000\000\000\000\ +\000\000\070\002\070\002\000\000\000\000\000\000\070\002\070\002\ +\000\000\070\002\000\000\000\000\000\000\000\000\000\000\000\000\ +\070\002\000\000\070\002\070\002\070\002\070\002\000\000\070\002\ +\070\002\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\070\002\070\002\ +\070\002\070\002\070\002\070\002\070\002\070\002\070\002\000\000\ +\000\000\000\000\000\000\000\000\000\000\070\002\070\002\000\000\ +\000\000\070\002\070\002\000\000\000\000\070\002\070\002\070\002\ +\070\002\070\002\000\000\000\000\070\002\000\000\000\000\070\002\ +\000\000\070\002\000\000\070\002\000\000\070\002\088\001\088\001\ +\088\001\070\002\088\001\000\000\000\000\000\000\000\000\000\000\ +\000\000\070\002\070\002\088\001\070\002\070\002\000\000\070\002\ +\088\001\000\000\000\000\000\000\000\000\000\000\000\000\088\001\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\088\001\ +\000\000\088\001\088\001\000\000\000\000\000\000\088\001\000\000\ +\000\000\000\000\088\001\083\001\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\088\001\088\001\088\001\088\001\088\001\ +\000\000\000\000\000\000\000\000\000\000\000\000\088\001\000\000\ +\088\001\000\000\000\000\000\000\000\000\000\000\088\001\088\001\ +\000\000\000\000\000\000\088\001\088\001\000\000\088\001\000\000\ +\000\000\000\000\000\000\000\000\000\000\088\001\000\000\088\001\ +\088\001\088\001\088\001\000\000\088\001\088\001\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\088\001\088\001\088\001\088\001\088\001\ +\088\001\088\001\088\001\088\001\000\000\000\000\000\000\000\000\ +\000\000\000\000\088\001\088\001\000\000\000\000\088\001\088\001\ +\000\000\000\000\088\001\088\001\088\001\088\001\088\001\000\000\ +\000\000\088\001\000\000\000\000\088\001\000\000\088\001\000\000\ +\088\001\000\000\088\001\000\000\000\000\000\000\088\001\000\000\ +\082\001\082\001\082\001\000\000\082\001\000\000\088\001\088\001\ +\000\000\088\001\088\001\000\000\088\001\082\001\000\000\000\000\ +\000\000\000\000\082\001\000\000\000\000\000\000\000\000\000\000\ +\000\000\082\001\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\082\001\000\000\082\001\082\001\000\000\000\000\000\000\ +\082\001\000\000\000\000\000\000\082\001\081\001\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\082\001\082\001\082\001\ +\082\001\082\001\000\000\000\000\000\000\000\000\000\000\000\000\ +\082\001\000\000\082\001\000\000\000\000\000\000\000\000\000\000\ +\082\001\082\001\000\000\000\000\000\000\082\001\082\001\000\000\ +\082\001\000\000\000\000\000\000\000\000\000\000\000\000\082\001\ +\000\000\082\001\082\001\082\001\082\001\000\000\082\001\082\001\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\082\001\082\001\082\001\ +\082\001\082\001\082\001\082\001\082\001\082\001\000\000\000\000\ +\000\000\000\000\000\000\000\000\082\001\082\001\000\000\000\000\ +\082\001\082\001\000\000\000\000\082\001\082\001\082\001\082\001\ +\082\001\000\000\000\000\082\001\000\000\000\000\082\001\000\000\ +\082\001\000\000\082\001\000\000\082\001\083\001\083\001\083\001\ +\082\001\083\001\000\000\000\000\000\000\000\000\000\000\000\000\ +\082\001\082\001\083\001\082\001\082\001\000\000\082\001\083\001\ +\000\000\000\000\000\000\000\000\000\000\000\000\083\001\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\083\001\000\000\ +\083\001\083\001\000\000\000\000\000\000\083\001\000\000\000\000\ +\000\000\083\001\000\000\000\000\000\000\000\000\000\000\062\001\ +\000\000\000\000\083\001\083\001\083\001\083\001\083\001\000\000\ +\000\000\000\000\000\000\000\000\000\000\083\001\000\000\083\001\ +\000\000\000\000\000\000\000\000\000\000\083\001\083\001\000\000\ +\000\000\000\000\083\001\083\001\000\000\083\001\000\000\000\000\ +\000\000\000\000\000\000\000\000\083\001\000\000\083\001\083\001\ +\083\001\083\001\000\000\083\001\083\001\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\083\001\083\001\083\001\083\001\083\001\083\001\ +\083\001\083\001\083\001\000\000\000\000\000\000\000\000\000\000\ +\000\000\083\001\083\001\000\000\000\000\083\001\083\001\000\000\ +\000\000\083\001\083\001\083\001\083\001\083\001\000\000\000\000\ +\083\001\000\000\000\000\083\001\000\000\083\001\000\000\083\001\ +\000\000\083\001\000\000\000\000\000\000\083\001\000\000\081\001\ +\081\001\081\001\000\000\081\001\000\000\083\001\083\001\000\000\ +\083\001\083\001\081\001\083\001\081\001\000\000\000\000\000\000\ +\000\000\081\001\000\000\000\000\000\000\000\000\000\000\000\000\ \081\001\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\081\001\000\000\081\001\081\001\000\000\000\000\000\000\081\001\ -\000\000\000\000\000\000\081\001\000\000\000\000\000\000\000\000\ -\000\000\060\001\000\000\000\000\081\001\081\001\081\001\081\001\ +\081\001\000\000\000\000\081\001\000\000\000\000\000\000\081\001\ +\000\000\000\000\000\000\081\001\000\000\224\001\000\000\000\000\ +\000\000\000\000\000\000\000\000\081\001\081\001\081\001\081\001\ \081\001\000\000\000\000\000\000\000\000\000\000\000\000\081\001\ \000\000\081\001\000\000\000\000\000\000\000\000\000\000\081\001\ \081\001\000\000\000\000\000\000\081\001\081\001\000\000\081\001\ @@ -2552,1748 +2615,1823 @@ let yytable = "\118\000\ \081\001\000\000\000\000\081\001\081\001\081\001\081\001\081\001\ \000\000\000\000\081\001\000\000\000\000\081\001\000\000\081\001\ \000\000\081\001\000\000\081\001\000\000\000\000\000\000\081\001\ -\000\000\079\001\079\001\079\001\000\000\079\001\000\000\081\001\ -\081\001\000\000\081\001\081\001\079\001\081\001\079\001\000\000\ -\000\000\000\000\000\000\079\001\000\000\000\000\000\000\000\000\ -\000\000\000\000\079\001\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\079\001\000\000\000\000\079\001\000\000\000\000\ -\000\000\079\001\000\000\000\000\000\000\079\001\000\000\222\001\ -\000\000\000\000\000\000\000\000\000\000\000\000\079\001\079\001\ -\079\001\079\001\079\001\000\000\000\000\000\000\000\000\000\000\ -\000\000\079\001\000\000\079\001\000\000\000\000\000\000\000\000\ -\000\000\079\001\079\001\000\000\000\000\000\000\079\001\079\001\ -\000\000\079\001\000\000\000\000\000\000\000\000\000\000\000\000\ -\079\001\000\000\079\001\079\001\079\001\079\001\000\000\079\001\ -\079\001\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\079\001\079\001\ -\079\001\079\001\079\001\079\001\079\001\079\001\079\001\000\000\ -\000\000\000\000\000\000\000\000\000\000\079\001\079\001\000\000\ -\000\000\079\001\079\001\000\000\000\000\079\001\079\001\079\001\ -\079\001\079\001\000\000\000\000\079\001\000\000\000\000\079\001\ -\000\000\079\001\000\000\079\001\000\000\079\001\000\000\000\000\ -\000\000\079\001\000\000\060\001\060\001\060\001\000\000\060\001\ -\000\000\079\001\079\001\000\000\079\001\079\001\060\001\079\001\ -\060\001\000\000\000\000\000\000\000\000\060\001\000\000\000\000\ -\000\000\000\000\000\000\139\001\060\001\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\060\001\000\000\060\001\060\001\ -\000\000\000\000\000\000\060\001\000\000\000\000\000\000\060\001\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\060\001\060\001\060\001\060\001\060\001\000\000\000\000\000\000\ -\000\000\000\000\000\000\060\001\000\000\060\001\000\000\000\000\ -\000\000\000\000\000\000\060\001\060\001\000\000\000\000\000\000\ -\060\001\060\001\000\000\060\001\000\000\000\000\000\000\000\000\ -\000\000\000\000\060\001\000\000\060\001\060\001\060\001\060\001\ -\000\000\060\001\060\001\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\060\001\060\001\060\001\060\001\060\001\060\001\060\001\060\001\ -\060\001\000\000\000\000\000\000\000\000\000\000\000\000\060\001\ -\060\001\000\000\000\000\060\001\060\001\000\000\000\000\060\001\ -\060\001\060\001\000\000\000\000\000\000\000\000\060\001\000\000\ -\000\000\060\001\000\000\060\001\000\000\060\001\000\000\060\001\ -\000\000\222\001\222\001\060\001\000\000\222\001\000\000\000\000\ -\000\000\000\000\000\000\060\001\060\001\222\001\060\001\060\001\ -\222\001\060\001\000\000\222\001\000\000\000\000\000\000\000\000\ -\053\000\000\000\222\001\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\222\001\000\000\000\000\000\000\000\000\222\001\ -\000\000\222\001\000\000\000\000\000\000\222\001\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\222\001\222\001\ -\222\001\222\001\222\001\000\000\000\000\000\000\000\000\000\000\ -\000\000\222\001\000\000\222\001\000\000\000\000\000\000\000\000\ -\000\000\222\001\222\001\000\000\000\000\000\000\000\000\222\001\ -\000\000\222\001\000\000\000\000\000\000\000\000\000\000\000\000\ -\222\001\000\000\222\001\222\001\222\001\222\001\000\000\222\001\ -\222\001\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\222\001\222\001\ -\222\001\222\001\222\001\222\001\222\001\222\001\222\001\000\000\ -\000\000\000\000\000\000\000\000\222\001\222\001\222\001\000\000\ -\000\000\222\001\222\001\000\000\000\000\139\001\139\001\000\000\ -\000\000\139\001\000\000\000\000\222\001\000\000\000\000\000\000\ -\000\000\222\001\000\000\000\000\139\001\222\001\000\000\139\001\ -\000\000\222\001\000\000\000\000\000\000\000\000\139\001\000\000\ -\000\000\222\001\222\001\000\000\222\001\222\001\139\001\222\001\ -\000\000\000\000\000\000\139\001\000\000\139\001\000\000\000\000\ -\000\000\139\001\000\000\000\000\000\000\000\000\034\001\000\000\ -\000\000\000\000\139\001\139\001\139\001\139\001\139\001\000\000\ -\000\000\000\000\000\000\000\000\000\000\139\001\000\000\139\001\ -\000\000\000\000\000\000\000\000\000\000\139\001\139\001\000\000\ -\000\000\000\000\000\000\139\001\000\000\139\001\000\000\000\000\ -\000\000\000\000\000\000\000\000\139\001\000\000\139\001\139\001\ -\139\001\139\001\000\000\139\001\139\001\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\139\001\139\001\139\001\139\001\139\001\139\001\ -\139\001\139\001\139\001\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\139\001\139\001\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\139\001\000\000\000\000\000\000\000\000\139\001\000\000\000\000\ -\000\000\139\001\053\000\000\000\000\000\139\001\053\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\139\001\139\001\000\000\ -\139\001\139\001\000\000\139\001\053\000\053\000\000\000\000\000\ -\000\000\000\000\000\000\053\000\000\000\000\000\053\000\000\000\ -\000\000\000\000\000\000\053\000\000\000\000\000\000\000\000\000\ -\053\000\000\000\053\000\000\000\000\000\000\000\053\000\000\000\ -\000\000\000\000\000\000\000\000\053\000\000\000\000\000\000\000\ -\053\000\053\000\053\000\053\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\053\000\000\000\053\000\000\000\000\000\000\000\ -\000\000\000\000\053\000\053\000\000\000\000\000\000\000\000\000\ -\053\000\000\000\053\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\053\000\000\000\053\000\053\000\053\000\053\000\000\000\ +\000\000\062\001\062\001\062\001\000\000\062\001\000\000\081\001\ +\081\001\000\000\081\001\081\001\062\001\081\001\062\001\000\000\ +\000\000\000\000\000\000\062\001\000\000\000\000\000\000\000\000\ +\000\000\000\000\062\001\000\000\000\000\000\000\053\000\000\000\ +\000\000\000\000\062\001\000\000\062\001\062\001\000\000\000\000\ +\000\000\062\001\000\000\000\000\000\000\062\001\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\062\001\062\001\ +\062\001\062\001\062\001\000\000\000\000\000\000\000\000\000\000\ +\000\000\062\001\000\000\062\001\000\000\000\000\000\000\000\000\ +\000\000\062\001\062\001\000\000\000\000\000\000\062\001\062\001\ +\000\000\062\001\000\000\000\000\000\000\000\000\000\000\000\000\ +\062\001\000\000\062\001\062\001\062\001\062\001\000\000\062\001\ +\062\001\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\062\001\062\001\ +\062\001\062\001\062\001\062\001\062\001\062\001\062\001\000\000\ +\000\000\000\000\000\000\000\000\000\000\062\001\062\001\000\000\ +\000\000\062\001\062\001\000\000\000\000\062\001\062\001\062\001\ +\000\000\000\000\000\000\000\000\062\001\000\000\000\000\062\001\ +\000\000\062\001\000\000\062\001\000\000\062\001\000\000\224\001\ +\224\001\062\001\000\000\224\001\000\000\000\000\000\000\000\000\ +\000\000\062\001\062\001\224\001\062\001\062\001\224\001\062\001\ +\000\000\224\001\000\000\000\000\000\000\000\000\000\000\000\000\ +\224\001\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\224\001\000\000\000\000\000\000\036\001\224\001\000\000\224\001\ +\000\000\000\000\000\000\224\001\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\224\001\224\001\224\001\224\001\ +\224\001\000\000\000\000\000\000\000\000\000\000\000\000\224\001\ +\000\000\224\001\000\000\000\000\000\000\000\000\000\000\224\001\ +\224\001\000\000\000\000\000\000\000\000\224\001\000\000\224\001\ +\000\000\000\000\000\000\000\000\000\000\000\000\224\001\000\000\ +\224\001\224\001\224\001\224\001\000\000\224\001\224\001\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\224\001\224\001\224\001\224\001\ +\224\001\224\001\224\001\224\001\224\001\000\000\000\000\000\000\ +\000\000\000\000\224\001\224\001\224\001\000\000\000\000\224\001\ +\224\001\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\053\000\000\000\224\001\000\000\053\000\000\000\000\000\224\001\ +\000\000\000\000\000\000\224\001\000\000\000\000\000\000\224\001\ +\000\000\000\000\053\000\053\000\000\000\000\000\000\000\224\001\ +\224\001\053\000\224\001\224\001\053\000\224\001\000\000\000\000\ +\000\000\053\000\000\000\000\000\000\000\000\000\053\000\000\000\ +\053\000\000\000\000\000\000\000\053\000\000\000\000\000\000\000\ +\000\000\000\000\053\000\000\000\000\000\000\000\053\000\053\000\ \053\000\053\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\053\000\ -\053\000\053\000\053\000\053\000\053\000\053\000\053\000\053\000\ +\053\000\000\000\053\000\000\000\000\000\000\000\000\000\000\000\ +\053\000\053\000\000\000\000\000\000\000\000\000\053\000\000\000\ +\053\000\000\000\000\000\000\000\000\000\000\000\000\000\053\000\ +\000\000\053\000\053\000\053\000\053\000\000\000\053\000\053\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\053\000\053\000\053\000\ +\053\000\053\000\053\000\053\000\053\000\053\000\000\000\000\000\ \000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\053\000\053\000\000\000\000\000\000\000\000\000\ +\053\000\053\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\053\000\000\000\000\000\000\000\000\000\ +\053\000\000\000\000\000\000\000\053\000\000\000\000\000\000\000\ +\053\000\000\000\000\000\000\000\000\000\000\000\000\000\036\001\ +\053\000\053\000\036\001\053\000\053\000\000\000\053\000\036\001\ +\021\000\036\001\000\000\000\000\000\000\000\000\000\000\036\001\ +\036\001\036\001\036\001\000\000\036\001\000\000\000\000\036\001\ +\000\000\036\001\000\000\000\000\000\000\000\000\000\000\036\001\ +\000\000\000\000\000\000\000\000\036\001\000\000\036\001\000\000\ +\036\001\000\000\036\001\000\000\000\000\036\001\036\001\000\000\ +\000\000\000\000\000\000\036\001\036\001\036\001\036\001\036\001\ +\036\001\036\001\036\001\036\001\036\001\000\000\036\001\000\000\ +\036\001\036\001\036\001\000\000\000\000\036\001\036\001\036\001\ +\036\001\000\000\036\001\036\001\036\001\000\000\036\001\000\000\ +\036\001\036\001\000\000\000\000\000\000\036\001\036\001\036\001\ +\036\001\036\001\036\001\000\000\036\001\036\001\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\036\001\036\001\036\001\036\001\036\001\ +\036\001\036\001\036\001\036\001\000\000\000\000\000\000\000\000\ +\036\001\036\001\036\001\036\001\036\001\000\000\036\001\036\001\ +\036\001\000\000\025\000\000\000\036\001\036\001\000\000\036\001\ +\036\001\036\001\036\001\036\001\000\000\036\001\036\001\000\000\ +\000\000\000\000\036\001\036\001\053\000\036\001\036\001\000\000\ +\053\000\036\001\036\001\036\001\036\001\036\001\036\001\036\001\ +\000\000\036\001\036\001\036\001\000\000\000\000\053\000\053\000\ \000\000\000\000\000\000\000\000\000\000\053\000\000\000\000\000\ -\000\000\000\000\053\000\000\000\000\000\000\000\053\000\000\000\ -\000\000\000\000\053\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\034\001\053\000\053\000\034\001\053\000\053\000\000\000\ -\053\000\034\001\021\000\034\001\000\000\000\000\000\000\000\000\ -\000\000\034\001\034\001\034\001\034\001\000\000\034\001\000\000\ -\000\000\034\001\000\000\034\001\000\000\000\000\000\000\000\000\ -\000\000\034\001\000\000\000\000\000\000\000\000\034\001\000\000\ -\034\001\000\000\034\001\000\000\034\001\000\000\000\000\034\001\ -\034\001\000\000\000\000\000\000\000\000\034\001\034\001\034\001\ -\034\001\034\001\034\001\034\001\034\001\034\001\034\001\000\000\ -\034\001\000\000\034\001\034\001\034\001\000\000\000\000\034\001\ -\034\001\034\001\034\001\000\000\034\001\034\001\034\001\000\000\ -\034\001\000\000\034\001\034\001\000\000\000\000\000\000\034\001\ -\034\001\034\001\034\001\034\001\034\001\000\000\034\001\034\001\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\034\001\034\001\034\001\ -\034\001\034\001\034\001\034\001\034\001\034\001\000\000\000\000\ -\000\000\000\000\034\001\034\001\034\001\034\001\034\001\000\000\ -\034\001\034\001\034\001\000\000\025\000\000\000\034\001\034\001\ -\000\000\034\001\034\001\034\001\034\001\034\001\000\000\034\001\ -\034\001\000\000\000\000\000\000\034\001\034\001\053\000\034\001\ -\034\001\000\000\053\000\034\001\034\001\034\001\034\001\034\001\ -\034\001\034\001\000\000\034\001\034\001\034\001\000\000\000\000\ -\053\000\053\000\000\000\000\000\000\000\000\000\000\000\053\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\053\000\ -\000\000\000\000\000\000\000\000\053\000\000\000\053\000\000\000\ -\000\000\000\000\053\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\053\000\053\000\053\000\053\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\053\000\000\000\ -\053\000\000\000\000\000\000\000\000\000\000\000\053\000\053\000\ +\000\000\000\000\000\000\000\000\000\000\053\000\000\000\000\000\ +\000\000\000\000\053\000\000\000\053\000\000\000\000\000\000\000\ +\053\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\053\000\053\000\053\000\053\000\000\000\000\000\ \000\000\000\000\000\000\000\000\053\000\000\000\053\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\053\000\000\000\053\000\ -\053\000\053\000\053\000\000\000\053\000\053\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\003\000\000\000\ -\000\000\000\000\000\000\053\000\053\000\053\000\053\000\053\000\ -\053\000\053\000\053\000\053\000\021\000\000\000\000\000\000\000\ -\021\000\000\000\000\000\000\000\000\000\000\000\053\000\053\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\021\000\000\000\ -\000\000\053\000\000\000\000\000\000\000\021\000\053\000\000\000\ -\000\000\000\000\053\000\000\000\000\000\021\000\053\000\000\000\ -\000\000\000\000\021\000\000\000\021\000\000\000\053\000\053\000\ -\021\000\053\000\053\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\021\000\021\000\021\000\021\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\021\000\000\000\021\000\000\000\ -\000\000\000\000\000\000\000\000\021\000\021\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\053\000\053\000\000\000\000\000\ +\000\000\000\000\053\000\000\000\053\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\053\000\000\000\053\000\053\000\053\000\ +\053\000\000\000\053\000\053\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\003\000\000\000\000\000\000\000\ +\000\000\053\000\053\000\053\000\053\000\053\000\053\000\053\000\ +\053\000\053\000\021\000\000\000\000\000\000\000\021\000\000\000\ +\000\000\000\000\000\000\000\000\053\000\053\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\021\000\000\000\000\000\053\000\ +\000\000\000\000\000\000\021\000\053\000\000\000\000\000\000\000\ +\053\000\000\000\000\000\021\000\053\000\000\000\000\000\000\000\ +\021\000\000\000\021\000\000\000\053\000\053\000\021\000\053\000\ +\053\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\021\000\021\000\021\000\021\000\000\000\000\000\000\000\000\000\ \000\000\000\000\021\000\000\000\021\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\021\000\163\000\021\000\021\000\021\000\ -\021\000\000\000\021\000\021\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\021\000\021\000\021\000\021\000\021\000\021\000\021\000\ -\021\000\021\000\000\000\000\000\000\000\000\000\025\000\000\000\ -\000\000\000\000\025\000\000\000\021\000\021\000\000\000\000\000\ -\000\000\025\000\000\000\000\000\000\000\000\000\000\000\021\000\ -\025\000\000\000\000\000\000\000\021\000\000\000\000\000\025\000\ -\021\000\000\000\000\000\000\000\021\000\000\000\000\000\025\000\ -\000\000\000\000\000\000\000\000\021\000\021\000\025\000\021\000\ -\021\000\000\000\025\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\025\000\025\000\025\000\025\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\025\000\000\000\ -\025\000\064\000\000\000\000\000\000\000\000\000\025\000\025\000\ -\000\000\000\000\000\000\000\000\025\000\000\000\025\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\025\000\000\000\025\000\ -\025\000\025\000\025\000\000\000\025\000\025\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\025\000\025\000\025\000\025\000\025\000\ -\025\000\025\000\025\000\025\000\000\000\000\000\000\000\000\000\ -\003\000\000\000\000\000\000\000\003\000\000\000\025\000\025\000\ -\000\000\000\000\000\000\003\000\000\000\000\000\000\000\000\000\ -\000\000\025\000\003\000\000\000\000\000\000\000\025\000\000\000\ -\000\000\003\000\025\000\000\000\000\000\000\000\025\000\000\000\ -\000\000\003\000\000\000\000\000\000\000\000\000\025\000\025\000\ -\003\000\025\000\025\000\000\000\003\000\000\000\033\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\003\000\003\000\ -\003\000\003\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\003\000\000\000\003\000\000\000\000\000\000\000\000\000\000\000\ -\003\000\003\000\000\000\000\000\000\000\000\000\003\000\000\000\ -\003\000\000\000\000\000\000\000\000\000\000\000\000\000\003\000\ -\000\000\003\000\003\000\003\000\003\000\000\000\003\000\003\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\163\000\000\000\ -\000\000\000\000\163\000\000\000\000\000\003\000\003\000\003\000\ -\003\000\003\000\003\000\003\000\003\000\003\000\000\000\000\000\ -\163\000\000\000\000\000\000\000\000\000\000\000\000\000\163\000\ -\003\000\003\000\000\000\000\000\000\000\000\000\000\000\163\000\ -\000\000\000\000\000\000\003\000\000\000\000\000\163\000\000\000\ -\003\000\000\000\163\000\001\000\003\000\000\000\000\000\000\000\ -\003\000\000\000\000\000\000\000\163\000\163\000\163\000\163\000\ -\003\000\003\000\000\000\003\000\003\000\000\000\163\000\000\000\ -\163\000\000\000\000\000\000\000\000\000\000\000\163\000\163\000\ -\000\000\000\000\000\000\000\000\163\000\000\000\163\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\163\000\000\000\163\000\ -\163\000\163\000\163\000\000\000\163\000\163\000\000\000\000\000\ -\000\000\000\000\000\000\064\000\000\000\000\000\000\000\064\000\ -\000\000\000\000\000\000\163\000\163\000\163\000\163\000\163\000\ -\163\000\163\000\163\000\163\000\000\000\064\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\064\000\000\000\163\000\163\000\ -\000\000\000\000\000\000\000\000\064\000\000\000\000\000\000\000\ -\000\000\163\000\000\000\064\000\000\000\000\000\163\000\064\000\ -\034\001\000\000\163\000\000\000\000\000\000\000\163\000\000\000\ -\000\000\064\000\064\000\064\000\064\000\000\000\163\000\163\000\ -\000\000\163\000\163\000\064\000\000\000\064\000\000\000\000\000\ -\000\000\000\000\000\000\064\000\064\000\000\000\000\000\000\000\ -\000\000\064\000\000\000\064\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\064\000\000\000\064\000\064\000\064\000\064\000\ -\000\000\064\000\064\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\033\000\000\000\000\000\000\000\ -\064\000\064\000\064\000\064\000\064\000\064\000\064\000\064\000\ -\064\000\000\000\033\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\033\000\000\000\064\000\064\000\000\000\000\000\000\000\ -\000\000\033\000\000\000\000\000\000\000\000\000\064\000\000\000\ -\033\000\000\000\000\000\064\000\033\000\094\000\000\000\064\000\ -\000\000\000\000\000\000\064\000\000\000\000\000\033\000\033\000\ -\033\000\033\000\000\000\064\000\064\000\000\000\064\000\064\000\ -\033\000\000\000\033\000\000\000\000\000\000\000\000\000\000\000\ -\033\000\033\000\000\000\000\000\000\000\000\000\033\000\000\000\ +\000\000\000\000\021\000\021\000\000\000\000\000\000\000\000\000\ +\021\000\000\000\021\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\021\000\165\000\021\000\021\000\021\000\021\000\000\000\ +\021\000\021\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\021\000\ +\021\000\021\000\021\000\021\000\021\000\021\000\021\000\021\000\ +\000\000\000\000\000\000\000\000\025\000\000\000\000\000\000\000\ +\025\000\000\000\021\000\021\000\000\000\000\000\000\000\025\000\ +\000\000\000\000\000\000\000\000\000\000\021\000\025\000\000\000\ +\000\000\000\000\021\000\000\000\000\000\025\000\021\000\000\000\ +\000\000\000\000\021\000\000\000\000\000\025\000\000\000\000\000\ +\000\000\000\000\021\000\021\000\025\000\021\000\021\000\000\000\ +\025\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\025\000\025\000\025\000\025\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\025\000\000\000\025\000\064\000\ +\000\000\000\000\000\000\000\000\025\000\025\000\000\000\000\000\ +\000\000\000\000\025\000\000\000\025\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\025\000\000\000\025\000\025\000\025\000\ +\025\000\000\000\025\000\025\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\025\000\025\000\025\000\025\000\025\000\025\000\025\000\ +\025\000\025\000\000\000\000\000\000\000\000\000\003\000\000\000\ +\000\000\000\000\003\000\000\000\025\000\025\000\000\000\000\000\ +\000\000\003\000\000\000\000\000\000\000\000\000\000\000\025\000\ +\003\000\000\000\000\000\000\000\025\000\000\000\000\000\003\000\ +\025\000\000\000\000\000\000\000\025\000\000\000\000\000\003\000\ +\000\000\000\000\000\000\000\000\025\000\025\000\003\000\025\000\ +\025\000\000\000\003\000\000\000\033\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\003\000\003\000\003\000\003\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\003\000\000\000\ +\003\000\000\000\000\000\000\000\000\000\000\000\003\000\003\000\ +\000\000\000\000\000\000\000\000\003\000\000\000\003\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\003\000\000\000\003\000\ +\003\000\003\000\003\000\000\000\003\000\003\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\165\000\000\000\000\000\000\000\ +\165\000\000\000\000\000\003\000\003\000\003\000\003\000\003\000\ +\003\000\003\000\003\000\003\000\000\000\000\000\165\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\165\000\003\000\003\000\ +\000\000\000\000\000\000\000\000\000\000\165\000\000\000\000\000\ +\000\000\003\000\000\000\000\000\165\000\000\000\003\000\000\000\ +\165\000\001\000\003\000\000\000\000\000\000\000\003\000\000\000\ +\000\000\000\000\165\000\165\000\165\000\165\000\003\000\003\000\ +\000\000\003\000\003\000\000\000\165\000\000\000\165\000\000\000\ +\000\000\000\000\000\000\000\000\165\000\165\000\000\000\000\000\ +\000\000\000\000\165\000\000\000\165\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\165\000\000\000\165\000\165\000\165\000\ +\165\000\000\000\165\000\165\000\000\000\000\000\000\000\000\000\ +\000\000\064\000\000\000\000\000\000\000\064\000\000\000\000\000\ +\000\000\165\000\165\000\165\000\165\000\165\000\165\000\165\000\ +\165\000\165\000\000\000\064\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\064\000\000\000\165\000\165\000\000\000\000\000\ +\000\000\000\000\064\000\000\000\000\000\000\000\000\000\165\000\ +\000\000\064\000\000\000\000\000\165\000\064\000\036\001\000\000\ +\165\000\000\000\000\000\000\000\165\000\000\000\000\000\064\000\ +\064\000\064\000\064\000\000\000\165\000\165\000\000\000\165\000\ +\165\000\064\000\000\000\064\000\000\000\000\000\000\000\000\000\ +\000\000\064\000\064\000\000\000\000\000\000\000\000\000\064\000\ +\000\000\064\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\064\000\000\000\064\000\064\000\064\000\064\000\000\000\064\000\ +\064\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\033\000\000\000\000\000\000\000\064\000\064\000\ +\064\000\064\000\064\000\064\000\064\000\064\000\064\000\000\000\ \033\000\000\000\000\000\000\000\000\000\000\000\000\000\033\000\ -\000\000\033\000\033\000\033\000\033\000\000\000\033\000\033\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\001\000\000\000\000\000\000\000\033\000\033\000\033\000\ -\033\000\033\000\033\000\033\000\033\000\033\000\000\000\001\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\001\000\000\000\ -\033\000\033\000\000\000\000\000\000\000\000\000\001\000\000\000\ -\000\000\000\000\000\000\033\000\000\000\001\000\000\000\000\000\ -\033\000\001\000\092\000\000\000\033\000\000\000\000\000\000\000\ -\033\000\000\000\000\000\001\000\001\000\001\000\001\000\000\000\ -\033\000\033\000\000\000\033\000\033\000\001\000\000\000\001\000\ -\000\000\000\000\000\000\000\000\000\000\001\000\001\000\000\000\ -\000\000\000\000\000\000\001\000\000\000\001\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\001\000\000\000\001\000\001\000\ -\001\000\001\000\000\000\001\000\001\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\034\001\000\000\ -\000\000\000\000\001\000\001\000\001\000\001\000\001\000\001\000\ -\001\000\001\000\001\000\000\000\034\001\000\000\000\000\000\000\ -\000\000\000\000\000\000\034\001\000\000\001\000\001\000\000\000\ -\000\000\000\000\000\000\034\001\000\000\000\000\000\000\000\000\ -\001\000\000\000\034\001\000\000\000\000\001\000\034\001\000\000\ -\000\000\001\000\000\000\000\000\000\000\001\000\000\000\000\000\ -\034\001\034\001\034\001\034\001\000\000\001\000\001\000\000\000\ -\001\000\001\000\034\001\000\000\034\001\000\000\000\000\000\000\ -\000\000\000\000\034\001\034\001\000\000\000\000\000\000\000\000\ -\034\001\000\000\034\001\000\000\000\000\000\000\000\000\000\000\ -\000\000\034\001\000\000\034\001\034\001\034\001\034\001\000\000\ -\034\001\034\001\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\094\000\000\000\000\000\000\000\034\001\ -\034\001\034\001\034\001\034\001\034\001\034\001\034\001\034\001\ -\000\000\094\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\094\000\000\000\034\001\034\001\000\000\000\000\000\000\000\000\ -\094\000\000\000\000\000\000\000\000\000\034\001\000\000\094\000\ -\000\000\000\000\034\001\094\000\000\000\000\000\034\001\000\000\ -\000\000\000\000\034\001\000\000\000\000\094\000\094\000\094\000\ -\094\000\000\000\034\001\034\001\000\000\034\001\034\001\094\000\ -\000\000\094\000\000\000\000\000\000\000\000\000\000\000\094\000\ -\094\000\000\000\000\000\000\000\000\000\094\000\000\000\094\000\ +\000\000\064\000\064\000\000\000\000\000\000\000\000\000\033\000\ +\000\000\000\000\000\000\000\000\064\000\000\000\033\000\000\000\ +\000\000\064\000\033\000\094\000\000\000\064\000\000\000\000\000\ +\000\000\064\000\000\000\000\000\033\000\033\000\033\000\033\000\ +\000\000\064\000\064\000\000\000\064\000\064\000\033\000\000\000\ +\033\000\000\000\000\000\000\000\000\000\000\000\033\000\033\000\ +\000\000\000\000\000\000\000\000\033\000\000\000\033\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\033\000\000\000\033\000\ +\033\000\033\000\033\000\000\000\033\000\033\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\001\000\ +\000\000\000\000\000\000\033\000\033\000\033\000\033\000\033\000\ +\033\000\033\000\033\000\033\000\000\000\001\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\001\000\000\000\033\000\033\000\ +\000\000\000\000\000\000\000\000\001\000\000\000\000\000\000\000\ +\000\000\033\000\000\000\001\000\000\000\000\000\033\000\001\000\ +\092\000\000\000\033\000\000\000\000\000\000\000\033\000\000\000\ +\000\000\001\000\001\000\001\000\001\000\000\000\033\000\033\000\ +\000\000\033\000\033\000\001\000\000\000\001\000\000\000\000\000\ +\000\000\000\000\000\000\001\000\001\000\000\000\000\000\000\000\ +\000\000\001\000\000\000\001\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\001\000\000\000\001\000\001\000\001\000\001\000\ +\000\000\001\000\001\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\036\001\000\000\000\000\000\000\ +\001\000\001\000\001\000\001\000\001\000\001\000\001\000\001\000\ +\001\000\000\000\036\001\000\000\000\000\000\000\000\000\000\000\ +\000\000\036\001\000\000\001\000\001\000\000\000\000\000\000\000\ +\000\000\036\001\000\000\000\000\000\000\000\000\001\000\000\000\ +\036\001\000\000\000\000\001\000\036\001\000\000\000\000\001\000\ +\000\000\000\000\000\000\001\000\000\000\000\000\036\001\036\001\ +\036\001\036\001\000\000\001\000\001\000\000\000\001\000\001\000\ +\036\001\000\000\036\001\000\000\000\000\000\000\000\000\000\000\ +\036\001\036\001\000\000\000\000\000\000\000\000\036\001\000\000\ +\036\001\000\000\000\000\000\000\000\000\000\000\000\000\036\001\ +\000\000\036\001\036\001\036\001\036\001\000\000\036\001\036\001\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\094\000\000\000\000\000\000\000\036\001\036\001\036\001\ +\036\001\036\001\036\001\036\001\036\001\036\001\000\000\094\000\ \000\000\000\000\000\000\000\000\000\000\000\000\094\000\000\000\ -\094\000\094\000\094\000\094\000\000\000\094\000\094\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\092\000\000\000\000\000\000\000\094\000\094\000\094\000\094\000\ -\094\000\094\000\094\000\094\000\094\000\000\000\092\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\092\000\000\000\094\000\ -\094\000\000\000\000\000\000\000\000\000\092\000\000\000\000\000\ -\000\000\000\000\094\000\000\000\092\000\000\000\000\000\094\000\ -\092\000\000\000\000\000\094\000\000\000\000\000\000\000\094\000\ -\000\000\000\000\092\000\092\000\092\000\092\000\000\000\094\000\ -\094\000\000\000\094\000\094\000\092\000\000\000\092\000\000\000\ -\000\000\000\000\000\000\000\000\092\000\092\000\000\000\000\000\ -\000\000\000\000\092\000\000\000\092\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\092\000\000\000\092\000\092\000\092\000\ -\092\000\000\000\092\000\092\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\196\000\000\000\000\000\ -\000\000\092\000\092\000\092\000\092\000\092\000\092\000\092\000\ -\092\000\092\000\000\000\196\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\196\000\000\000\092\000\092\000\000\000\000\000\ -\000\000\000\000\196\000\000\000\000\000\000\000\000\000\092\000\ -\000\000\196\000\000\000\000\000\092\000\196\000\000\000\000\000\ -\092\000\000\000\000\000\000\000\092\000\000\000\000\000\196\000\ -\196\000\196\000\196\000\000\000\092\000\092\000\000\000\092\000\ -\092\000\196\000\000\000\196\000\000\000\000\000\000\000\000\000\ -\000\000\196\000\196\000\000\000\000\000\000\000\000\000\196\000\ -\000\000\196\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\196\000\000\000\196\000\196\000\196\000\196\000\000\000\196\000\ -\196\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\196\000\196\000\ -\196\000\196\000\196\000\196\000\196\000\196\000\196\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\196\000\196\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\196\000\000\000\000\000\000\000\ -\000\000\196\000\000\000\000\000\000\000\196\000\000\000\000\000\ -\000\000\196\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\196\000\196\000\000\000\196\000\196\000\013\001\000\000\ -\031\000\032\000\033\000\034\000\035\000\036\000\037\000\038\000\ -\000\000\000\000\000\000\039\000\000\000\040\000\041\000\000\000\ -\000\000\014\001\015\001\000\000\016\001\042\000\000\000\017\001\ -\043\000\000\000\000\000\000\000\000\000\000\000\000\000\044\000\ -\000\000\000\000\045\000\018\001\000\000\000\000\046\000\047\000\ -\048\000\000\000\049\000\050\000\051\000\052\000\053\000\019\001\ -\054\000\020\001\000\000\000\000\000\000\000\000\000\000\055\000\ -\056\000\057\000\058\000\059\000\060\000\000\000\000\000\000\000\ -\061\000\062\000\000\000\000\000\063\000\010\000\011\000\064\000\ -\065\000\066\000\000\000\067\000\021\001\000\000\000\000\069\000\ -\070\000\071\000\072\000\073\000\000\000\074\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\022\001\023\001\024\001\ -\025\001\026\001\027\001\028\001\029\001\030\001\031\001\032\001\ -\076\000\033\001\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\077\000\000\000\078\000\ -\000\000\000\000\000\000\079\000\000\000\000\000\000\000\080\000\ -\081\000\000\000\000\000\000\000\000\000\000\000\082\000\083\000\ -\000\000\000\000\084\000\000\000\000\000\000\000\085\000\000\000\ -\034\001\000\000\087\000\088\000\089\000\000\000\000\000\000\000\ -\090\000\091\000\092\000\093\000\094\000\242\001\000\000\031\000\ -\032\000\033\000\034\000\035\000\036\000\037\000\038\000\000\000\ -\000\000\000\000\039\000\000\000\040\000\041\000\000\000\000\000\ -\014\001\015\001\000\000\016\001\042\000\000\000\017\001\043\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\044\000\000\000\ -\000\000\045\000\018\001\000\000\000\000\046\000\047\000\048\000\ -\000\000\049\000\050\000\051\000\052\000\053\000\019\001\054\000\ -\020\001\000\000\000\000\000\000\000\000\000\000\055\000\056\000\ -\057\000\058\000\059\000\060\000\000\000\000\000\000\000\061\000\ -\062\000\000\000\000\000\063\000\010\000\011\000\064\000\065\000\ -\066\000\000\000\067\000\243\001\000\000\000\000\069\000\070\000\ -\071\000\072\000\073\000\000\000\074\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\022\001\023\001\024\001\025\001\ -\026\001\027\001\028\001\029\001\244\001\031\001\032\001\076\000\ -\033\001\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\077\000\000\000\078\000\000\000\ -\000\000\000\000\079\000\000\000\000\000\000\000\080\000\081\000\ -\000\000\000\000\000\000\000\000\000\000\082\000\083\000\000\000\ -\000\000\084\000\000\000\000\000\000\000\085\000\000\000\245\001\ -\000\000\087\000\088\000\089\000\000\000\000\000\000\000\090\000\ -\091\000\092\000\093\000\094\000\086\004\000\000\031\000\032\000\ -\033\000\034\000\035\000\036\000\037\000\038\000\000\000\000\000\ -\000\000\039\000\000\000\040\000\041\000\000\000\000\000\014\001\ -\015\001\000\000\016\001\042\000\000\000\017\001\043\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\044\000\000\000\000\000\ -\045\000\018\001\000\000\000\000\046\000\047\000\048\000\000\000\ -\049\000\050\000\051\000\052\000\053\000\019\001\054\000\020\001\ -\000\000\000\000\000\000\000\000\000\000\055\000\056\000\057\000\ -\058\000\059\000\060\000\000\000\000\000\000\000\061\000\062\000\ -\000\000\000\000\063\000\010\000\011\000\064\000\065\000\066\000\ -\000\000\067\000\087\004\000\000\000\000\069\000\070\000\071\000\ -\072\000\073\000\000\000\074\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\022\001\023\001\024\001\025\001\026\001\ -\027\001\028\001\029\001\088\004\031\001\032\001\076\000\033\001\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\077\000\000\000\078\000\000\000\000\000\ -\000\000\079\000\000\000\000\000\000\000\080\000\081\000\000\000\ -\000\000\000\000\000\000\000\000\082\000\083\000\000\000\000\000\ -\084\000\000\000\000\000\210\001\085\000\000\000\089\004\000\000\ -\087\000\088\000\089\000\217\000\000\000\000\000\090\000\091\000\ -\092\000\093\000\094\000\041\000\000\000\000\000\014\001\015\001\ -\000\000\016\001\000\000\000\000\017\001\043\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\018\001\000\000\000\000\000\000\047\000\000\000\000\000\000\000\ -\000\000\218\000\141\000\000\000\019\001\000\000\020\001\000\000\ -\000\000\000\000\000\000\000\000\055\000\056\000\057\000\058\000\ -\059\000\000\000\000\000\000\000\000\000\219\000\000\000\000\000\ -\000\000\221\000\000\000\000\000\064\000\000\000\222\000\000\000\ -\000\000\211\001\000\000\000\000\237\000\070\000\000\000\000\000\ -\000\000\000\000\074\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\022\001\023\001\024\001\025\001\026\001\027\001\ -\028\001\029\001\212\001\031\001\032\001\000\000\033\001\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\078\000\000\000\000\000\000\000\ -\079\000\000\000\000\000\000\000\080\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\082\000\083\000\000\000\000\000\084\000\ -\000\000\000\000\103\003\000\000\000\000\213\001\000\000\087\000\ -\000\000\089\000\217\000\000\000\000\000\090\000\091\000\092\000\ -\093\000\224\000\041\000\000\000\000\000\014\001\015\001\000\000\ -\016\001\000\000\000\000\017\001\043\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\018\001\ -\000\000\000\000\000\000\047\000\000\000\000\000\000\000\000\000\ -\218\000\141\000\000\000\019\001\000\000\020\001\000\000\000\000\ -\000\000\000\000\000\000\055\000\056\000\057\000\058\000\059\000\ -\000\000\000\000\000\000\000\000\219\000\000\000\000\000\000\000\ -\221\000\000\000\000\000\064\000\000\000\222\000\000\000\000\000\ -\104\003\000\000\000\000\237\000\070\000\000\000\000\000\000\000\ -\000\000\074\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\022\001\023\001\024\001\025\001\026\001\027\001\028\001\ -\029\001\105\003\031\001\032\001\000\000\033\001\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\078\000\000\000\000\000\000\000\079\000\ -\000\000\000\000\000\000\080\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\082\000\083\000\000\000\000\000\084\000\000\000\ -\000\000\000\000\000\000\000\000\106\003\000\000\087\000\000\000\ -\089\000\000\000\000\000\000\000\090\000\091\000\092\000\093\000\ -\224\000\031\000\032\000\033\000\034\000\035\000\036\000\037\000\ -\038\000\000\000\000\000\000\000\039\000\000\000\040\000\041\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\042\000\000\000\ -\000\000\043\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\044\000\000\000\000\000\045\000\000\000\000\000\000\000\046\000\ -\047\000\048\000\000\000\049\000\050\000\051\000\052\000\053\000\ -\000\000\054\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\055\000\056\000\057\000\058\000\059\000\060\000\000\000\000\000\ -\000\000\061\000\062\000\000\000\000\000\063\000\010\000\011\000\ -\064\000\065\000\066\000\000\000\067\000\068\000\000\000\000\000\ -\069\000\070\000\071\000\072\000\073\000\000\000\074\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\075\000\000\000\ -\000\000\076\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\077\000\000\000\ -\078\000\000\000\000\000\000\000\079\000\000\000\000\000\000\000\ -\080\000\081\000\000\000\000\000\000\000\000\000\000\000\082\000\ -\083\000\000\000\000\000\084\000\000\000\000\000\000\000\085\000\ -\000\000\086\000\000\000\087\000\088\000\089\000\000\000\000\000\ -\000\000\090\000\091\000\092\000\093\000\094\000\110\001\110\001\ -\110\001\110\001\110\001\110\001\110\001\110\001\000\000\000\000\ -\000\000\110\001\000\000\110\001\110\001\000\000\000\000\000\000\ -\000\000\000\000\000\000\110\001\000\000\000\000\110\001\000\000\ -\000\000\000\000\000\000\000\000\000\000\110\001\000\000\000\000\ -\110\001\000\000\000\000\000\000\110\001\110\001\110\001\000\000\ -\110\001\110\001\110\001\110\001\110\001\000\000\110\001\000\000\ -\000\000\000\000\000\000\000\000\000\000\110\001\110\001\110\001\ -\110\001\110\001\110\001\000\000\000\000\000\000\110\001\110\001\ -\000\000\000\000\110\001\110\001\110\001\110\001\110\001\110\001\ -\000\000\110\001\110\001\000\000\000\000\110\001\110\001\110\001\ -\110\001\110\001\000\000\110\001\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\110\001\000\000\000\000\110\001\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\110\001\000\000\110\001\000\000\000\000\ -\000\000\110\001\000\000\000\000\000\000\110\001\110\001\000\000\ -\000\000\000\000\000\000\000\000\110\001\110\001\000\000\000\000\ -\110\001\000\000\000\000\000\000\110\001\000\000\110\001\000\000\ -\110\001\110\001\110\001\000\000\000\000\000\000\110\001\110\001\ -\110\001\110\001\110\001\032\001\000\000\000\000\000\000\032\001\ -\000\000\000\000\032\001\000\000\020\001\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\032\001\032\001\032\001\ -\032\001\032\001\000\000\000\000\000\000\000\000\000\000\032\001\ -\000\000\000\000\000\000\032\001\000\000\000\000\000\000\000\000\ -\032\001\032\001\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\032\001\032\001\032\001\032\001\032\001\ -\000\000\000\000\000\000\000\000\032\001\000\000\000\000\000\000\ -\032\001\000\000\000\000\000\000\032\001\032\001\000\000\000\000\ -\000\000\000\000\000\000\032\001\032\001\000\000\000\000\000\000\ -\000\000\032\001\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\032\001\000\000\000\000\032\001\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\032\001\000\000\032\001\000\000\032\001\ -\000\000\000\000\000\000\032\001\000\000\000\000\000\000\000\000\ -\032\001\000\000\032\001\032\001\000\000\000\000\032\001\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\032\001\000\000\ -\032\001\000\000\000\000\032\001\032\001\032\001\032\001\032\001\ -\032\001\027\000\000\000\032\001\000\000\027\000\000\000\032\001\ +\036\001\036\001\000\000\000\000\000\000\000\000\094\000\000\000\ +\000\000\000\000\000\000\036\001\000\000\094\000\000\000\000\000\ +\036\001\094\000\000\000\000\000\036\001\000\000\000\000\000\000\ +\036\001\000\000\000\000\094\000\094\000\094\000\094\000\000\000\ +\036\001\036\001\000\000\036\001\036\001\094\000\000\000\094\000\ +\000\000\000\000\000\000\000\000\000\000\094\000\094\000\000\000\ +\000\000\000\000\000\000\094\000\000\000\094\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\094\000\000\000\094\000\094\000\ +\094\000\094\000\000\000\094\000\094\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\092\000\000\000\ +\000\000\000\000\094\000\094\000\094\000\094\000\094\000\094\000\ +\094\000\094\000\094\000\000\000\092\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\092\000\000\000\094\000\094\000\000\000\ +\000\000\000\000\000\000\092\000\000\000\000\000\000\000\000\000\ +\094\000\000\000\092\000\000\000\000\000\094\000\092\000\000\000\ +\000\000\094\000\000\000\000\000\000\000\094\000\000\000\000\000\ +\092\000\092\000\092\000\092\000\000\000\094\000\094\000\000\000\ +\094\000\094\000\092\000\000\000\092\000\000\000\000\000\000\000\ +\000\000\000\000\092\000\092\000\000\000\000\000\000\000\000\000\ +\092\000\000\000\092\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\092\000\000\000\092\000\092\000\092\000\092\000\000\000\ +\092\000\092\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\092\000\ +\092\000\092\000\092\000\092\000\092\000\092\000\092\000\092\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\092\000\092\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\092\000\000\000\000\000\ +\000\000\000\000\092\000\000\000\000\000\000\000\092\000\000\000\ +\000\000\000\000\092\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\092\000\092\000\000\000\092\000\092\000\017\001\ +\000\000\035\000\036\000\037\000\038\000\039\000\040\000\041\000\ +\042\000\000\000\000\000\000\000\043\000\000\000\044\000\045\000\ +\000\000\000\000\018\001\019\001\000\000\020\001\046\000\000\000\ +\021\001\047\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\048\000\000\000\000\000\049\000\022\001\000\000\000\000\050\000\ +\051\000\052\000\000\000\053\000\054\000\055\000\056\000\057\000\ +\023\001\058\000\024\001\000\000\000\000\000\000\000\000\000\000\ +\059\000\060\000\061\000\062\000\063\000\064\000\000\000\000\000\ +\000\000\065\000\066\000\000\000\000\000\067\000\011\000\012\000\ +\068\000\069\000\070\000\000\000\071\000\025\001\000\000\000\000\ +\073\000\074\000\075\000\076\000\077\000\000\000\078\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\026\001\027\001\ +\028\001\029\001\030\001\031\001\032\001\033\001\034\001\035\001\ +\036\001\080\000\037\001\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\081\000\000\000\ +\082\000\000\000\000\000\000\000\083\000\000\000\000\000\000\000\ +\084\000\085\000\000\000\000\000\000\000\000\000\000\000\086\000\ +\087\000\000\000\000\000\088\000\000\000\000\000\000\000\089\000\ +\000\000\038\001\000\000\091\000\092\000\093\000\000\000\000\000\ +\000\000\094\000\095\000\096\000\097\000\098\000\246\001\000\000\ +\035\000\036\000\037\000\038\000\039\000\040\000\041\000\042\000\ +\000\000\000\000\000\000\043\000\000\000\044\000\045\000\000\000\ +\000\000\018\001\019\001\000\000\020\001\046\000\000\000\021\001\ +\047\000\000\000\000\000\000\000\000\000\000\000\000\000\048\000\ +\000\000\000\000\049\000\022\001\000\000\000\000\050\000\051\000\ +\052\000\000\000\053\000\054\000\055\000\056\000\057\000\023\001\ +\058\000\024\001\000\000\000\000\000\000\000\000\000\000\059\000\ +\060\000\061\000\062\000\063\000\064\000\000\000\000\000\000\000\ +\065\000\066\000\000\000\000\000\067\000\011\000\012\000\068\000\ +\069\000\070\000\000\000\071\000\247\001\000\000\000\000\073\000\ +\074\000\075\000\076\000\077\000\000\000\078\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\026\001\027\001\028\001\ +\029\001\030\001\031\001\032\001\033\001\248\001\035\001\036\001\ +\080\000\037\001\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\081\000\000\000\082\000\ +\000\000\000\000\000\000\083\000\000\000\000\000\000\000\084\000\ +\085\000\000\000\000\000\000\000\000\000\000\000\086\000\087\000\ +\000\000\000\000\088\000\000\000\000\000\000\000\089\000\000\000\ +\249\001\000\000\091\000\092\000\093\000\000\000\000\000\000\000\ +\094\000\095\000\096\000\097\000\098\000\090\004\000\000\035\000\ +\036\000\037\000\038\000\039\000\040\000\041\000\042\000\000\000\ +\000\000\000\000\043\000\000\000\044\000\045\000\000\000\000\000\ +\018\001\019\001\000\000\020\001\046\000\000\000\021\001\047\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\048\000\000\000\ +\000\000\049\000\022\001\000\000\000\000\050\000\051\000\052\000\ +\000\000\053\000\054\000\055\000\056\000\057\000\023\001\058\000\ +\024\001\000\000\000\000\000\000\000\000\000\000\059\000\060\000\ +\061\000\062\000\063\000\064\000\000\000\000\000\000\000\065\000\ +\066\000\000\000\000\000\067\000\011\000\012\000\068\000\069\000\ +\070\000\000\000\071\000\091\004\000\000\000\000\073\000\074\000\ +\075\000\076\000\077\000\000\000\078\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\026\001\027\001\028\001\029\001\ +\030\001\031\001\032\001\033\001\092\004\035\001\036\001\080\000\ +\037\001\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\081\000\000\000\082\000\000\000\ +\000\000\000\000\083\000\000\000\000\000\000\000\084\000\085\000\ +\000\000\000\000\000\000\000\000\000\000\086\000\087\000\000\000\ +\000\000\088\000\000\000\000\000\214\001\089\000\000\000\093\004\ +\000\000\091\000\092\000\093\000\221\000\000\000\000\000\094\000\ +\095\000\096\000\097\000\098\000\045\000\000\000\000\000\018\001\ +\019\001\000\000\020\001\000\000\000\000\021\001\047\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\022\001\000\000\000\000\000\000\051\000\000\000\000\000\ +\000\000\000\000\222\000\145\000\000\000\023\001\000\000\024\001\ +\000\000\000\000\000\000\000\000\000\000\059\000\060\000\061\000\ +\062\000\063\000\000\000\000\000\000\000\000\000\223\000\000\000\ +\000\000\000\000\225\000\000\000\000\000\068\000\000\000\226\000\ +\000\000\000\000\215\001\000\000\000\000\241\000\074\000\000\000\ +\000\000\000\000\000\000\078\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\026\001\027\001\028\001\029\001\030\001\ +\031\001\032\001\033\001\216\001\035\001\036\001\000\000\037\001\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\082\000\000\000\000\000\ +\000\000\083\000\000\000\000\000\000\000\084\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\086\000\087\000\000\000\000\000\ +\088\000\000\000\000\000\107\003\000\000\000\000\217\001\000\000\ +\091\000\000\000\093\000\221\000\000\000\000\000\094\000\095\000\ +\096\000\097\000\228\000\045\000\000\000\000\000\018\001\019\001\ +\000\000\020\001\000\000\000\000\021\001\047\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\022\001\000\000\000\000\000\000\051\000\000\000\000\000\000\000\ +\000\000\222\000\145\000\000\000\023\001\000\000\024\001\000\000\ +\000\000\000\000\000\000\000\000\059\000\060\000\061\000\062\000\ +\063\000\000\000\000\000\000\000\000\000\223\000\000\000\000\000\ +\000\000\225\000\000\000\000\000\068\000\000\000\226\000\000\000\ +\000\000\108\003\000\000\000\000\241\000\074\000\000\000\000\000\ +\000\000\000\000\078\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\026\001\027\001\028\001\029\001\030\001\031\001\ +\032\001\033\001\109\003\035\001\036\001\000\000\037\001\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\082\000\000\000\000\000\000\000\ +\083\000\000\000\000\000\000\000\084\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\086\000\087\000\000\000\000\000\088\000\ +\000\000\000\000\000\000\000\000\000\000\110\003\000\000\091\000\ +\000\000\093\000\000\000\000\000\000\000\094\000\095\000\096\000\ +\097\000\228\000\035\000\036\000\037\000\038\000\039\000\040\000\ +\041\000\042\000\000\000\000\000\000\000\043\000\000\000\044\000\ +\045\000\000\000\000\000\000\000\000\000\000\000\000\000\046\000\ +\000\000\000\000\047\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\048\000\000\000\000\000\049\000\000\000\000\000\000\000\ +\050\000\051\000\052\000\000\000\053\000\054\000\055\000\056\000\ +\057\000\000\000\058\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\059\000\060\000\061\000\062\000\063\000\064\000\000\000\ +\000\000\000\000\065\000\066\000\000\000\000\000\067\000\011\000\ +\012\000\068\000\069\000\070\000\000\000\071\000\072\000\000\000\ +\000\000\073\000\074\000\075\000\076\000\077\000\000\000\078\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\079\000\ +\000\000\000\000\080\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\081\000\ +\000\000\082\000\000\000\000\000\000\000\083\000\000\000\000\000\ +\000\000\084\000\085\000\000\000\000\000\000\000\000\000\000\000\ +\086\000\087\000\000\000\000\000\088\000\000\000\000\000\000\000\ +\089\000\000\000\090\000\000\000\091\000\092\000\093\000\000\000\ +\000\000\000\000\094\000\095\000\096\000\097\000\098\000\112\001\ +\112\001\112\001\112\001\112\001\112\001\112\001\112\001\000\000\ +\000\000\000\000\112\001\000\000\112\001\112\001\000\000\000\000\ +\000\000\000\000\000\000\000\000\112\001\000\000\000\000\112\001\ +\000\000\000\000\000\000\000\000\000\000\000\000\112\001\000\000\ +\000\000\112\001\000\000\000\000\000\000\112\001\112\001\112\001\ +\000\000\112\001\112\001\112\001\112\001\112\001\000\000\112\001\ +\000\000\000\000\000\000\000\000\000\000\000\000\112\001\112\001\ +\112\001\112\001\112\001\112\001\000\000\000\000\000\000\112\001\ +\112\001\000\000\000\000\112\001\112\001\112\001\112\001\112\001\ +\112\001\000\000\112\001\112\001\000\000\000\000\112\001\112\001\ +\112\001\112\001\112\001\000\000\112\001\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\112\001\000\000\000\000\112\001\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\112\001\000\000\112\001\000\000\ +\000\000\000\000\112\001\000\000\000\000\000\000\112\001\112\001\ +\000\000\000\000\000\000\000\000\000\000\112\001\112\001\000\000\ +\000\000\112\001\000\000\000\000\000\000\112\001\000\000\112\001\ +\000\000\112\001\112\001\112\001\000\000\000\000\000\000\112\001\ +\112\001\112\001\112\001\112\001\034\001\000\000\000\000\000\000\ +\034\001\000\000\000\000\034\001\000\000\022\001\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\034\001\034\001\ +\034\001\034\001\034\001\000\000\000\000\000\000\000\000\000\000\ +\034\001\000\000\000\000\000\000\034\001\000\000\000\000\000\000\ +\000\000\034\001\034\001\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\034\001\034\001\034\001\034\001\ +\034\001\000\000\000\000\000\000\000\000\034\001\000\000\000\000\ +\000\000\034\001\000\000\000\000\000\000\034\001\034\001\000\000\ +\000\000\000\000\000\000\000\000\034\001\034\001\000\000\000\000\ +\000\000\000\000\034\001\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\034\001\000\000\000\000\034\001\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\034\001\000\000\034\001\000\000\ +\034\001\000\000\000\000\000\000\034\001\000\000\000\000\000\000\ +\000\000\034\001\000\000\034\001\034\001\000\000\000\000\034\001\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\034\001\ +\000\000\034\001\000\000\000\000\034\001\034\001\034\001\034\001\ +\034\001\034\001\027\000\000\000\034\001\000\000\027\000\000\000\ +\034\001\027\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\027\000\027\000\027\000\027\000\ +\027\000\000\000\000\000\000\000\000\000\000\000\040\001\000\000\ +\000\000\000\000\027\000\000\000\000\000\000\000\000\000\027\000\ \027\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\027\000\027\000\027\000\027\000\027\000\ -\000\000\000\000\000\000\000\000\000\000\038\001\000\000\000\000\ -\000\000\027\000\000\000\000\000\000\000\000\000\027\000\027\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\027\000\027\000\027\000\027\000\027\000\000\000\000\000\ -\000\000\000\000\027\000\000\000\000\000\000\000\027\000\000\000\ +\000\000\000\000\027\000\027\000\027\000\027\000\027\000\000\000\ +\000\000\000\000\000\000\027\000\000\000\000\000\000\000\027\000\ +\000\000\000\000\000\000\027\000\027\000\000\000\000\000\000\000\ \000\000\000\000\027\000\027\000\000\000\000\000\000\000\000\000\ -\000\000\027\000\027\000\000\000\000\000\000\000\000\000\027\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\027\000\ -\000\000\000\000\027\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\027\000\000\000\038\001\000\000\027\000\000\000\000\000\ -\000\000\027\000\000\000\000\000\000\000\000\000\038\001\000\000\ -\027\000\027\000\000\000\000\000\027\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\027\000\000\000\027\000\000\000\ -\000\000\000\000\027\000\027\000\027\000\027\000\027\000\249\001\ -\000\000\027\000\000\000\249\001\000\000\027\000\249\001\000\000\ -\024\001\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\249\001\249\001\249\001\249\001\000\000\000\000\ -\000\000\000\000\000\000\042\001\000\000\000\000\000\000\249\001\ -\000\000\000\000\000\000\000\000\249\001\249\001\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\249\001\ -\249\001\249\001\249\001\249\001\000\000\000\000\000\000\000\000\ -\249\001\000\000\000\000\000\000\249\001\000\000\000\000\000\000\ -\249\001\249\001\000\000\000\000\000\000\000\000\000\000\249\001\ -\249\001\000\000\000\000\000\000\000\000\249\001\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\249\001\000\000\000\000\ -\249\001\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\249\001\ -\000\000\042\001\000\000\249\001\000\000\000\000\000\000\249\001\ -\000\000\000\000\000\000\000\000\042\001\000\000\249\001\249\001\ -\000\000\000\000\249\001\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\249\001\000\000\249\001\000\000\000\000\000\000\ -\249\001\249\001\249\001\249\001\249\001\250\001\000\000\249\001\ -\000\000\250\001\000\000\249\001\250\001\000\000\025\001\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\250\001\250\001\250\001\250\001\000\000\000\000\000\000\000\000\ -\000\000\043\001\000\000\000\000\000\000\250\001\000\000\000\000\ -\000\000\000\000\250\001\250\001\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\250\001\250\001\250\001\ -\250\001\250\001\000\000\000\000\000\000\000\000\250\001\000\000\ -\000\000\000\000\250\001\000\000\000\000\000\000\250\001\250\001\ -\000\000\000\000\000\000\000\000\000\000\250\001\250\001\000\000\ -\000\000\000\000\000\000\250\001\000\000\000\000\000\000\000\000\ +\027\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ \000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\250\001\000\000\000\000\250\001\000\000\ +\027\000\000\000\000\000\027\000\000\000\000\000\000\000\000\000\ \000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\250\001\000\000\043\001\ -\000\000\250\001\000\000\000\000\000\000\250\001\000\000\000\000\ -\000\000\000\000\043\001\000\000\250\001\250\001\000\000\000\000\ -\250\001\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\250\001\000\000\250\001\000\000\000\000\000\000\250\001\250\001\ -\250\001\250\001\250\001\246\001\000\000\250\001\000\000\246\001\ -\000\000\250\001\246\001\000\000\021\001\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\246\001\246\001\ -\246\001\246\001\000\000\000\000\000\000\000\000\000\000\039\001\ -\000\000\000\000\000\000\246\001\000\000\000\000\000\000\000\000\ -\246\001\246\001\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\246\001\246\001\246\001\246\001\246\001\ -\000\000\000\000\000\000\000\000\246\001\000\000\000\000\000\000\ -\246\001\000\000\000\000\000\000\246\001\246\001\000\000\000\000\ -\000\000\000\000\000\000\246\001\246\001\000\000\000\000\000\000\ -\000\000\246\001\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\246\001\000\000\000\000\246\001\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\246\001\000\000\039\001\000\000\246\001\ -\000\000\000\000\000\000\246\001\000\000\000\000\000\000\000\000\ -\039\001\000\000\246\001\246\001\000\000\000\000\246\001\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\246\001\000\000\ -\246\001\000\000\000\000\000\000\246\001\246\001\246\001\246\001\ -\246\001\248\001\000\000\246\001\000\000\248\001\000\000\246\001\ -\248\001\000\000\023\001\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\027\000\000\000\040\001\000\000\027\000\000\000\ +\000\000\000\000\027\000\000\000\000\000\000\000\000\000\040\001\ +\000\000\027\000\027\000\000\000\000\000\027\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\027\000\000\000\027\000\ +\000\000\000\000\000\000\027\000\027\000\027\000\027\000\027\000\ +\251\001\000\000\027\000\000\000\251\001\000\000\027\000\251\001\ +\000\000\026\001\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\251\001\251\001\251\001\251\001\000\000\ +\000\000\000\000\000\000\000\000\044\001\000\000\000\000\000\000\ +\251\001\000\000\000\000\000\000\000\000\251\001\251\001\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\251\001\251\001\251\001\251\001\251\001\000\000\000\000\000\000\ +\000\000\251\001\000\000\000\000\000\000\251\001\000\000\000\000\ +\000\000\251\001\251\001\000\000\000\000\000\000\000\000\000\000\ +\251\001\251\001\000\000\000\000\000\000\000\000\251\001\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\251\001\000\000\ +\000\000\251\001\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\251\001\000\000\044\001\000\000\251\001\000\000\000\000\000\000\ +\251\001\000\000\000\000\000\000\000\000\044\001\000\000\251\001\ +\251\001\000\000\000\000\251\001\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\251\001\000\000\251\001\000\000\000\000\ +\000\000\251\001\251\001\251\001\251\001\251\001\252\001\000\000\ +\251\001\000\000\252\001\000\000\251\001\252\001\000\000\027\001\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\252\001\252\001\252\001\252\001\000\000\000\000\000\000\ +\000\000\000\000\045\001\000\000\000\000\000\000\252\001\000\000\ +\000\000\000\000\000\000\252\001\252\001\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\252\001\252\001\ +\252\001\252\001\252\001\000\000\000\000\000\000\000\000\252\001\ +\000\000\000\000\000\000\252\001\000\000\000\000\000\000\252\001\ +\252\001\000\000\000\000\000\000\000\000\000\000\252\001\252\001\ +\000\000\000\000\000\000\000\000\252\001\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\252\001\000\000\000\000\252\001\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\252\001\000\000\ +\045\001\000\000\252\001\000\000\000\000\000\000\252\001\000\000\ +\000\000\000\000\000\000\045\001\000\000\252\001\252\001\000\000\ +\000\000\252\001\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\252\001\000\000\252\001\000\000\000\000\000\000\252\001\ +\252\001\252\001\252\001\252\001\248\001\000\000\252\001\000\000\ +\248\001\000\000\252\001\248\001\000\000\023\001\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\248\001\ +\248\001\248\001\248\001\000\000\000\000\000\000\000\000\000\000\ +\041\001\000\000\000\000\000\000\248\001\000\000\000\000\000\000\ +\000\000\248\001\248\001\000\000\000\000\000\000\000\000\000\000\ \000\000\000\000\000\000\000\000\248\001\248\001\248\001\248\001\ -\000\000\000\000\000\000\000\000\000\000\041\001\000\000\000\000\ -\000\000\248\001\000\000\000\000\000\000\000\000\248\001\248\001\ +\248\001\000\000\000\000\000\000\000\000\248\001\000\000\000\000\ +\000\000\248\001\000\000\000\000\000\000\248\001\248\001\000\000\ +\000\000\000\000\000\000\000\000\248\001\248\001\000\000\000\000\ +\000\000\000\000\248\001\000\000\000\000\000\000\000\000\000\000\ \000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\248\001\248\001\248\001\248\001\248\001\000\000\000\000\ -\000\000\000\000\248\001\000\000\000\000\000\000\248\001\000\000\ -\000\000\000\000\248\001\248\001\000\000\000\000\000\000\000\000\ -\000\000\248\001\248\001\000\000\000\000\000\000\000\000\248\001\ +\000\000\000\000\248\001\000\000\000\000\248\001\000\000\000\000\ \000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\248\001\000\000\041\001\000\000\ +\248\001\000\000\000\000\000\000\248\001\000\000\000\000\000\000\ +\000\000\041\001\000\000\248\001\248\001\000\000\000\000\248\001\ \000\000\000\000\000\000\000\000\000\000\000\000\000\000\248\001\ -\000\000\000\000\248\001\000\000\000\000\000\000\000\000\000\000\ +\000\000\248\001\000\000\000\000\000\000\248\001\248\001\248\001\ +\248\001\248\001\250\001\000\000\248\001\000\000\250\001\000\000\ +\248\001\250\001\000\000\025\001\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\250\001\250\001\250\001\ +\250\001\000\000\000\000\000\000\000\000\000\000\043\001\000\000\ +\000\000\000\000\250\001\000\000\000\000\000\000\000\000\250\001\ +\250\001\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\250\001\250\001\250\001\250\001\250\001\000\000\ +\000\000\000\000\000\000\250\001\000\000\000\000\000\000\250\001\ +\000\000\000\000\000\000\250\001\250\001\000\000\000\000\000\000\ +\000\000\000\000\250\001\250\001\000\000\000\000\000\000\000\000\ +\250\001\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ \000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\248\001\000\000\041\001\000\000\248\001\000\000\000\000\ -\000\000\248\001\000\000\000\000\000\000\000\000\041\001\000\000\ -\248\001\248\001\000\000\000\000\248\001\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\248\001\000\000\248\001\000\000\ -\000\000\000\000\248\001\248\001\248\001\248\001\248\001\247\001\ -\000\000\248\001\000\000\247\001\000\000\248\001\247\001\000\000\ -\022\001\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\247\001\247\001\247\001\247\001\000\000\000\000\ -\000\000\000\000\000\000\040\001\000\000\000\000\000\000\247\001\ -\000\000\000\000\000\000\000\000\247\001\247\001\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\247\001\ -\247\001\247\001\247\001\247\001\000\000\000\000\000\000\000\000\ -\247\001\000\000\000\000\000\000\247\001\000\000\000\000\000\000\ -\247\001\247\001\000\000\000\000\000\000\000\000\000\000\247\001\ -\247\001\000\000\000\000\000\000\000\000\247\001\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\247\001\000\000\000\000\ -\247\001\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\247\001\ -\000\000\040\001\000\000\247\001\000\000\000\000\000\000\247\001\ -\000\000\000\000\000\000\000\000\040\001\000\000\247\001\247\001\ -\000\000\000\000\247\001\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\247\001\000\000\247\001\000\000\000\000\000\000\ -\247\001\247\001\247\001\247\001\247\001\249\001\000\000\247\001\ -\000\000\249\001\000\000\247\001\249\001\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\249\001\249\001\249\001\249\001\000\000\000\000\000\000\000\000\ -\000\000\042\001\000\000\000\000\000\000\249\001\000\000\000\000\ -\000\000\000\000\249\001\249\001\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\249\001\249\001\249\001\ -\249\001\249\001\000\000\000\000\000\000\000\000\249\001\000\000\ -\000\000\000\000\249\001\000\000\000\000\000\000\249\001\249\001\ -\000\000\000\000\000\000\000\000\000\000\249\001\249\001\000\000\ -\000\000\000\000\000\000\249\001\000\000\000\000\000\000\000\000\ +\250\001\000\000\000\000\250\001\000\000\000\000\000\000\000\000\ \000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\249\001\000\000\000\000\249\001\000\000\ +\000\000\000\000\250\001\000\000\043\001\000\000\250\001\000\000\ +\000\000\000\000\250\001\000\000\000\000\000\000\000\000\043\001\ +\000\000\250\001\250\001\000\000\000\000\250\001\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\250\001\000\000\250\001\ +\000\000\000\000\000\000\250\001\250\001\250\001\250\001\250\001\ +\249\001\000\000\250\001\000\000\249\001\000\000\250\001\249\001\ +\000\000\024\001\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\249\001\249\001\249\001\249\001\000\000\ +\000\000\000\000\000\000\000\000\042\001\000\000\000\000\000\000\ +\249\001\000\000\000\000\000\000\000\000\249\001\249\001\000\000\ \000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\249\001\000\000\042\001\ +\249\001\249\001\249\001\249\001\249\001\000\000\000\000\000\000\ \000\000\249\001\000\000\000\000\000\000\249\001\000\000\000\000\ -\000\000\000\000\042\001\000\000\249\001\249\001\000\000\000\000\ -\249\001\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\249\001\000\000\249\001\000\000\000\000\000\000\249\001\249\001\ -\249\001\249\001\249\001\250\001\000\000\249\001\000\000\250\001\ -\000\000\249\001\250\001\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\250\001\250\001\ -\250\001\250\001\000\000\000\000\000\000\000\000\000\000\043\001\ -\000\000\000\000\000\000\250\001\000\000\000\000\000\000\000\000\ -\250\001\250\001\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\250\001\250\001\250\001\250\001\250\001\ -\000\000\000\000\000\000\000\000\250\001\000\000\000\000\000\000\ -\250\001\000\000\000\000\000\000\250\001\250\001\000\000\000\000\ -\000\000\000\000\000\000\250\001\250\001\000\000\000\000\000\000\ -\000\000\250\001\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\249\001\249\001\000\000\000\000\000\000\000\000\000\000\ +\249\001\249\001\000\000\000\000\000\000\000\000\249\001\000\000\ \000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\250\001\000\000\000\000\250\001\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\249\001\000\000\ +\000\000\249\001\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\249\001\000\000\042\001\000\000\249\001\000\000\000\000\000\000\ +\249\001\000\000\000\000\000\000\000\000\042\001\000\000\249\001\ +\249\001\000\000\000\000\249\001\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\249\001\000\000\249\001\000\000\000\000\ +\000\000\249\001\249\001\249\001\249\001\249\001\251\001\000\000\ +\249\001\000\000\251\001\000\000\249\001\251\001\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\251\001\251\001\251\001\251\001\000\000\000\000\000\000\ +\000\000\000\000\044\001\000\000\000\000\000\000\251\001\000\000\ +\000\000\000\000\000\000\251\001\251\001\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\251\001\251\001\ +\251\001\251\001\251\001\000\000\000\000\000\000\000\000\251\001\ +\000\000\000\000\000\000\251\001\000\000\000\000\000\000\251\001\ +\251\001\000\000\000\000\000\000\000\000\000\000\251\001\251\001\ +\000\000\000\000\000\000\000\000\251\001\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\251\001\000\000\000\000\251\001\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\251\001\000\000\ +\044\001\000\000\251\001\000\000\000\000\000\000\251\001\000\000\ +\000\000\000\000\000\000\044\001\000\000\251\001\251\001\000\000\ +\000\000\251\001\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\251\001\000\000\251\001\000\000\000\000\000\000\251\001\ +\251\001\251\001\251\001\251\001\252\001\000\000\251\001\000\000\ +\252\001\000\000\251\001\252\001\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\252\001\ +\252\001\252\001\252\001\000\000\000\000\000\000\000\000\000\000\ +\045\001\000\000\000\000\000\000\252\001\000\000\000\000\000\000\ +\000\000\252\001\252\001\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\252\001\252\001\252\001\252\001\ +\252\001\000\000\000\000\000\000\000\000\252\001\000\000\000\000\ +\000\000\252\001\000\000\000\000\000\000\252\001\252\001\000\000\ +\000\000\000\000\000\000\000\000\252\001\252\001\000\000\000\000\ +\000\000\000\000\252\001\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\252\001\000\000\000\000\252\001\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\252\001\000\000\045\001\000\000\ +\252\001\000\000\000\000\000\000\252\001\000\000\000\000\000\000\ +\000\000\045\001\000\000\252\001\252\001\000\000\000\000\252\001\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\252\001\ +\000\000\252\001\000\000\000\000\000\000\252\001\252\001\252\001\ +\252\001\252\001\248\001\000\000\252\001\000\000\248\001\000\000\ +\252\001\248\001\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\248\001\248\001\248\001\ +\248\001\000\000\000\000\000\000\000\000\000\000\041\001\000\000\ +\000\000\000\000\248\001\000\000\000\000\000\000\000\000\248\001\ +\248\001\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\248\001\248\001\248\001\248\001\248\001\000\000\ +\000\000\000\000\000\000\248\001\000\000\000\000\000\000\248\001\ +\000\000\000\000\000\000\248\001\248\001\000\000\000\000\000\000\ +\000\000\000\000\248\001\248\001\000\000\000\000\000\000\000\000\ +\248\001\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ \000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\250\001\000\000\043\001\000\000\250\001\ -\000\000\000\000\000\000\250\001\000\000\000\000\000\000\000\000\ -\043\001\000\000\250\001\250\001\000\000\000\000\250\001\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\250\001\000\000\ -\250\001\000\000\000\000\000\000\250\001\250\001\250\001\250\001\ -\250\001\246\001\000\000\250\001\000\000\246\001\000\000\250\001\ -\246\001\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\246\001\246\001\246\001\246\001\ -\000\000\000\000\000\000\000\000\000\000\039\001\000\000\000\000\ -\000\000\246\001\000\000\000\000\000\000\000\000\246\001\246\001\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\246\001\246\001\246\001\246\001\246\001\000\000\000\000\ -\000\000\000\000\246\001\000\000\000\000\000\000\246\001\000\000\ -\000\000\000\000\246\001\246\001\000\000\000\000\000\000\000\000\ -\000\000\246\001\246\001\000\000\000\000\000\000\000\000\246\001\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\246\001\ -\000\000\000\000\246\001\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\246\001\000\000\039\001\000\000\246\001\000\000\000\000\ -\000\000\246\001\000\000\000\000\000\000\000\000\039\001\000\000\ -\246\001\246\001\000\000\000\000\246\001\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\246\001\000\000\246\001\000\000\ -\000\000\000\000\246\001\246\001\246\001\246\001\246\001\248\001\ -\000\000\246\001\000\000\248\001\000\000\246\001\248\001\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\248\001\248\001\248\001\248\001\000\000\000\000\ -\000\000\000\000\000\000\041\001\000\000\000\000\000\000\248\001\ -\000\000\000\000\000\000\000\000\248\001\248\001\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\248\001\ -\248\001\248\001\248\001\248\001\000\000\000\000\000\000\000\000\ -\248\001\000\000\000\000\000\000\248\001\000\000\000\000\000\000\ -\248\001\248\001\000\000\000\000\000\000\000\000\000\000\248\001\ -\248\001\000\000\000\000\000\000\000\000\248\001\000\000\000\000\ +\248\001\000\000\000\000\248\001\000\000\000\000\000\000\000\000\ \000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\248\001\000\000\000\000\ -\248\001\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\248\001\ -\000\000\041\001\000\000\248\001\000\000\000\000\000\000\248\001\ -\000\000\000\000\000\000\000\000\041\001\000\000\248\001\248\001\ -\000\000\000\000\248\001\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\248\001\000\000\248\001\000\000\000\000\000\000\ -\248\001\248\001\248\001\248\001\248\001\247\001\000\000\248\001\ -\000\000\247\001\000\000\248\001\247\001\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\247\001\247\001\247\001\247\001\000\000\000\000\000\000\000\000\ -\000\000\040\001\000\000\000\000\000\000\247\001\000\000\000\000\ -\000\000\000\000\247\001\247\001\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\247\001\247\001\247\001\ -\247\001\247\001\000\000\000\000\000\000\000\000\247\001\000\000\ -\000\000\000\000\247\001\000\000\000\000\000\000\247\001\247\001\ -\000\000\000\000\000\000\000\000\000\000\247\001\247\001\000\000\ -\000\000\000\000\000\000\247\001\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\247\001\000\000\000\000\247\001\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\247\001\000\000\040\001\ -\000\000\247\001\000\000\174\003\000\000\247\001\000\000\038\000\ -\000\000\000\000\040\001\039\000\247\001\247\001\041\000\000\000\ -\247\001\014\001\015\001\000\000\016\001\000\000\000\000\017\001\ -\247\001\000\000\247\001\000\000\000\000\000\000\247\001\247\001\ -\247\001\247\001\247\001\018\001\000\000\247\001\000\000\047\000\ -\000\000\247\001\000\000\000\000\000\000\141\000\000\000\019\001\ -\000\000\020\001\000\000\000\000\000\000\000\000\000\000\055\000\ -\056\000\057\000\058\000\059\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\063\000\000\000\000\000\000\000\ -\065\000\066\000\000\000\000\000\000\000\000\000\000\000\142\000\ -\070\000\000\000\000\000\000\000\000\000\074\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\022\001\023\001\024\001\ -\025\001\026\001\027\001\028\001\029\001\075\000\031\001\032\001\ -\076\000\033\001\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\078\000\ -\000\000\000\000\000\000\079\000\000\000\000\000\000\000\080\000\ -\034\000\035\000\036\000\037\000\038\000\000\000\082\000\083\000\ -\039\000\000\000\084\000\041\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\087\000\000\000\089\000\043\000\000\000\000\000\ -\090\000\091\000\092\000\093\000\143\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\046\000\047\000\048\000\000\000\049\000\ -\000\000\051\000\052\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\055\000\056\000\057\000\058\000\ -\059\000\000\000\000\000\000\000\000\000\061\000\062\000\000\000\ -\195\000\063\000\000\000\000\000\064\000\065\000\066\000\000\000\ -\000\000\000\000\000\000\000\000\069\000\070\000\000\000\000\000\ -\073\000\000\000\074\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\075\000\000\000\000\000\076\000\000\000\000\000\ +\000\000\000\000\248\001\000\000\041\001\000\000\248\001\000\000\ +\000\000\000\000\248\001\000\000\000\000\000\000\000\000\041\001\ +\000\000\248\001\248\001\000\000\000\000\248\001\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\248\001\000\000\248\001\ +\000\000\000\000\000\000\248\001\248\001\248\001\248\001\248\001\ +\250\001\000\000\248\001\000\000\250\001\000\000\248\001\250\001\ \000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\077\000\000\000\078\000\000\000\000\000\000\000\ -\079\000\000\000\000\000\000\000\080\000\034\000\035\000\036\000\ -\037\000\038\000\242\003\082\000\083\000\039\000\000\000\084\000\ -\041\000\000\000\000\000\000\000\000\000\086\000\000\000\087\000\ -\000\000\089\000\043\000\000\000\000\000\090\000\091\000\092\000\ -\093\000\184\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\046\000\047\000\048\000\000\000\049\000\000\000\051\000\052\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\055\000\056\000\057\000\058\000\059\000\000\000\000\000\ -\000\000\000\000\243\003\062\000\000\000\000\000\063\000\000\000\ -\000\000\064\000\065\000\066\000\000\000\000\000\000\000\000\000\ -\000\000\069\000\070\000\000\000\000\000\073\000\000\000\074\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\075\000\ -\000\000\000\000\076\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\077\000\ -\000\000\078\000\000\000\000\000\000\000\079\000\000\000\000\000\ -\000\000\080\000\034\000\035\000\036\000\037\000\038\000\000\000\ -\082\000\083\000\039\000\000\000\084\000\041\000\000\000\000\000\ -\000\000\000\000\086\000\000\000\087\000\000\000\089\000\043\000\ -\000\000\000\000\090\000\091\000\092\000\093\000\184\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\046\000\047\000\048\000\ -\000\000\049\000\000\000\051\000\052\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\055\000\056\000\ -\057\000\058\000\059\000\000\000\000\000\000\000\000\000\061\000\ -\062\000\000\000\000\000\063\000\000\000\000\000\064\000\065\000\ -\066\000\000\000\000\000\000\000\000\000\000\000\069\000\070\000\ -\000\000\000\000\073\000\000\000\074\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\075\000\000\000\000\000\076\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\077\000\000\000\078\000\000\000\ -\000\000\000\000\079\000\000\000\000\000\000\000\080\000\034\000\ -\035\000\036\000\037\000\038\000\000\000\082\000\083\000\039\000\ -\000\000\084\000\041\000\000\000\000\000\000\000\000\000\086\000\ -\000\000\087\000\000\000\089\000\043\000\000\000\000\000\090\000\ -\091\000\092\000\093\000\184\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\047\000\000\000\000\000\049\000\000\000\ -\051\000\052\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\055\000\056\000\057\000\058\000\059\000\ -\000\000\000\000\000\000\000\000\061\000\062\000\000\000\000\000\ -\063\000\000\000\000\000\064\000\065\000\066\000\000\000\000\000\ -\000\000\000\000\000\000\069\000\070\000\000\000\000\000\073\000\ -\000\000\074\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\075\000\000\000\000\000\076\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\077\000\000\000\078\000\000\000\000\000\000\000\079\000\ -\000\000\000\000\000\000\080\000\170\001\171\001\172\001\173\001\ -\038\000\000\000\082\000\083\000\039\000\000\000\084\000\041\000\ -\000\000\000\000\000\000\000\000\086\000\000\000\087\000\000\000\ -\089\000\043\000\000\000\000\000\090\000\091\000\092\000\093\000\ -\184\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\047\000\000\000\000\000\000\000\000\000\051\000\052\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\055\000\056\000\057\000\058\000\059\000\000\000\000\000\000\000\ -\000\000\174\001\175\001\000\000\000\000\063\000\000\000\000\000\ -\064\000\065\000\066\000\000\000\000\000\000\000\000\000\000\000\ -\142\000\070\000\000\000\000\000\176\001\000\000\074\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\075\000\000\000\ -\000\000\076\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\177\001\000\000\ -\078\000\000\000\000\000\000\000\079\000\000\000\000\000\000\000\ -\080\000\034\000\035\000\036\000\037\000\038\000\000\000\082\000\ -\083\000\039\000\000\000\084\000\041\000\000\000\000\000\000\000\ -\000\000\178\001\000\000\087\000\000\000\089\000\043\000\000\000\ -\000\000\090\000\091\000\092\000\093\000\184\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\047\000\000\000\000\000\ -\000\000\000\000\051\000\052\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\055\000\056\000\057\000\ -\058\000\059\000\000\000\000\000\000\000\000\000\061\000\062\000\ -\000\000\000\000\063\000\000\000\000\000\064\000\065\000\066\000\ -\000\000\000\000\000\000\000\000\000\000\069\000\070\000\000\000\ -\000\000\073\000\000\000\074\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\075\000\000\000\000\000\076\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\077\000\000\000\078\000\000\000\000\000\ -\000\000\079\000\000\000\000\000\000\000\080\000\170\001\171\001\ -\172\001\173\001\038\000\000\000\082\000\083\000\039\000\000\000\ -\084\000\041\000\000\000\000\000\000\000\000\000\086\000\000\000\ -\087\000\000\000\089\000\043\000\000\000\000\000\090\000\091\000\ -\092\000\093\000\184\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\047\000\000\000\000\000\000\000\000\000\051\000\ -\141\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\055\000\056\000\057\000\058\000\059\000\000\000\ -\000\000\000\000\000\000\174\001\175\001\000\000\000\000\063\000\ -\000\000\000\000\064\000\065\000\066\000\000\000\000\000\000\000\ -\000\000\000\000\142\000\070\000\000\000\000\000\176\001\000\000\ -\074\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\075\000\000\000\000\000\076\000\034\000\035\000\036\000\037\000\ -\038\000\000\000\000\000\000\000\039\000\000\000\000\000\041\000\ -\177\001\000\000\078\000\000\000\000\000\000\000\079\000\000\000\ -\000\000\000\000\080\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\082\000\083\000\000\000\000\000\084\000\000\000\198\000\ -\047\000\199\000\000\000\178\001\000\000\087\000\052\000\089\000\ -\000\000\000\000\000\000\090\000\091\000\092\000\093\000\143\000\ -\055\000\056\000\057\000\058\000\059\000\000\000\000\000\000\000\ -\000\000\061\000\000\000\000\000\000\000\063\000\000\000\000\000\ -\000\000\065\000\066\000\000\000\000\000\000\000\000\000\000\000\ -\142\000\070\000\000\000\000\000\073\000\000\000\074\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\075\000\000\000\ -\000\000\076\000\034\000\035\000\036\000\037\000\038\000\000\000\ -\000\000\000\000\039\000\000\000\000\000\041\000\077\000\000\000\ -\078\000\000\000\000\000\000\000\079\000\000\000\000\000\000\000\ -\080\000\000\000\000\000\000\000\000\000\000\000\000\000\082\000\ -\083\000\000\000\000\000\084\000\000\000\248\000\047\000\249\000\ -\000\000\086\000\000\000\087\000\052\000\089\000\000\000\000\000\ -\000\000\090\000\091\000\092\000\093\000\184\000\055\000\056\000\ -\057\000\058\000\059\000\000\000\000\000\000\000\000\000\061\000\ -\000\000\000\000\000\000\063\000\000\000\000\000\000\000\065\000\ -\066\000\000\000\000\000\000\000\000\000\000\000\142\000\070\000\ -\000\000\000\000\073\000\000\000\074\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\075\000\000\000\000\000\076\000\ -\034\000\035\000\036\000\037\000\038\000\000\000\000\000\000\000\ -\039\000\000\000\000\000\041\000\077\000\000\000\078\000\000\000\ -\000\000\000\000\079\000\000\000\000\000\000\000\080\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\082\000\083\000\000\000\ -\000\000\084\000\000\000\000\000\047\000\000\000\000\000\086\000\ -\000\000\087\000\052\000\089\000\000\000\000\000\000\000\090\000\ -\091\000\092\000\093\000\184\000\055\000\056\000\057\000\058\000\ -\059\000\000\000\000\000\000\000\000\000\061\000\000\000\000\000\ -\000\000\063\000\000\000\000\000\064\000\065\000\066\000\000\000\ -\000\000\000\000\000\000\000\000\142\000\070\000\000\000\000\000\ -\073\000\000\000\074\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\075\000\000\000\000\000\076\000\170\001\171\001\ -\172\001\173\001\038\000\000\000\000\000\000\000\039\000\000\000\ -\000\000\041\000\077\000\000\000\078\000\000\000\000\000\000\000\ -\079\000\000\000\000\000\000\000\080\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\082\000\083\000\000\000\000\000\084\000\ -\000\000\000\000\047\000\000\000\000\000\086\000\000\000\087\000\ -\141\000\089\000\000\000\000\000\000\000\090\000\091\000\092\000\ -\093\000\184\000\055\000\056\000\057\000\058\000\059\000\000\000\ -\000\000\000\000\000\000\174\001\000\000\000\000\000\000\063\000\ -\000\000\000\000\064\000\065\000\066\000\000\000\000\000\000\000\ -\000\000\000\000\142\000\070\000\000\000\000\000\176\001\000\000\ -\074\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\075\000\000\000\000\000\076\000\034\000\035\000\036\000\037\000\ -\038\000\000\000\000\000\000\000\039\000\000\000\000\000\041\000\ -\177\001\000\000\078\000\000\000\000\000\000\000\079\000\000\000\ -\000\000\000\000\080\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\082\000\083\000\000\000\000\000\084\000\000\000\000\000\ -\047\000\000\000\000\000\178\001\000\000\087\000\052\000\089\000\ -\000\000\000\000\000\000\090\000\091\000\092\000\093\000\143\000\ -\055\000\056\000\057\000\058\000\059\000\000\000\000\000\000\000\ -\000\000\061\000\000\000\000\000\000\000\063\000\000\000\000\000\ -\000\000\065\000\066\000\000\000\000\000\000\000\000\000\000\000\ -\142\000\070\000\000\000\000\000\073\000\000\000\074\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\075\000\000\000\ -\000\000\076\000\170\001\171\001\172\001\173\001\038\000\000\000\ -\000\000\000\000\039\000\000\000\000\000\041\000\077\000\000\000\ -\078\000\000\000\000\000\000\000\079\000\000\000\000\000\000\000\ -\080\000\000\000\000\000\000\000\000\000\000\000\000\000\082\000\ -\083\000\000\000\000\000\084\000\000\000\000\000\047\000\000\000\ -\000\000\086\000\000\000\087\000\141\000\089\000\000\000\000\000\ -\000\000\090\000\091\000\092\000\093\000\184\000\055\000\056\000\ -\057\000\058\000\059\000\000\000\000\000\000\000\000\000\174\001\ -\000\000\000\000\000\000\063\000\000\000\000\000\000\000\065\000\ -\066\000\000\000\000\000\000\000\000\000\000\000\142\000\070\000\ -\000\000\000\000\176\001\000\000\074\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\075\000\000\000\000\000\076\000\ -\034\000\035\000\036\000\037\000\038\000\000\000\000\000\000\000\ -\039\000\000\000\000\000\041\000\177\001\000\000\078\000\000\000\ -\000\000\000\000\079\000\000\000\000\000\000\000\080\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\082\000\083\000\000\000\ -\000\000\084\000\000\000\000\000\047\000\000\000\000\000\178\001\ -\000\000\087\000\052\000\089\000\000\000\000\000\000\000\090\000\ -\091\000\092\000\093\000\143\000\055\000\056\000\057\000\058\000\ -\059\000\000\000\000\000\000\000\000\000\050\004\000\000\000\000\ -\000\000\063\000\000\000\000\000\000\000\065\000\066\000\000\000\ -\000\000\000\000\000\000\000\000\142\000\070\000\000\000\000\000\ -\073\000\000\000\074\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\075\000\000\000\000\000\076\000\217\000\217\000\ -\217\000\217\000\217\000\000\000\000\000\000\000\217\000\000\000\ -\000\000\217\000\077\000\000\000\078\000\000\000\000\000\000\000\ -\079\000\000\000\000\000\000\000\080\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\082\000\083\000\000\000\000\000\084\000\ -\000\000\000\000\217\000\000\000\000\000\086\000\000\000\087\000\ -\217\000\089\000\000\000\000\000\000\000\090\000\091\000\092\000\ -\093\000\184\000\217\000\217\000\217\000\217\000\217\000\000\000\ -\000\000\000\000\000\000\217\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\217\000\217\000\217\000\000\000\000\000\000\000\ -\000\000\000\000\217\000\217\000\000\000\000\000\217\000\000\000\ -\217\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\217\000\000\000\000\000\217\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\217\000\000\000\217\000\000\000\000\000\000\000\217\000\000\000\ -\038\000\000\000\217\000\000\000\039\000\000\000\000\000\041\000\ -\000\000\217\000\217\000\000\000\000\000\217\000\000\000\000\000\ -\000\000\000\000\000\000\217\000\000\000\217\000\000\000\217\000\ -\000\000\000\000\000\000\217\000\217\000\217\000\217\000\217\000\ -\047\000\000\000\000\000\000\000\000\000\096\001\141\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\055\000\056\000\057\000\058\000\059\000\000\000\000\000\000\000\ -\000\000\097\001\000\000\000\000\000\000\063\000\000\000\000\000\ -\000\000\065\000\066\000\000\000\000\000\000\000\000\000\000\000\ -\142\000\070\000\000\000\000\000\000\000\000\000\074\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\075\000\000\000\ -\000\000\076\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\038\000\000\000\000\000\000\000\039\000\000\000\ -\078\000\041\000\000\000\000\000\079\000\000\000\000\000\000\000\ -\080\000\000\000\000\000\000\000\000\000\000\000\000\000\082\000\ -\083\000\000\000\000\000\084\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\047\000\087\000\000\000\089\000\000\000\157\002\ -\141\000\090\000\091\000\092\000\093\000\143\000\000\000\000\000\ -\098\001\000\000\055\000\056\000\057\000\058\000\059\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\063\000\ -\000\000\000\000\000\000\065\000\066\000\000\000\000\000\000\000\ -\000\000\000\000\142\000\070\000\000\000\000\000\000\000\000\000\ -\074\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\075\000\000\000\000\000\076\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\037\000\038\000\000\000\000\000\000\000\ -\039\000\000\000\078\000\041\000\000\000\000\000\079\000\000\000\ -\000\000\000\000\080\000\000\000\000\000\043\000\000\000\000\000\ -\000\000\082\000\083\000\000\000\000\000\084\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\047\000\087\000\000\000\089\000\ -\000\000\051\000\052\000\090\000\091\000\092\000\093\000\143\000\ -\000\000\000\000\098\001\000\000\055\000\056\000\057\000\058\000\ -\059\000\000\000\000\000\000\000\000\000\061\000\222\002\000\000\ -\000\000\063\000\000\000\000\000\064\000\065\000\066\000\000\000\ -\000\000\000\000\000\000\000\000\223\002\070\000\000\000\000\000\ -\000\000\000\000\074\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\075\000\000\000\000\000\076\000\000\000\000\000\ -\000\000\037\000\038\000\000\000\000\000\000\000\039\000\000\000\ -\000\000\041\000\000\000\000\000\078\000\000\000\000\000\000\000\ -\079\000\000\000\000\000\000\000\080\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\082\000\083\000\000\000\000\000\084\000\ -\000\000\000\000\047\000\000\000\000\000\086\000\000\000\087\000\ -\052\000\089\000\000\000\000\000\000\000\090\000\091\000\092\000\ -\093\000\184\000\055\000\056\000\057\000\058\000\059\000\000\000\ -\000\000\000\000\000\000\061\000\000\000\000\000\000\000\063\000\ -\000\000\000\000\064\000\065\000\066\000\000\000\000\000\000\000\ -\000\000\000\000\142\000\070\000\000\000\000\000\000\000\000\000\ -\074\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\075\000\000\000\000\000\076\000\000\000\000\000\000\000\037\000\ -\038\000\000\000\000\000\000\000\039\000\000\000\000\000\041\000\ -\000\000\000\000\078\000\000\000\000\000\000\000\079\000\000\000\ +\000\000\000\000\000\000\250\001\250\001\250\001\250\001\000\000\ +\000\000\000\000\000\000\000\000\043\001\000\000\000\000\000\000\ +\250\001\000\000\000\000\000\000\000\000\250\001\250\001\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\250\001\250\001\250\001\250\001\250\001\000\000\000\000\000\000\ +\000\000\250\001\000\000\000\000\000\000\250\001\000\000\000\000\ +\000\000\250\001\250\001\000\000\000\000\000\000\000\000\000\000\ +\250\001\250\001\000\000\000\000\000\000\000\000\250\001\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\250\001\000\000\ +\000\000\250\001\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\250\001\000\000\043\001\000\000\250\001\000\000\000\000\000\000\ +\250\001\000\000\000\000\000\000\000\000\043\001\000\000\250\001\ +\250\001\000\000\000\000\250\001\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\250\001\000\000\250\001\000\000\000\000\ +\000\000\250\001\250\001\250\001\250\001\250\001\249\001\000\000\ +\250\001\000\000\249\001\000\000\250\001\249\001\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\249\001\249\001\249\001\249\001\000\000\000\000\000\000\ +\000\000\000\000\042\001\000\000\000\000\000\000\249\001\000\000\ +\000\000\000\000\000\000\249\001\249\001\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\249\001\249\001\ +\249\001\249\001\249\001\000\000\000\000\000\000\000\000\249\001\ +\000\000\000\000\000\000\249\001\000\000\000\000\000\000\249\001\ +\249\001\000\000\000\000\000\000\000\000\000\000\249\001\249\001\ +\000\000\000\000\000\000\000\000\249\001\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\249\001\000\000\000\000\249\001\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\249\001\000\000\ +\042\001\000\000\249\001\000\000\178\003\000\000\249\001\000\000\ +\042\000\000\000\000\000\042\001\043\000\249\001\249\001\045\000\ +\000\000\249\001\018\001\019\001\000\000\020\001\000\000\000\000\ +\021\001\249\001\000\000\249\001\000\000\000\000\000\000\249\001\ +\249\001\249\001\249\001\249\001\022\001\000\000\249\001\000\000\ +\051\000\000\000\249\001\000\000\000\000\000\000\145\000\000\000\ +\023\001\000\000\024\001\000\000\000\000\000\000\000\000\000\000\ +\059\000\060\000\061\000\062\000\063\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\067\000\000\000\000\000\ +\000\000\069\000\070\000\000\000\000\000\000\000\000\000\000\000\ +\146\000\074\000\000\000\000\000\000\000\000\000\078\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\026\001\027\001\ +\028\001\029\001\030\001\031\001\032\001\033\001\079\000\035\001\ +\036\001\080\000\037\001\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\082\000\000\000\000\000\000\000\083\000\000\000\000\000\000\000\ +\084\000\038\000\039\000\040\000\041\000\042\000\000\000\086\000\ +\087\000\043\000\000\000\088\000\045\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\091\000\000\000\093\000\047\000\000\000\ +\000\000\094\000\095\000\096\000\097\000\147\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\050\000\051\000\052\000\000\000\ +\053\000\000\000\055\000\056\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\059\000\060\000\061\000\ +\062\000\063\000\000\000\000\000\000\000\000\000\065\000\066\000\ +\000\000\199\000\067\000\000\000\000\000\068\000\069\000\070\000\ +\000\000\000\000\000\000\000\000\000\000\073\000\074\000\000\000\ +\000\000\077\000\000\000\078\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\079\000\000\000\000\000\080\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\081\000\000\000\082\000\000\000\000\000\ +\000\000\083\000\000\000\000\000\000\000\084\000\038\000\039\000\ +\040\000\041\000\042\000\246\003\086\000\087\000\043\000\000\000\ +\088\000\045\000\000\000\000\000\000\000\000\000\090\000\000\000\ +\091\000\000\000\093\000\047\000\000\000\000\000\094\000\095\000\ +\096\000\097\000\188\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\050\000\051\000\052\000\000\000\053\000\000\000\055\000\ +\056\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\059\000\060\000\061\000\062\000\063\000\000\000\ +\000\000\000\000\000\000\247\003\066\000\000\000\000\000\067\000\ +\000\000\000\000\068\000\069\000\070\000\000\000\000\000\000\000\ +\000\000\000\000\073\000\074\000\000\000\000\000\077\000\000\000\ +\078\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\079\000\000\000\000\000\080\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\081\000\000\000\082\000\000\000\000\000\000\000\083\000\000\000\ +\000\000\000\000\084\000\038\000\039\000\040\000\041\000\042\000\ +\000\000\086\000\087\000\043\000\000\000\088\000\045\000\000\000\ +\000\000\000\000\000\000\090\000\000\000\091\000\000\000\093\000\ +\047\000\000\000\000\000\094\000\095\000\096\000\097\000\188\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\050\000\051\000\ +\052\000\000\000\053\000\000\000\055\000\056\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\059\000\ +\060\000\061\000\062\000\063\000\000\000\000\000\000\000\000\000\ +\065\000\066\000\000\000\000\000\067\000\000\000\000\000\068\000\ +\069\000\070\000\000\000\000\000\000\000\000\000\000\000\073\000\ +\074\000\000\000\000\000\077\000\000\000\078\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\079\000\000\000\000\000\ +\080\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\081\000\000\000\082\000\ +\000\000\000\000\000\000\083\000\000\000\000\000\000\000\084\000\ +\038\000\039\000\040\000\041\000\042\000\000\000\086\000\087\000\ +\043\000\000\000\088\000\045\000\000\000\000\000\000\000\000\000\ +\090\000\000\000\091\000\000\000\093\000\047\000\000\000\000\000\ +\094\000\095\000\096\000\097\000\188\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\051\000\000\000\000\000\053\000\ +\000\000\055\000\056\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\059\000\060\000\061\000\062\000\ +\063\000\000\000\000\000\000\000\000\000\065\000\066\000\000\000\ +\000\000\067\000\000\000\000\000\068\000\069\000\070\000\000\000\ +\000\000\000\000\000\000\000\000\073\000\074\000\000\000\000\000\ +\077\000\000\000\078\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\079\000\000\000\000\000\080\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\081\000\000\000\082\000\000\000\000\000\000\000\ +\083\000\000\000\000\000\000\000\084\000\174\001\175\001\176\001\ +\177\001\042\000\000\000\086\000\087\000\043\000\000\000\088\000\ +\045\000\000\000\000\000\000\000\000\000\090\000\000\000\091\000\ +\000\000\093\000\047\000\000\000\000\000\094\000\095\000\096\000\ +\097\000\188\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\051\000\000\000\000\000\000\000\000\000\055\000\056\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\059\000\060\000\061\000\062\000\063\000\000\000\000\000\ +\000\000\000\000\178\001\179\001\000\000\000\000\067\000\000\000\ +\000\000\068\000\069\000\070\000\000\000\000\000\000\000\000\000\ +\000\000\146\000\074\000\000\000\000\000\180\001\000\000\078\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\079\000\ \000\000\000\000\080\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\082\000\083\000\000\000\000\000\084\000\000\000\000\000\ -\047\000\000\000\000\000\086\000\000\000\087\000\052\000\089\000\ -\000\000\000\000\000\000\090\000\091\000\092\000\093\000\184\000\ -\055\000\056\000\057\000\058\000\059\000\000\000\000\000\000\000\ -\000\000\061\000\000\000\000\000\000\000\063\000\000\000\000\000\ -\000\000\065\000\066\000\000\000\000\000\000\000\000\000\000\000\ -\142\000\070\000\000\000\000\000\000\000\000\000\074\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\075\000\000\000\ -\000\000\076\000\000\000\000\000\000\000\173\001\038\000\000\000\ -\000\000\000\000\039\000\000\000\000\000\041\000\000\000\000\000\ -\078\000\000\000\000\000\000\000\079\000\000\000\000\000\000\000\ -\080\000\000\000\000\000\000\000\000\000\000\000\000\000\082\000\ -\083\000\000\000\000\000\084\000\000\000\000\000\047\000\000\000\ -\000\000\086\000\000\000\087\000\141\000\089\000\000\000\000\000\ -\000\000\090\000\091\000\092\000\093\000\184\000\055\000\056\000\ -\057\000\058\000\059\000\000\000\000\000\000\000\000\000\174\001\ -\000\000\000\000\000\000\063\000\000\000\000\000\000\000\065\000\ -\066\000\000\000\000\000\000\000\000\000\000\000\142\000\070\000\ -\000\000\000\000\000\000\000\000\074\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\075\000\000\000\000\000\076\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\217\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\078\000\041\000\ -\000\000\000\000\079\000\000\000\000\000\000\000\080\000\000\000\ -\000\000\043\000\000\000\000\000\000\000\082\000\083\000\000\000\ -\000\000\084\000\000\000\000\000\000\000\000\000\000\000\178\001\ -\047\000\087\000\000\000\089\000\000\000\218\000\141\000\090\000\ -\091\000\092\000\093\000\143\000\000\000\000\000\000\000\000\000\ -\055\000\056\000\057\000\058\000\059\000\000\000\000\000\000\000\ -\000\000\219\000\220\000\000\000\000\000\221\000\000\000\000\000\ -\064\000\000\000\222\000\000\000\000\000\000\000\000\000\000\000\ -\223\000\070\000\000\000\000\000\000\000\000\000\074\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\217\000\000\000\236\000\000\000\000\000\000\000\000\000\ -\078\000\041\000\000\000\000\000\079\000\000\000\000\000\000\000\ -\080\000\000\000\000\000\043\000\000\000\000\000\000\000\082\000\ -\083\000\000\000\000\000\084\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\047\000\087\000\000\000\089\000\000\000\218\000\ -\141\000\090\000\091\000\092\000\093\000\224\000\000\000\000\000\ -\000\000\000\000\055\000\056\000\057\000\058\000\059\000\000\000\ -\000\000\000\000\000\000\219\000\000\000\000\000\000\000\221\000\ -\000\000\000\000\064\000\000\000\222\000\000\000\000\000\000\000\ -\000\000\000\000\237\000\070\000\000\000\000\000\000\000\000\000\ -\074\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\217\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\078\000\041\000\000\000\000\000\079\000\000\000\ -\000\000\000\000\080\000\000\000\000\000\043\000\000\000\000\000\ -\000\000\082\000\083\000\000\000\000\000\084\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\047\000\087\000\000\000\089\000\ -\000\000\218\000\141\000\090\000\091\000\092\000\093\000\224\000\ -\000\000\000\000\000\000\000\000\055\000\056\000\057\000\058\000\ -\059\000\000\000\000\000\000\000\000\000\219\000\000\000\000\000\ -\000\000\221\000\000\000\000\000\064\000\000\000\222\000\000\000\ -\000\000\000\000\000\000\000\000\237\000\070\000\000\000\000\000\ -\000\000\000\000\074\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\217\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\078\000\041\000\000\000\000\000\ -\079\000\000\000\000\000\000\000\080\000\000\000\000\000\043\000\ -\011\002\000\000\000\000\082\000\083\000\000\000\000\000\084\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\047\000\087\000\ -\000\000\089\000\000\000\218\000\141\000\090\000\091\000\092\000\ -\093\000\224\000\000\000\000\000\000\000\000\000\055\000\056\000\ -\057\000\058\000\059\000\000\000\000\000\000\000\000\000\219\000\ -\000\000\000\000\000\000\221\000\000\000\000\000\064\000\000\000\ -\222\000\000\000\000\000\000\000\000\000\000\000\237\000\070\000\ -\000\000\000\000\000\000\000\000\074\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\217\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\078\000\041\000\ -\000\000\000\000\079\000\000\000\000\000\000\000\080\000\000\000\ -\000\000\043\000\200\002\000\000\000\000\082\000\083\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\181\001\ +\000\000\082\000\000\000\000\000\000\000\083\000\000\000\000\000\ +\000\000\084\000\038\000\039\000\040\000\041\000\042\000\000\000\ +\086\000\087\000\043\000\000\000\088\000\045\000\000\000\000\000\ +\000\000\000\000\182\001\000\000\091\000\000\000\093\000\047\000\ +\000\000\000\000\094\000\095\000\096\000\097\000\188\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\051\000\000\000\ +\000\000\000\000\000\000\055\000\056\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\059\000\060\000\ +\061\000\062\000\063\000\000\000\000\000\000\000\000\000\065\000\ +\066\000\000\000\000\000\067\000\000\000\000\000\068\000\069\000\ +\070\000\000\000\000\000\000\000\000\000\000\000\073\000\074\000\ +\000\000\000\000\077\000\000\000\078\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\079\000\000\000\000\000\080\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\081\000\000\000\082\000\000\000\ +\000\000\000\000\083\000\000\000\000\000\000\000\084\000\174\001\ +\175\001\176\001\177\001\042\000\000\000\086\000\087\000\043\000\ +\000\000\088\000\045\000\000\000\000\000\000\000\000\000\090\000\ +\000\000\091\000\000\000\093\000\047\000\000\000\000\000\094\000\ +\095\000\096\000\097\000\188\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\051\000\000\000\000\000\000\000\000\000\ +\055\000\145\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\059\000\060\000\061\000\062\000\063\000\ +\000\000\000\000\000\000\000\000\178\001\179\001\000\000\000\000\ +\067\000\000\000\000\000\068\000\069\000\070\000\000\000\000\000\ +\000\000\000\000\000\000\146\000\074\000\000\000\000\000\180\001\ +\000\000\078\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\079\000\000\000\000\000\080\000\038\000\039\000\040\000\ +\041\000\042\000\000\000\000\000\000\000\043\000\000\000\000\000\ +\045\000\181\001\000\000\082\000\000\000\000\000\000\000\083\000\ +\000\000\000\000\000\000\084\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\086\000\087\000\000\000\000\000\088\000\000\000\ +\202\000\051\000\203\000\000\000\182\001\000\000\091\000\056\000\ +\093\000\000\000\000\000\000\000\094\000\095\000\096\000\097\000\ +\147\000\059\000\060\000\061\000\062\000\063\000\000\000\000\000\ +\000\000\000\000\065\000\000\000\000\000\000\000\067\000\000\000\ +\000\000\000\000\069\000\070\000\000\000\000\000\000\000\000\000\ +\000\000\146\000\074\000\000\000\000\000\077\000\000\000\078\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\079\000\ +\000\000\000\000\080\000\038\000\039\000\040\000\041\000\042\000\ +\000\000\000\000\000\000\043\000\000\000\000\000\045\000\081\000\ +\000\000\082\000\000\000\000\000\000\000\083\000\000\000\000\000\ \000\000\084\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\047\000\087\000\000\000\089\000\000\000\218\000\141\000\090\000\ -\091\000\092\000\093\000\224\000\000\000\000\000\000\000\000\000\ -\055\000\056\000\057\000\058\000\059\000\000\000\000\000\000\000\ -\000\000\219\000\000\000\000\000\000\000\221\000\000\000\000\000\ -\064\000\000\000\222\000\000\000\000\000\000\000\000\000\000\000\ -\237\000\070\000\000\000\000\000\000\000\000\000\074\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\215\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\078\000\215\000\000\000\000\000\079\000\000\000\000\000\000\000\ -\080\000\000\000\000\000\215\000\124\004\000\000\000\000\082\000\ -\083\000\000\000\000\000\084\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\215\000\087\000\000\000\089\000\000\000\215\000\ -\215\000\090\000\091\000\092\000\093\000\224\000\000\000\000\000\ -\000\000\000\000\215\000\215\000\215\000\215\000\215\000\000\000\ -\000\000\000\000\000\000\215\000\000\000\000\000\000\000\215\000\ -\000\000\000\000\215\000\000\000\215\000\000\000\000\000\000\000\ -\000\000\000\000\215\000\215\000\000\000\000\000\000\000\000\000\ -\215\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\217\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\215\000\041\000\000\000\000\000\215\000\000\000\ -\000\000\000\000\215\000\000\000\000\000\043\000\000\000\000\000\ -\000\000\215\000\215\000\000\000\215\000\215\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\047\000\215\000\000\000\215\000\ -\000\000\218\000\141\000\215\000\215\000\215\000\215\000\215\000\ -\000\000\000\000\000\000\000\000\055\000\056\000\057\000\058\000\ -\059\000\000\000\000\000\000\000\000\000\219\000\000\000\000\000\ -\000\000\221\000\000\000\000\000\064\000\000\000\222\000\000\000\ -\000\000\000\000\000\000\000\000\237\000\070\000\000\000\000\000\ -\000\000\000\000\074\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\217\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\078\000\041\000\000\000\000\000\ -\079\000\000\000\000\000\000\000\080\000\000\000\000\000\043\000\ -\000\000\000\000\000\000\082\000\083\000\000\000\000\000\084\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\047\000\087\000\ -\000\000\089\000\000\000\218\000\141\000\090\000\091\000\092\000\ -\093\000\224\000\000\000\000\000\000\000\000\000\055\000\056\000\ -\057\000\058\000\059\000\000\000\000\000\000\000\000\000\219\000\ -\000\000\000\000\000\000\221\000\000\000\000\000\064\000\000\000\ -\222\000\000\000\000\000\000\000\000\000\000\000\217\002\070\000\ -\000\000\000\000\000\000\000\000\074\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\037\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\078\000\037\000\ -\000\000\000\000\079\000\000\000\000\000\000\000\080\000\000\000\ -\000\000\037\000\000\000\000\000\000\000\082\000\083\000\000\000\ +\086\000\087\000\000\000\000\000\088\000\000\000\252\000\051\000\ +\253\000\000\000\090\000\000\000\091\000\056\000\093\000\000\000\ +\000\000\000\000\094\000\095\000\096\000\097\000\188\000\059\000\ +\060\000\061\000\062\000\063\000\000\000\000\000\000\000\000\000\ +\065\000\000\000\000\000\000\000\067\000\000\000\000\000\000\000\ +\069\000\070\000\000\000\000\000\000\000\000\000\000\000\146\000\ +\074\000\000\000\000\000\077\000\000\000\078\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\079\000\000\000\000\000\ +\080\000\038\000\039\000\040\000\041\000\042\000\000\000\000\000\ +\000\000\043\000\000\000\000\000\045\000\081\000\000\000\082\000\ +\000\000\000\000\000\000\083\000\000\000\000\000\000\000\084\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\086\000\087\000\ +\000\000\000\000\088\000\000\000\000\000\051\000\000\000\000\000\ +\090\000\000\000\091\000\056\000\093\000\000\000\000\000\000\000\ +\094\000\095\000\096\000\097\000\188\000\059\000\060\000\061\000\ +\062\000\063\000\000\000\000\000\000\000\000\000\065\000\000\000\ +\000\000\000\000\067\000\000\000\000\000\068\000\069\000\070\000\ +\000\000\000\000\000\000\000\000\000\000\146\000\074\000\000\000\ +\000\000\077\000\000\000\078\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\079\000\000\000\000\000\080\000\174\001\ +\175\001\176\001\177\001\042\000\000\000\000\000\000\000\043\000\ +\000\000\000\000\045\000\081\000\000\000\082\000\000\000\000\000\ +\000\000\083\000\000\000\000\000\000\000\084\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\086\000\087\000\000\000\000\000\ +\088\000\000\000\000\000\051\000\000\000\000\000\090\000\000\000\ +\091\000\145\000\093\000\000\000\000\000\000\000\094\000\095\000\ +\096\000\097\000\188\000\059\000\060\000\061\000\062\000\063\000\ +\000\000\000\000\000\000\000\000\178\001\000\000\000\000\000\000\ +\067\000\000\000\000\000\068\000\069\000\070\000\000\000\000\000\ +\000\000\000\000\000\000\146\000\074\000\000\000\000\000\180\001\ +\000\000\078\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\079\000\000\000\000\000\080\000\038\000\039\000\040\000\ +\041\000\042\000\000\000\000\000\000\000\043\000\000\000\000\000\ +\045\000\181\001\000\000\082\000\000\000\000\000\000\000\083\000\ +\000\000\000\000\000\000\084\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\086\000\087\000\000\000\000\000\088\000\000\000\ +\000\000\051\000\000\000\000\000\182\001\000\000\091\000\056\000\ +\093\000\000\000\000\000\000\000\094\000\095\000\096\000\097\000\ +\147\000\059\000\060\000\061\000\062\000\063\000\000\000\000\000\ +\000\000\000\000\065\000\000\000\000\000\000\000\067\000\000\000\ +\000\000\000\000\069\000\070\000\000\000\000\000\000\000\000\000\ +\000\000\146\000\074\000\000\000\000\000\077\000\000\000\078\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\079\000\ +\000\000\000\000\080\000\174\001\175\001\176\001\177\001\042\000\ +\000\000\000\000\000\000\043\000\000\000\000\000\045\000\081\000\ +\000\000\082\000\000\000\000\000\000\000\083\000\000\000\000\000\ \000\000\084\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\037\000\087\000\000\000\089\000\000\000\037\000\037\000\090\000\ -\091\000\092\000\093\000\224\000\000\000\000\000\000\000\000\000\ -\037\000\037\000\037\000\037\000\037\000\000\000\000\000\000\000\ -\000\000\037\000\000\000\000\000\000\000\037\000\000\000\000\000\ -\037\000\000\000\037\000\000\000\000\000\000\000\000\000\000\000\ -\037\000\037\000\000\000\000\000\000\000\000\000\037\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\038\000\000\000\000\000\000\000\039\000\000\000\ -\037\000\041\000\000\000\000\000\037\000\000\000\000\000\000\000\ -\037\000\000\000\000\000\043\000\000\000\000\000\000\000\037\000\ -\037\000\000\000\000\000\037\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\047\000\037\000\000\000\037\000\000\000\051\000\ -\141\000\037\000\037\000\037\000\037\000\037\000\000\000\000\000\ -\000\000\000\000\055\000\056\000\057\000\058\000\059\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\063\000\ -\000\000\000\000\064\000\065\000\066\000\000\000\000\000\000\000\ -\000\000\000\000\042\004\070\000\000\000\000\000\000\000\000\000\ -\074\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\075\000\000\000\000\000\076\000\000\000\000\000\000\000\000\000\ -\038\000\000\000\000\000\000\000\039\000\000\000\000\000\041\000\ -\000\000\000\000\078\000\000\000\000\000\000\000\079\000\000\000\ +\086\000\087\000\000\000\000\000\088\000\000\000\000\000\051\000\ +\000\000\000\000\090\000\000\000\091\000\145\000\093\000\000\000\ +\000\000\000\000\094\000\095\000\096\000\097\000\188\000\059\000\ +\060\000\061\000\062\000\063\000\000\000\000\000\000\000\000\000\ +\178\001\000\000\000\000\000\000\067\000\000\000\000\000\000\000\ +\069\000\070\000\000\000\000\000\000\000\000\000\000\000\146\000\ +\074\000\000\000\000\000\180\001\000\000\078\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\079\000\000\000\000\000\ +\080\000\038\000\039\000\040\000\041\000\042\000\000\000\000\000\ +\000\000\043\000\000\000\000\000\045\000\181\001\000\000\082\000\ +\000\000\000\000\000\000\083\000\000\000\000\000\000\000\084\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\086\000\087\000\ +\000\000\000\000\088\000\000\000\000\000\051\000\000\000\000\000\ +\182\001\000\000\091\000\056\000\093\000\000\000\000\000\000\000\ +\094\000\095\000\096\000\097\000\147\000\059\000\060\000\061\000\ +\062\000\063\000\000\000\000\000\000\000\000\000\054\004\000\000\ +\000\000\000\000\067\000\000\000\000\000\000\000\069\000\070\000\ +\000\000\000\000\000\000\000\000\000\000\146\000\074\000\000\000\ +\000\000\077\000\000\000\078\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\079\000\000\000\000\000\080\000\219\000\ +\219\000\219\000\219\000\219\000\000\000\000\000\000\000\219\000\ +\000\000\000\000\219\000\081\000\000\000\082\000\000\000\000\000\ +\000\000\083\000\000\000\000\000\000\000\084\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\086\000\087\000\000\000\000\000\ +\088\000\000\000\000\000\219\000\000\000\000\000\090\000\000\000\ +\091\000\219\000\093\000\000\000\000\000\000\000\094\000\095\000\ +\096\000\097\000\188\000\219\000\219\000\219\000\219\000\219\000\ +\000\000\000\000\000\000\000\000\219\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\219\000\219\000\219\000\000\000\000\000\ +\000\000\000\000\000\000\219\000\219\000\000\000\000\000\219\000\ +\000\000\219\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\219\000\000\000\000\000\219\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\219\000\000\000\219\000\000\000\000\000\000\000\219\000\ +\000\000\042\000\000\000\219\000\000\000\043\000\000\000\000\000\ +\045\000\000\000\219\000\219\000\000\000\000\000\219\000\000\000\ +\000\000\000\000\000\000\000\000\219\000\000\000\219\000\000\000\ +\219\000\000\000\000\000\000\000\219\000\219\000\219\000\219\000\ +\219\000\051\000\000\000\000\000\000\000\000\000\100\001\145\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\059\000\060\000\061\000\062\000\063\000\000\000\000\000\ +\000\000\000\000\101\001\000\000\000\000\000\000\067\000\000\000\ +\000\000\000\000\069\000\070\000\000\000\000\000\000\000\000\000\ +\000\000\146\000\074\000\000\000\000\000\000\000\000\000\078\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\079\000\ \000\000\000\000\080\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\082\000\083\000\000\000\000\000\084\000\000\000\000\000\ -\047\000\000\000\000\000\000\000\000\000\087\000\141\000\089\000\ -\000\000\000\000\000\000\090\000\091\000\092\000\093\000\143\000\ -\055\000\056\000\057\000\058\000\059\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\063\000\000\000\000\000\ -\000\000\065\000\066\000\000\000\000\000\000\000\000\000\000\000\ -\142\000\070\000\000\000\000\000\000\000\000\000\074\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\039\001\000\000\ -\000\000\076\000\000\000\000\000\000\000\000\000\038\000\000\000\ -\000\000\000\000\039\000\000\000\000\000\041\000\000\000\000\000\ -\078\000\000\000\000\000\000\000\079\000\000\000\000\000\000\000\ -\080\000\000\000\000\000\000\000\012\002\000\000\000\000\082\000\ -\083\000\000\000\000\000\084\000\000\000\000\000\047\000\000\000\ -\000\000\000\000\000\000\087\000\141\000\089\000\000\000\000\000\ -\000\000\090\000\091\000\092\000\093\000\143\000\055\000\056\000\ -\057\000\058\000\059\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\063\000\000\000\000\000\000\000\065\000\ -\066\000\000\000\000\000\000\000\000\000\000\000\142\000\070\000\ -\000\000\000\000\000\000\000\000\074\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\075\000\000\000\000\000\076\000\ -\000\000\000\000\000\000\000\000\038\000\000\000\000\000\000\000\ -\039\000\000\000\000\000\041\000\000\000\000\000\078\000\000\000\ -\000\000\000\000\079\000\000\000\000\000\000\000\080\000\000\000\ -\000\000\000\000\013\002\000\000\000\000\082\000\083\000\000\000\ -\000\000\084\000\000\000\000\000\047\000\000\000\000\000\000\000\ -\000\000\087\000\141\000\089\000\000\000\000\000\000\000\090\000\ -\091\000\092\000\093\000\143\000\055\000\056\000\057\000\058\000\ -\059\000\000\000\000\000\000\000\000\000\072\002\000\000\000\000\ -\000\000\063\000\000\000\000\000\000\000\065\000\066\000\000\000\ -\000\000\000\000\000\000\000\000\142\000\070\000\000\000\000\000\ -\000\000\000\000\074\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\075\000\000\000\000\000\076\000\000\000\000\000\ -\000\000\000\000\038\000\000\000\000\000\000\000\039\000\000\000\ -\000\000\041\000\000\000\000\000\078\000\000\000\000\000\000\000\ -\079\000\000\000\000\000\000\000\080\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\082\000\083\000\000\000\000\000\084\000\ -\000\000\000\000\047\000\000\000\000\000\000\000\000\000\087\000\ -\141\000\089\000\000\000\000\000\000\000\090\000\091\000\092\000\ -\093\000\143\000\055\000\056\000\057\000\058\000\059\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\063\000\ -\000\000\000\000\000\000\065\000\066\000\000\000\000\000\000\000\ -\000\000\000\000\142\000\070\000\000\000\000\000\000\000\000\000\ -\074\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\039\001\000\000\000\000\076\000\000\000\000\000\000\000\000\000\ -\038\000\000\000\000\000\000\000\039\000\000\000\000\000\041\000\ -\000\000\000\000\078\000\000\000\000\000\000\000\079\000\000\000\ -\000\000\000\000\080\000\000\000\000\000\000\000\201\002\000\000\ -\000\000\082\000\083\000\000\000\000\000\084\000\000\000\000\000\ -\047\000\000\000\000\000\000\000\000\000\087\000\141\000\089\000\ -\000\000\000\000\000\000\090\000\091\000\092\000\093\000\143\000\ -\055\000\056\000\057\000\058\000\059\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\063\000\000\000\000\000\ -\000\000\065\000\066\000\000\000\000\000\000\000\000\000\000\000\ -\142\000\070\000\000\000\000\000\000\000\000\000\074\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\075\000\000\000\ -\000\000\076\000\000\000\000\000\000\000\000\000\038\000\000\000\ -\000\000\000\000\039\000\000\000\000\000\041\000\000\000\000\000\ -\078\000\000\000\000\000\000\000\079\000\000\000\000\000\000\000\ -\080\000\000\000\000\000\000\000\202\002\000\000\000\000\082\000\ -\083\000\000\000\000\000\084\000\000\000\000\000\047\000\000\000\ -\000\000\000\000\000\000\087\000\141\000\089\000\000\000\000\000\ -\000\000\090\000\091\000\092\000\093\000\143\000\055\000\056\000\ -\057\000\058\000\059\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\063\000\000\000\000\000\000\000\065\000\ -\066\000\000\000\000\000\000\000\000\000\000\000\142\000\070\000\ -\000\000\000\000\000\000\000\000\074\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\039\001\000\000\000\000\076\000\ -\000\000\000\000\000\000\000\000\038\000\000\000\000\000\000\000\ -\039\000\000\000\000\000\041\000\000\000\000\000\078\000\000\000\ -\000\000\000\000\079\000\000\000\000\000\000\000\080\000\000\000\ -\000\000\000\000\125\004\000\000\000\000\082\000\083\000\000\000\ +\000\000\000\000\000\000\042\000\000\000\000\000\000\000\043\000\ +\000\000\082\000\045\000\000\000\000\000\083\000\000\000\000\000\ +\000\000\084\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\086\000\087\000\000\000\000\000\088\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\051\000\091\000\000\000\093\000\000\000\ +\161\002\145\000\094\000\095\000\096\000\097\000\147\000\000\000\ +\000\000\102\001\000\000\059\000\060\000\061\000\062\000\063\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\067\000\000\000\000\000\000\000\069\000\070\000\000\000\000\000\ +\000\000\000\000\000\000\146\000\074\000\000\000\000\000\000\000\ +\000\000\078\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\079\000\000\000\000\000\080\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\041\000\042\000\000\000\000\000\ +\000\000\043\000\000\000\082\000\045\000\000\000\000\000\083\000\ +\000\000\000\000\000\000\084\000\000\000\000\000\047\000\000\000\ +\000\000\000\000\086\000\087\000\000\000\000\000\088\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\051\000\091\000\000\000\ +\093\000\000\000\055\000\056\000\094\000\095\000\096\000\097\000\ +\147\000\000\000\000\000\102\001\000\000\059\000\060\000\061\000\ +\062\000\063\000\000\000\000\000\000\000\000\000\065\000\226\002\ +\000\000\000\000\067\000\000\000\000\000\068\000\069\000\070\000\ +\000\000\000\000\000\000\000\000\000\000\227\002\074\000\000\000\ +\000\000\000\000\000\000\078\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\079\000\000\000\000\000\080\000\000\000\ +\000\000\000\000\041\000\042\000\000\000\000\000\000\000\043\000\ +\000\000\000\000\045\000\000\000\000\000\082\000\000\000\000\000\ +\000\000\083\000\000\000\000\000\000\000\084\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\086\000\087\000\000\000\000\000\ +\088\000\000\000\000\000\051\000\000\000\000\000\090\000\000\000\ +\091\000\056\000\093\000\000\000\000\000\000\000\094\000\095\000\ +\096\000\097\000\188\000\059\000\060\000\061\000\062\000\063\000\ +\000\000\000\000\000\000\000\000\065\000\000\000\000\000\000\000\ +\067\000\000\000\000\000\068\000\069\000\070\000\000\000\000\000\ +\000\000\000\000\000\000\146\000\074\000\000\000\000\000\000\000\ +\000\000\078\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\079\000\000\000\000\000\080\000\000\000\000\000\000\000\ +\041\000\042\000\000\000\000\000\000\000\043\000\000\000\000\000\ +\045\000\000\000\000\000\082\000\000\000\000\000\000\000\083\000\ +\000\000\000\000\000\000\084\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\086\000\087\000\000\000\000\000\088\000\000\000\ +\000\000\051\000\000\000\000\000\090\000\000\000\091\000\056\000\ +\093\000\000\000\000\000\000\000\094\000\095\000\096\000\097\000\ +\188\000\059\000\060\000\061\000\062\000\063\000\000\000\000\000\ +\000\000\000\000\065\000\000\000\000\000\000\000\067\000\000\000\ +\000\000\000\000\069\000\070\000\000\000\000\000\000\000\000\000\ +\000\000\146\000\074\000\000\000\000\000\000\000\000\000\078\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\079\000\ +\000\000\000\000\080\000\000\000\000\000\000\000\177\001\042\000\ +\000\000\000\000\000\000\043\000\000\000\000\000\045\000\000\000\ +\000\000\082\000\000\000\000\000\000\000\083\000\000\000\000\000\ +\000\000\084\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\086\000\087\000\000\000\000\000\088\000\000\000\000\000\051\000\ +\000\000\000\000\090\000\000\000\091\000\145\000\093\000\000\000\ +\000\000\000\000\094\000\095\000\096\000\097\000\188\000\059\000\ +\060\000\061\000\062\000\063\000\000\000\000\000\000\000\000\000\ +\178\001\000\000\000\000\000\000\067\000\000\000\000\000\000\000\ +\069\000\070\000\000\000\000\000\000\000\000\000\000\000\146\000\ +\074\000\000\000\000\000\000\000\000\000\078\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\079\000\000\000\000\000\ +\080\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\221\000\000\000\240\000\000\000\000\000\000\000\000\000\082\000\ +\045\000\000\000\000\000\083\000\000\000\000\000\000\000\084\000\ +\000\000\000\000\047\000\000\000\000\000\000\000\086\000\087\000\ +\000\000\000\000\088\000\000\000\000\000\000\000\000\000\000\000\ +\182\001\051\000\091\000\000\000\093\000\000\000\222\000\145\000\ +\094\000\095\000\096\000\097\000\147\000\000\000\000\000\000\000\ +\000\000\059\000\060\000\061\000\062\000\063\000\000\000\000\000\ +\000\000\000\000\223\000\000\000\000\000\000\000\225\000\000\000\ +\000\000\068\000\000\000\226\000\000\000\000\000\000\000\000\000\ +\000\000\241\000\074\000\000\000\000\000\000\000\000\000\078\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\221\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\082\000\045\000\000\000\000\000\083\000\000\000\000\000\ \000\000\084\000\000\000\000\000\047\000\000\000\000\000\000\000\ -\000\000\087\000\141\000\089\000\000\000\000\000\000\000\090\000\ -\091\000\092\000\093\000\143\000\055\000\056\000\057\000\058\000\ -\059\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\063\000\000\000\000\000\000\000\065\000\066\000\000\000\ -\000\000\000\000\000\000\000\000\142\000\070\000\000\000\000\000\ -\000\000\000\000\074\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\075\000\000\000\000\000\076\000\000\000\000\000\ -\000\000\000\000\038\000\000\000\000\000\000\000\039\000\000\000\ -\000\000\041\000\000\000\000\000\078\000\000\000\000\000\000\000\ -\079\000\000\000\000\000\000\000\080\000\000\000\000\000\000\000\ -\126\004\000\000\000\000\082\000\083\000\000\000\000\000\084\000\ -\000\000\000\000\047\000\000\000\000\000\000\000\000\000\087\000\ -\141\000\089\000\000\000\000\000\000\000\090\000\091\000\092\000\ -\093\000\143\000\055\000\056\000\057\000\058\000\059\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\063\000\ -\000\000\000\000\000\000\065\000\066\000\000\000\000\000\000\000\ -\000\000\000\000\142\000\070\000\000\000\000\000\000\000\000\000\ -\074\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\075\000\000\000\000\000\076\000\000\000\000\000\000\000\000\000\ -\038\000\000\000\000\000\000\000\039\000\000\000\000\000\041\000\ -\000\000\000\000\078\000\000\000\000\000\000\000\079\000\000\000\ -\000\000\000\000\080\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\082\000\083\000\000\000\000\000\084\000\000\000\000\000\ -\047\000\000\000\000\000\000\000\000\000\087\000\141\000\089\000\ -\000\000\000\000\000\000\090\000\091\000\092\000\093\000\143\000\ -\055\000\056\000\057\000\058\000\059\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\063\000\000\000\000\000\ -\000\000\065\000\066\000\000\000\000\000\000\000\000\000\000\000\ -\252\000\070\000\000\000\000\000\000\000\000\000\074\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\075\000\000\000\ -\000\000\076\000\000\000\000\000\000\000\000\000\038\000\000\000\ -\000\000\000\000\039\000\000\000\000\000\041\000\000\000\000\000\ -\078\000\000\000\000\000\000\000\079\000\000\000\000\000\000\000\ -\080\000\000\000\000\000\000\000\000\000\000\000\000\000\082\000\ -\083\000\000\000\000\000\084\000\000\000\000\000\047\000\000\000\ -\000\000\000\000\000\000\087\000\141\000\089\000\000\000\000\000\ -\000\000\090\000\091\000\092\000\093\000\143\000\055\000\056\000\ -\057\000\058\000\059\000\000\000\000\000\106\002\000\000\000\000\ -\000\000\000\000\000\000\063\000\000\000\000\000\000\000\065\000\ -\066\000\000\000\000\000\000\000\000\000\000\000\142\000\070\000\ -\014\001\015\001\000\000\016\001\074\000\000\000\017\001\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\090\003\ -\000\000\000\000\018\001\000\000\039\001\000\000\000\000\076\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\019\001\000\000\ -\020\001\000\000\014\001\015\001\000\000\016\001\078\000\000\000\ -\017\001\000\000\079\000\000\000\000\000\000\000\080\000\000\000\ -\000\000\000\000\000\000\000\000\018\001\082\000\083\000\000\000\ -\000\000\084\000\000\000\107\002\000\000\000\000\000\000\000\000\ -\019\001\087\000\020\001\089\000\000\000\000\000\000\000\090\000\ -\091\000\092\000\093\000\143\000\022\001\023\001\024\001\025\001\ -\026\001\027\001\028\001\029\001\108\002\031\001\032\001\000\000\ -\033\001\000\000\000\000\000\000\000\000\091\003\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\128\004\000\000\000\000\000\000\000\000\000\000\022\001\023\001\ -\024\001\025\001\026\001\027\001\028\001\029\001\092\003\031\001\ -\032\001\000\000\033\001\014\001\015\001\000\000\016\001\109\002\ -\000\000\017\001\000\000\000\000\136\004\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\018\001\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\014\001\ -\015\001\019\001\016\001\020\001\000\000\017\001\000\000\000\000\ -\000\000\093\003\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\018\001\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\019\001\129\004\020\001\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\022\001\ -\023\001\024\001\025\001\026\001\027\001\028\001\029\001\130\004\ -\031\001\032\001\137\004\033\001\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\160\004\000\000\000\000\ -\000\000\000\000\000\000\022\001\023\001\024\001\025\001\026\001\ -\027\001\028\001\029\001\138\004\031\001\032\001\000\000\033\001\ -\014\001\015\001\000\000\016\001\000\000\000\000\017\001\000\000\ -\000\000\177\004\131\004\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\018\001\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\014\001\015\001\019\001\016\001\ -\020\001\000\000\017\001\000\000\000\000\000\000\139\004\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\018\001\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\019\001\161\004\020\001\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\022\001\023\001\024\001\025\001\ -\026\001\027\001\028\001\029\001\162\004\031\001\032\001\178\004\ -\033\001\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\022\001\023\001\024\001\025\001\026\001\027\001\028\001\029\001\ -\179\004\031\001\032\001\000\000\033\001\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\163\004\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ -\000\000\000\000\000\000\180\004" +\086\000\087\000\000\000\000\000\088\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\051\000\091\000\000\000\093\000\000\000\ +\222\000\145\000\094\000\095\000\096\000\097\000\228\000\000\000\ +\000\000\000\000\000\000\059\000\060\000\061\000\062\000\063\000\ +\000\000\000\000\000\000\000\000\223\000\000\000\000\000\000\000\ +\225\000\000\000\000\000\068\000\000\000\226\000\000\000\000\000\ +\000\000\000\000\000\000\241\000\074\000\000\000\000\000\000\000\ +\000\000\078\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\221\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\082\000\045\000\000\000\000\000\083\000\ +\000\000\000\000\000\000\084\000\000\000\000\000\047\000\015\002\ +\000\000\000\000\086\000\087\000\000\000\000\000\088\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\051\000\091\000\000\000\ +\093\000\000\000\222\000\145\000\094\000\095\000\096\000\097\000\ +\228\000\000\000\000\000\000\000\000\000\059\000\060\000\061\000\ +\062\000\063\000\000\000\000\000\000\000\000\000\223\000\000\000\ +\000\000\000\000\225\000\000\000\000\000\068\000\000\000\226\000\ +\000\000\000\000\000\000\000\000\000\000\241\000\074\000\000\000\ +\000\000\000\000\000\000\078\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\221\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\082\000\045\000\000\000\ +\000\000\083\000\000\000\000\000\000\000\084\000\000\000\000\000\ +\047\000\204\002\000\000\000\000\086\000\087\000\000\000\000\000\ +\088\000\000\000\000\000\000\000\000\000\000\000\000\000\051\000\ +\091\000\000\000\093\000\000\000\222\000\145\000\094\000\095\000\ +\096\000\097\000\228\000\000\000\000\000\000\000\000\000\059\000\ +\060\000\061\000\062\000\063\000\000\000\000\000\000\000\000\000\ +\223\000\000\000\000\000\000\000\225\000\000\000\000\000\068\000\ +\000\000\226\000\000\000\000\000\000\000\000\000\000\000\241\000\ +\074\000\000\000\000\000\000\000\000\000\078\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\217\000\000\000\000\000\000\000\000\000\000\000\000\000\082\000\ +\217\000\000\000\000\000\083\000\000\000\000\000\000\000\084\000\ +\000\000\000\000\217\000\128\004\000\000\000\000\086\000\087\000\ +\000\000\000\000\088\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\217\000\091\000\000\000\093\000\000\000\217\000\217\000\ +\094\000\095\000\096\000\097\000\228\000\000\000\000\000\000\000\ +\000\000\217\000\217\000\217\000\217\000\217\000\000\000\000\000\ +\000\000\000\000\217\000\000\000\000\000\000\000\217\000\000\000\ +\000\000\217\000\000\000\217\000\000\000\000\000\000\000\000\000\ +\000\000\217\000\217\000\000\000\000\000\000\000\000\000\217\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\221\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\217\000\045\000\000\000\000\000\217\000\000\000\000\000\ +\000\000\217\000\000\000\000\000\047\000\000\000\000\000\000\000\ +\217\000\217\000\000\000\217\000\217\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\051\000\217\000\000\000\217\000\000\000\ +\222\000\145\000\217\000\217\000\217\000\217\000\217\000\000\000\ +\000\000\000\000\000\000\059\000\060\000\061\000\062\000\063\000\ +\000\000\000\000\000\000\000\000\223\000\000\000\000\000\000\000\ +\225\000\000\000\000\000\068\000\000\000\226\000\000\000\000\000\ +\000\000\000\000\000\000\241\000\074\000\000\000\000\000\000\000\ +\000\000\078\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\221\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\082\000\045\000\000\000\000\000\083\000\ +\000\000\000\000\000\000\084\000\000\000\000\000\047\000\000\000\ +\000\000\000\000\086\000\087\000\000\000\000\000\088\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\051\000\091\000\000\000\ +\093\000\000\000\222\000\145\000\094\000\095\000\096\000\097\000\ +\228\000\000\000\000\000\000\000\000\000\059\000\060\000\061\000\ +\062\000\063\000\000\000\000\000\000\000\000\000\223\000\000\000\ +\000\000\000\000\225\000\000\000\000\000\068\000\000\000\226\000\ +\000\000\000\000\000\000\000\000\000\000\221\002\074\000\000\000\ +\000\000\000\000\000\000\078\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\037\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\082\000\037\000\000\000\ +\000\000\083\000\000\000\000\000\000\000\084\000\000\000\000\000\ +\037\000\000\000\000\000\000\000\086\000\087\000\000\000\000\000\ +\088\000\000\000\000\000\000\000\000\000\000\000\000\000\037\000\ +\091\000\000\000\093\000\000\000\037\000\037\000\094\000\095\000\ +\096\000\097\000\228\000\000\000\000\000\000\000\000\000\037\000\ +\037\000\037\000\037\000\037\000\000\000\000\000\000\000\000\000\ +\037\000\000\000\000\000\000\000\037\000\000\000\000\000\037\000\ +\000\000\037\000\000\000\000\000\000\000\000\000\000\000\037\000\ +\037\000\000\000\000\000\000\000\000\000\037\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\042\000\000\000\000\000\000\000\043\000\000\000\037\000\ +\045\000\000\000\000\000\037\000\000\000\000\000\000\000\037\000\ +\000\000\000\000\047\000\000\000\000\000\000\000\037\000\037\000\ +\000\000\000\000\037\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\051\000\037\000\000\000\037\000\000\000\055\000\145\000\ +\037\000\037\000\037\000\037\000\037\000\000\000\000\000\000\000\ +\000\000\059\000\060\000\061\000\062\000\063\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\067\000\000\000\ +\000\000\068\000\069\000\070\000\000\000\000\000\000\000\000\000\ +\000\000\046\004\074\000\000\000\000\000\000\000\000\000\078\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\079\000\ +\000\000\000\000\080\000\000\000\000\000\000\000\000\000\042\000\ +\000\000\000\000\000\000\043\000\000\000\000\000\045\000\000\000\ +\000\000\082\000\000\000\000\000\000\000\083\000\000\000\000\000\ +\000\000\084\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\086\000\087\000\000\000\000\000\088\000\000\000\000\000\051\000\ +\000\000\000\000\000\000\000\000\091\000\145\000\093\000\000\000\ +\000\000\000\000\094\000\095\000\096\000\097\000\147\000\059\000\ +\060\000\061\000\062\000\063\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\067\000\000\000\000\000\000\000\ +\069\000\070\000\000\000\000\000\000\000\000\000\000\000\146\000\ +\074\000\000\000\000\000\000\000\000\000\078\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\043\001\000\000\000\000\ +\080\000\000\000\000\000\000\000\000\000\042\000\000\000\000\000\ +\000\000\043\000\000\000\000\000\045\000\000\000\000\000\082\000\ +\000\000\000\000\000\000\083\000\000\000\000\000\000\000\084\000\ +\000\000\000\000\000\000\016\002\000\000\000\000\086\000\087\000\ +\000\000\000\000\088\000\000\000\000\000\051\000\000\000\000\000\ +\000\000\000\000\091\000\145\000\093\000\000\000\000\000\000\000\ +\094\000\095\000\096\000\097\000\147\000\059\000\060\000\061\000\ +\062\000\063\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\067\000\000\000\000\000\000\000\069\000\070\000\ +\000\000\000\000\000\000\000\000\000\000\146\000\074\000\000\000\ +\000\000\000\000\000\000\078\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\079\000\000\000\000\000\080\000\000\000\ +\000\000\000\000\000\000\042\000\000\000\000\000\000\000\043\000\ +\000\000\000\000\045\000\000\000\000\000\082\000\000\000\000\000\ +\000\000\083\000\000\000\000\000\000\000\084\000\000\000\000\000\ +\000\000\017\002\000\000\000\000\086\000\087\000\000\000\000\000\ +\088\000\000\000\000\000\051\000\000\000\000\000\000\000\000\000\ +\091\000\145\000\093\000\000\000\000\000\000\000\094\000\095\000\ +\096\000\097\000\147\000\059\000\060\000\061\000\062\000\063\000\ +\000\000\000\000\000\000\000\000\076\002\000\000\000\000\000\000\ +\067\000\000\000\000\000\000\000\069\000\070\000\000\000\000\000\ +\000\000\000\000\000\000\146\000\074\000\000\000\000\000\000\000\ +\000\000\078\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\079\000\000\000\000\000\080\000\000\000\000\000\000\000\ +\000\000\042\000\000\000\000\000\000\000\043\000\000\000\000\000\ +\045\000\000\000\000\000\082\000\000\000\000\000\000\000\083\000\ +\000\000\000\000\000\000\084\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\086\000\087\000\000\000\000\000\088\000\000\000\ +\000\000\051\000\000\000\000\000\000\000\000\000\091\000\145\000\ +\093\000\000\000\000\000\000\000\094\000\095\000\096\000\097\000\ +\147\000\059\000\060\000\061\000\062\000\063\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\067\000\000\000\ +\000\000\000\000\069\000\070\000\000\000\000\000\000\000\000\000\ +\000\000\146\000\074\000\000\000\000\000\000\000\000\000\078\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\043\001\ +\000\000\000\000\080\000\000\000\000\000\000\000\000\000\042\000\ +\000\000\000\000\000\000\043\000\000\000\000\000\045\000\000\000\ +\000\000\082\000\000\000\000\000\000\000\083\000\000\000\000\000\ +\000\000\084\000\000\000\000\000\000\000\205\002\000\000\000\000\ +\086\000\087\000\000\000\000\000\088\000\000\000\000\000\051\000\ +\000\000\000\000\000\000\000\000\091\000\145\000\093\000\000\000\ +\000\000\000\000\094\000\095\000\096\000\097\000\147\000\059\000\ +\060\000\061\000\062\000\063\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\067\000\000\000\000\000\000\000\ +\069\000\070\000\000\000\000\000\000\000\000\000\000\000\146\000\ +\074\000\000\000\000\000\000\000\000\000\078\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\079\000\000\000\000\000\ +\080\000\000\000\000\000\000\000\000\000\042\000\000\000\000\000\ +\000\000\043\000\000\000\000\000\045\000\000\000\000\000\082\000\ +\000\000\000\000\000\000\083\000\000\000\000\000\000\000\084\000\ +\000\000\000\000\000\000\206\002\000\000\000\000\086\000\087\000\ +\000\000\000\000\088\000\000\000\000\000\051\000\000\000\000\000\ +\000\000\000\000\091\000\145\000\093\000\000\000\000\000\000\000\ +\094\000\095\000\096\000\097\000\147\000\059\000\060\000\061\000\ +\062\000\063\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\067\000\000\000\000\000\000\000\069\000\070\000\ +\000\000\000\000\000\000\000\000\000\000\146\000\074\000\000\000\ +\000\000\000\000\000\000\078\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\043\001\000\000\000\000\080\000\000\000\ +\000\000\000\000\000\000\042\000\000\000\000\000\000\000\043\000\ +\000\000\000\000\045\000\000\000\000\000\082\000\000\000\000\000\ +\000\000\083\000\000\000\000\000\000\000\084\000\000\000\000\000\ +\000\000\129\004\000\000\000\000\086\000\087\000\000\000\000\000\ +\088\000\000\000\000\000\051\000\000\000\000\000\000\000\000\000\ +\091\000\145\000\093\000\000\000\000\000\000\000\094\000\095\000\ +\096\000\097\000\147\000\059\000\060\000\061\000\062\000\063\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\067\000\000\000\000\000\000\000\069\000\070\000\000\000\000\000\ +\000\000\000\000\000\000\146\000\074\000\000\000\000\000\000\000\ +\000\000\078\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\079\000\000\000\000\000\080\000\000\000\000\000\000\000\ +\000\000\042\000\000\000\000\000\000\000\043\000\000\000\000\000\ +\045\000\000\000\000\000\082\000\000\000\000\000\000\000\083\000\ +\000\000\000\000\000\000\084\000\000\000\000\000\000\000\130\004\ +\000\000\000\000\086\000\087\000\000\000\000\000\088\000\000\000\ +\000\000\051\000\000\000\000\000\000\000\000\000\091\000\145\000\ +\093\000\000\000\000\000\000\000\094\000\095\000\096\000\097\000\ +\147\000\059\000\060\000\061\000\062\000\063\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\067\000\000\000\ +\000\000\000\000\069\000\070\000\000\000\000\000\000\000\000\000\ +\000\000\146\000\074\000\000\000\000\000\000\000\000\000\078\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\079\000\ +\000\000\000\000\080\000\000\000\000\000\000\000\000\000\042\000\ +\000\000\000\000\000\000\043\000\000\000\000\000\045\000\000\000\ +\000\000\082\000\000\000\000\000\000\000\083\000\000\000\000\000\ +\000\000\084\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\086\000\087\000\000\000\000\000\088\000\000\000\000\000\051\000\ +\000\000\000\000\000\000\000\000\091\000\145\000\093\000\000\000\ +\000\000\000\000\094\000\095\000\096\000\097\000\147\000\059\000\ +\060\000\061\000\062\000\063\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\067\000\000\000\000\000\000\000\ +\069\000\070\000\000\000\000\000\000\000\000\000\000\000\000\001\ +\074\000\000\000\000\000\000\000\000\000\078\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\079\000\000\000\000\000\ +\080\000\000\000\000\000\000\000\000\000\042\000\000\000\000\000\ +\000\000\043\000\000\000\000\000\045\000\000\000\000\000\082\000\ +\000\000\000\000\000\000\083\000\000\000\000\000\000\000\084\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\086\000\087\000\ +\000\000\000\000\088\000\000\000\000\000\051\000\000\000\000\000\ +\000\000\000\000\091\000\145\000\093\000\000\000\000\000\000\000\ +\094\000\095\000\096\000\097\000\147\000\059\000\060\000\061\000\ +\062\000\063\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\067\000\000\000\000\000\000\000\069\000\070\000\ +\000\000\000\000\000\000\000\000\000\000\146\000\074\000\000\000\ +\000\000\000\000\000\000\078\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\043\001\000\000\068\000\080\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\068\000\082\000\000\000\000\000\ +\000\000\083\000\000\000\068\000\000\000\084\000\000\000\068\000\ +\000\000\000\000\000\000\000\000\086\000\087\000\000\000\000\000\ +\088\000\068\000\000\000\000\000\068\000\000\000\000\000\000\000\ +\091\000\000\000\093\000\000\000\000\000\068\000\094\000\095\000\ +\096\000\097\000\147\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\068\000\000\000\094\003\000\000\000\000\000\000\000\000\ +\000\000\000\000\068\000\000\000\000\000\068\000\000\000\000\000\ +\000\000\000\000\068\000\000\000\000\000\000\000\018\001\019\001\ +\000\000\020\001\000\000\000\000\021\001\000\000\000\000\132\004\ +\068\000\068\000\068\000\068\000\068\000\068\000\068\000\068\000\ +\022\001\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\018\001\019\001\023\001\020\001\024\001\000\000\ +\021\001\000\000\000\000\000\000\000\000\000\000\068\000\000\000\ +\000\000\000\000\000\000\068\000\022\001\000\000\000\000\000\000\ +\000\000\000\000\000\000\068\000\000\000\000\000\000\000\000\000\ +\023\001\095\003\024\001\000\000\000\000\000\000\000\000\068\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\026\001\027\001\028\001\029\001\030\001\031\001\ +\032\001\033\001\096\003\035\001\036\001\133\004\037\001\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\140\004\000\000\000\000\000\000\000\000\000\000\026\001\027\001\ +\028\001\029\001\030\001\031\001\032\001\033\001\134\004\035\001\ +\036\001\000\000\037\001\018\001\019\001\000\000\020\001\000\000\ +\000\000\021\001\000\000\000\000\164\004\097\003\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\022\001\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\018\001\ +\019\001\023\001\020\001\024\001\000\000\021\001\000\000\000\000\ +\000\000\135\004\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\022\001\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\023\001\141\004\024\001\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\026\001\ +\027\001\028\001\029\001\030\001\031\001\032\001\033\001\142\004\ +\035\001\036\001\165\004\037\001\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\181\004\000\000\000\000\ +\000\000\000\000\000\000\026\001\027\001\028\001\029\001\030\001\ +\031\001\032\001\033\001\166\004\035\001\036\001\000\000\037\001\ +\018\001\019\001\000\000\020\001\000\000\000\000\021\001\000\000\ +\000\000\000\000\143\004\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\022\001\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\023\001\000\000\ +\024\001\000\000\000\000\000\000\000\000\000\000\167\004\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\182\004\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\026\001\027\001\028\001\029\001\ +\030\001\031\001\032\001\033\001\183\004\035\001\036\001\000\000\ +\037\001\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\ +\000\000\000\000\000\000\000\000\000\000\000\000\000\000\184\004" -let yycheck = "\002\000\ -\010\001\247\000\068\000\044\000\049\000\220\000\236\000\010\000\ -\011\000\145\000\075\000\100\001\098\001\046\000\001\000\048\000\ -\016\002\061\000\223\000\060\000\108\000\062\000\011\000\056\001\ -\077\001\145\001\044\000\250\001\031\000\032\000\033\000\229\000\ -\035\000\036\000\037\000\222\000\039\000\040\000\205\000\026\000\ -\097\001\028\000\060\000\013\003\063\000\064\000\065\000\188\001\ -\009\004\034\000\152\002\153\002\154\002\155\002\156\002\054\001\ -\226\001\076\000\111\002\175\001\063\000\064\000\065\000\179\001\ -\180\001\056\001\069\000\078\001\028\001\110\000\111\000\108\003\ -\075\000\076\000\063\000\064\000\065\000\157\000\028\001\159\000\ -\025\001\034\001\085\000\086\000\021\001\088\000\050\001\076\000\ -\073\000\015\001\029\001\042\000\077\000\067\003\045\000\015\001\ -\060\001\094\004\098\002\120\001\121\001\073\001\063\004\028\001\ -\025\001\028\001\050\001\125\001\103\001\083\003\049\001\050\001\ -\115\000\116\000\124\001\042\000\160\000\174\001\045\000\060\001\ -\084\001\125\001\021\001\022\001\142\001\024\001\090\001\116\000\ -\014\001\008\001\081\000\029\001\028\001\038\000\025\003\156\001\ -\014\001\076\001\142\001\142\000\029\001\136\001\145\000\050\001\ -\028\001\097\001\090\001\103\001\118\003\167\001\028\001\049\001\ -\050\001\151\001\081\000\221\000\061\000\223\000\169\001\050\001\ -\049\001\050\001\223\000\167\001\014\001\198\000\199\000\120\001\ -\121\001\060\001\028\001\237\000\070\001\023\001\171\004\219\000\ -\237\000\129\001\076\001\182\000\136\001\070\001\123\001\069\001\ -\229\000\021\001\084\001\076\001\028\001\161\001\160\003\252\000\ -\195\000\084\001\041\001\084\001\014\001\161\001\243\000\136\001\ -\125\001\127\001\125\001\156\001\040\001\023\001\195\000\127\001\ -\168\001\028\001\002\000\021\001\076\002\134\001\217\000\248\000\ -\249\000\142\001\168\001\069\001\194\003\030\001\161\001\168\001\ -\090\001\116\003\105\001\188\004\219\001\125\001\241\001\055\001\ -\029\003\030\003\031\003\032\003\033\003\136\001\035\002\125\001\ -\037\002\125\001\023\004\024\004\025\004\026\004\027\004\039\000\ -\098\001\252\000\136\001\069\001\037\001\223\003\161\001\153\001\ -\255\001\227\003\156\001\233\001\174\002\008\001\028\001\161\001\ -\153\001\214\001\014\001\050\001\126\001\021\001\161\001\072\002\ -\161\001\158\002\014\001\125\001\030\002\137\002\063\001\190\001\ -\098\001\099\001\014\001\030\001\028\001\125\001\136\001\034\001\ -\054\001\125\001\069\001\023\001\039\001\076\001\225\001\000\000\ -\088\000\050\004\145\001\021\001\008\001\048\001\049\001\050\001\ -\051\001\040\001\125\001\125\001\040\001\097\001\057\001\090\001\ -\050\001\215\002\076\001\103\001\030\004\136\001\136\001\066\001\ -\034\004\069\001\003\001\115\000\071\001\072\001\073\001\074\001\ -\081\004\069\001\002\001\083\001\126\001\066\001\069\001\237\002\ -\073\001\069\001\076\001\116\001\186\003\050\001\119\001\137\001\ -\017\001\150\001\084\001\069\001\050\001\112\001\069\001\060\001\ -\081\001\082\001\083\001\084\001\085\001\086\001\087\001\088\001\ -\089\001\090\001\109\001\001\001\078\004\112\001\098\001\099\001\ -\100\001\111\002\008\001\090\001\156\001\095\001\055\001\084\001\ -\109\001\125\001\161\001\112\001\001\001\125\001\084\001\168\001\ -\095\003\125\001\174\001\008\001\136\001\050\001\167\002\050\001\ -\136\001\125\001\103\001\003\001\136\001\105\001\069\001\070\001\ -\070\001\115\004\074\001\075\001\136\001\167\001\134\001\135\001\ -\079\001\156\001\157\001\158\001\159\001\210\002\084\001\028\001\ -\163\001\076\001\026\003\163\003\137\001\138\001\166\003\084\001\ -\171\001\172\001\173\001\136\001\005\003\090\001\163\001\178\001\ -\040\001\243\001\008\003\244\001\024\001\140\002\141\002\228\001\ -\168\001\054\001\076\001\240\002\020\002\050\001\231\001\055\001\ -\097\001\003\001\004\002\005\002\161\001\029\001\001\001\240\001\ -\090\001\160\002\153\001\213\002\248\001\008\001\250\001\001\000\ -\002\000\003\000\104\001\105\001\171\002\028\001\217\001\029\001\ -\008\001\049\001\050\001\022\001\080\001\239\001\181\002\084\001\ -\050\001\030\001\153\001\104\001\105\001\232\001\040\001\180\002\ -\222\002\018\002\224\002\049\001\050\001\227\002\014\001\054\001\ -\028\001\244\001\245\001\025\001\076\001\055\001\249\001\023\001\ -\025\001\021\001\076\001\028\001\084\001\134\001\135\001\052\001\ -\161\001\049\001\050\001\088\001\249\001\002\000\076\001\030\001\ -\040\001\057\001\054\001\148\001\040\001\174\001\054\001\000\000\ -\003\002\102\001\103\001\249\002\218\002\054\001\021\001\071\001\ -\072\001\073\001\074\001\022\001\072\002\024\001\150\003\029\003\ -\030\003\031\003\032\003\033\003\126\001\069\001\031\001\032\001\ -\033\001\034\001\039\000\134\001\135\001\104\001\105\001\137\001\ -\127\001\046\000\030\001\048\000\170\001\040\001\054\004\010\000\ -\089\002\146\003\176\001\177\001\040\001\094\002\022\001\149\003\ -\024\001\003\001\098\001\099\001\100\001\101\001\123\003\127\001\ -\069\000\088\001\067\001\013\001\031\000\032\000\033\000\126\001\ -\035\000\036\000\037\000\021\001\022\001\040\000\024\001\102\001\ -\103\001\016\002\137\001\088\000\102\003\125\001\137\001\110\001\ -\031\001\032\001\033\001\034\001\166\002\088\001\040\001\021\001\ -\136\001\208\000\209\000\173\002\156\001\157\001\158\001\159\001\ -\099\001\100\001\101\001\102\001\103\001\055\001\115\000\122\002\ -\123\002\088\001\081\003\110\001\114\003\169\002\129\001\069\001\ -\084\001\119\003\085\000\086\000\067\001\122\002\123\002\102\001\ -\103\001\001\001\143\001\088\001\091\002\079\001\001\001\204\002\ -\008\001\142\000\073\001\147\004\017\001\008\001\126\001\217\002\ -\101\001\102\001\103\001\038\001\217\002\024\001\022\001\069\001\ -\163\002\110\001\165\002\022\001\091\002\025\001\125\001\072\002\ -\158\003\217\001\050\001\098\002\001\001\218\002\163\002\228\003\ -\165\002\021\001\022\001\008\001\024\001\030\001\182\004\013\001\ -\232\001\030\001\013\001\125\001\126\001\127\001\145\000\194\002\ -\195\002\022\001\030\001\126\001\040\001\025\001\136\001\137\001\ -\088\001\198\000\199\000\203\004\142\001\194\002\195\002\204\000\ -\205\000\212\002\168\001\208\000\209\000\101\001\102\001\103\001\ -\213\000\214\000\221\002\040\001\223\002\137\001\110\001\220\000\ -\127\001\228\002\127\001\182\000\215\002\167\001\220\003\015\001\ -\235\002\024\001\088\001\040\003\142\001\003\001\021\001\242\002\ -\104\001\105\001\245\002\040\001\136\001\104\001\105\001\250\002\ -\102\001\103\001\237\002\248\000\249\000\242\002\144\001\252\000\ -\110\001\162\002\232\003\136\001\168\001\250\002\217\000\031\001\ -\032\001\033\001\034\001\008\001\004\001\211\003\212\003\213\003\ -\214\003\215\003\127\001\104\001\105\001\095\001\025\003\040\001\ -\142\002\143\002\144\002\145\002\146\002\147\002\148\002\149\002\ -\150\002\151\002\145\001\088\001\142\001\040\003\002\001\126\001\ -\043\003\044\003\045\003\067\001\251\002\181\003\253\002\100\001\ -\101\001\102\001\103\001\095\003\049\001\050\001\043\003\044\003\ -\045\003\110\001\030\001\030\001\057\001\197\003\014\001\040\001\ -\247\002\248\002\017\003\021\001\251\002\125\001\253\002\021\001\ -\071\001\114\003\071\001\072\001\073\001\074\001\040\001\034\001\ -\088\001\084\003\030\001\040\001\035\003\126\001\089\003\014\001\ -\021\001\022\001\017\003\024\001\137\001\048\001\102\001\103\001\ -\051\001\100\003\021\001\022\001\089\003\024\001\021\001\014\001\ -\127\001\137\001\109\003\040\001\035\003\136\001\136\001\088\001\ -\023\001\116\003\127\001\136\001\177\003\040\001\136\001\116\001\ -\136\001\030\001\119\001\100\001\101\001\102\001\103\001\030\001\ -\040\001\040\001\174\003\088\001\136\001\110\001\137\003\136\001\ -\139\003\125\001\141\003\055\001\143\003\098\001\099\001\100\001\ -\101\001\102\001\103\001\191\003\137\003\055\001\139\003\024\001\ -\141\003\110\001\143\003\136\001\014\001\000\000\069\001\156\001\ -\157\001\158\001\159\001\136\001\212\002\023\001\211\003\212\003\ -\213\003\214\003\215\003\155\001\069\001\221\002\177\003\134\001\ -\137\001\180\003\181\003\220\003\228\002\168\001\040\001\031\001\ -\032\001\033\001\034\001\098\001\099\001\100\001\135\001\180\003\ -\148\001\190\001\197\003\198\003\199\003\245\002\195\001\243\003\ -\031\001\032\001\033\001\034\001\017\001\156\003\247\003\017\001\ -\015\001\198\003\199\003\069\001\127\001\079\001\125\001\218\003\ -\171\001\172\001\173\001\067\001\217\001\105\001\136\001\178\001\ -\126\001\136\001\021\001\230\003\153\001\156\003\233\003\021\001\ -\022\001\010\004\024\001\232\001\163\003\042\004\040\001\166\003\ -\098\001\099\001\100\001\101\001\233\003\021\001\022\001\002\001\ -\024\001\024\001\040\001\127\001\110\001\136\001\046\000\137\001\ -\048\000\049\000\050\000\006\004\021\001\136\001\050\004\015\001\ -\040\001\136\001\136\001\125\001\031\001\032\001\033\001\034\001\ -\136\001\127\001\021\004\022\004\068\000\087\004\136\001\088\004\ -\079\001\028\004\137\001\188\003\189\003\136\001\126\001\092\004\ -\021\004\022\004\245\001\096\004\084\003\021\001\022\001\042\004\ -\024\001\134\001\014\001\021\001\168\001\022\001\137\001\254\003\ -\255\003\000\004\001\004\002\004\100\003\079\001\055\001\002\001\ -\040\001\118\004\017\001\023\001\125\001\109\003\126\001\111\000\ -\067\004\125\001\155\001\100\004\024\001\040\001\168\001\254\003\ -\255\003\000\004\001\004\002\004\040\001\021\001\067\004\040\001\ -\141\004\134\001\243\003\168\001\014\001\088\004\089\004\105\001\ -\003\001\137\001\028\001\126\001\089\002\137\001\014\001\168\001\ -\079\001\094\002\013\001\090\001\168\001\040\001\105\004\136\001\ -\107\004\108\004\021\001\110\004\111\004\112\004\113\004\114\004\ -\136\001\136\001\136\001\136\001\105\004\014\001\107\004\108\004\ -\136\001\136\001\136\001\136\001\088\001\040\001\136\001\054\004\ -\125\001\080\004\125\001\082\004\125\001\055\001\098\001\099\001\ -\100\001\101\001\102\001\103\001\055\001\126\001\013\001\037\001\ -\095\004\050\004\110\001\136\001\151\004\050\001\198\000\199\000\ -\021\001\080\004\021\001\082\004\204\000\205\000\021\001\021\001\ -\208\000\209\000\151\004\021\001\079\001\069\001\214\000\136\001\ -\095\004\126\001\218\003\136\001\220\000\221\000\222\000\223\000\ -\168\001\020\001\168\001\030\001\002\001\229\000\230\003\040\001\ -\125\001\136\001\013\001\235\000\236\000\237\000\125\001\110\001\ -\035\001\126\001\168\001\243\000\040\001\008\001\168\001\042\001\ -\248\000\249\000\136\001\046\001\207\004\002\001\137\001\136\001\ -\125\001\021\001\125\001\126\001\127\001\056\001\006\004\212\002\ -\059\001\126\001\207\004\126\001\147\004\136\001\137\001\030\001\ -\221\002\068\001\223\002\142\001\030\001\021\001\040\001\228\002\ -\040\001\125\001\040\001\040\001\028\004\080\001\136\001\040\001\ -\136\001\126\001\126\001\194\004\137\001\037\001\089\001\040\001\ -\245\002\092\001\030\001\126\001\167\001\136\001\097\001\182\004\ -\136\001\136\001\136\001\136\001\050\001\137\001\136\001\136\001\ -\136\001\040\001\136\001\194\004\111\001\112\001\113\001\114\001\ -\115\001\116\001\117\001\118\001\203\004\136\001\002\001\021\001\ -\235\002\137\001\136\001\136\001\136\001\040\001\136\001\023\001\ -\024\001\025\001\136\001\136\001\028\001\030\001\136\001\136\001\ -\003\001\136\001\141\001\136\001\126\001\247\000\127\001\146\001\ -\040\001\126\001\013\001\127\001\030\001\015\001\126\001\154\001\ -\127\001\126\001\021\001\022\001\055\001\024\001\110\004\111\004\ -\112\004\113\004\114\004\166\001\116\001\000\000\025\003\119\001\ -\040\001\127\001\021\001\126\001\046\000\040\001\048\000\049\000\ -\050\000\136\001\021\001\127\001\003\001\040\003\126\001\084\003\ -\055\001\168\001\155\001\088\003\055\001\021\001\013\001\040\001\ -\088\001\127\001\068\000\069\000\126\001\021\001\021\001\100\003\ -\126\001\024\001\098\001\099\001\100\001\101\001\102\001\103\001\ -\109\003\155\001\014\001\136\001\079\001\126\001\110\001\126\001\ -\126\001\040\001\022\001\023\001\040\001\002\001\126\001\126\001\ -\253\002\126\001\126\001\219\002\180\001\125\001\085\004\126\001\ -\055\001\068\001\168\000\170\000\040\001\161\001\190\001\088\001\ -\136\001\091\001\092\001\093\001\094\001\095\001\220\002\158\002\ -\026\003\098\001\099\001\100\001\101\001\102\001\103\001\026\000\ -\079\001\116\003\125\001\126\001\127\001\110\001\214\001\173\002\ -\081\002\069\001\005\002\177\003\028\000\136\001\137\001\246\003\ -\224\001\225\001\226\001\142\001\228\001\180\002\190\001\228\001\ -\040\001\233\001\182\000\105\002\088\001\149\003\020\003\152\000\ -\224\001\235\002\066\000\243\001\189\003\188\003\098\001\099\001\ -\100\001\101\001\102\001\103\001\167\001\146\004\125\001\126\001\ -\127\001\153\003\110\001\119\001\004\002\005\002\161\000\007\002\ -\219\000\136\001\137\001\022\002\197\001\218\003\177\003\142\001\ -\190\002\125\001\181\003\080\001\198\000\199\000\022\002\065\001\ -\088\001\230\003\204\000\205\000\136\001\150\003\208\000\209\000\ -\006\001\108\001\197\003\099\001\100\001\101\001\102\001\103\001\ -\167\001\191\003\220\000\221\000\222\000\223\000\110\001\014\001\ -\171\004\095\003\023\001\229\000\248\001\027\001\255\255\255\255\ -\023\001\006\004\236\000\237\000\009\004\010\004\255\255\255\255\ -\255\255\243\000\255\255\040\001\255\255\255\255\248\000\249\000\ -\255\255\040\001\255\255\255\255\255\255\255\255\255\255\028\004\ -\255\255\255\255\255\255\057\001\058\001\255\255\255\255\255\255\ -\255\255\089\002\255\255\255\255\066\001\042\004\094\002\255\255\ -\255\255\255\255\255\255\021\001\074\001\075\001\069\001\255\255\ -\255\255\255\255\255\255\255\255\082\001\007\002\255\002\000\003\ -\001\003\002\003\063\004\088\001\255\255\091\001\255\255\093\001\ -\094\001\088\001\096\001\255\255\255\255\098\001\099\001\100\001\ -\101\001\102\001\103\001\098\001\099\001\100\001\101\001\102\001\ -\103\001\110\001\255\255\255\255\140\002\141\002\255\255\110\001\ -\255\255\119\001\255\255\036\003\037\003\038\003\039\003\100\004\ -\255\255\255\255\255\255\255\255\130\001\131\001\125\001\255\255\ -\160\002\110\004\111\004\112\004\113\004\114\004\166\002\255\255\ -\255\255\136\001\255\255\171\002\255\255\173\002\174\002\255\255\ -\150\001\255\255\255\255\255\255\180\002\181\002\089\004\255\255\ -\255\255\255\255\020\001\255\255\162\001\163\001\255\255\165\001\ -\255\255\255\255\116\001\255\255\255\255\119\001\255\255\255\255\ -\255\255\035\001\046\000\255\255\048\000\049\000\050\000\255\255\ +let yycheck = "\003\000\ +\079\000\065\000\149\000\251\000\053\000\072\000\149\001\011\000\ +\012\000\240\000\102\001\081\001\227\000\112\000\104\001\014\001\ +\050\000\060\001\052\000\001\000\002\000\254\001\012\000\224\000\ +\226\000\209\000\048\000\233\000\192\001\101\001\058\001\035\000\ +\036\000\037\000\082\001\039\000\040\000\041\000\048\000\043\000\ +\044\000\230\001\064\000\060\001\066\000\027\000\115\002\029\000\ +\067\000\068\000\069\000\179\001\028\001\013\004\064\000\183\001\ +\184\001\025\001\028\001\042\000\017\003\080\000\028\001\067\000\ +\068\000\069\000\021\001\034\001\015\001\073\000\156\002\157\002\ +\158\002\159\002\160\002\079\000\080\000\067\000\068\000\069\000\ +\103\001\029\001\065\000\038\000\015\001\089\000\090\000\161\000\ +\092\000\163\000\080\000\020\002\114\000\115\000\046\000\025\001\ +\028\001\049\000\000\000\041\001\164\000\049\001\050\001\071\003\ +\060\001\112\003\178\001\067\004\028\001\090\001\003\001\021\001\ +\022\001\136\001\024\001\119\000\120\000\050\001\046\000\087\003\ +\227\000\049\000\077\000\029\001\125\001\173\001\081\000\060\001\ +\076\001\128\001\120\000\014\001\028\001\085\000\241\000\014\001\ +\028\001\050\001\050\001\155\001\023\001\142\001\146\000\049\001\ +\050\001\149\000\029\003\060\001\073\001\000\001\050\001\084\001\ +\103\001\120\001\121\001\125\001\017\001\085\000\225\000\223\000\ +\227\000\028\001\055\001\156\001\070\001\122\003\167\001\123\001\ +\202\000\203\000\076\001\084\001\142\001\102\002\241\000\028\001\ +\076\001\090\001\084\001\073\001\127\001\136\001\186\000\024\001\ +\233\000\136\001\069\001\008\001\098\004\156\001\069\001\080\001\ +\168\001\080\002\126\001\199\000\127\001\245\001\247\000\008\001\ +\168\001\014\001\168\001\145\001\134\001\125\001\223\001\164\003\ +\030\001\199\000\011\000\034\001\028\001\161\001\025\001\098\001\ +\099\001\221\000\252\000\253\000\033\003\034\003\035\003\036\003\ +\037\003\039\002\136\001\041\002\161\001\125\001\198\003\003\002\ +\192\004\125\001\035\000\036\000\037\000\120\003\039\000\040\000\ +\041\000\059\001\125\001\044\000\136\001\178\002\125\001\153\001\ +\161\001\161\001\156\001\219\002\000\001\136\001\162\002\161\001\ +\069\001\136\001\125\001\218\001\076\002\161\001\237\001\021\001\ +\012\001\175\004\194\001\161\001\050\001\141\002\227\003\028\001\ +\125\001\241\002\231\003\041\001\027\004\028\004\029\004\030\004\ +\031\004\021\001\040\001\229\001\105\001\014\001\034\001\025\001\ +\089\000\090\000\038\001\034\002\097\001\101\001\023\001\043\001\ +\105\001\054\001\021\001\107\001\050\001\067\001\084\001\154\001\ +\052\001\053\001\054\001\055\001\090\001\125\001\174\001\029\001\ +\125\001\061\001\120\001\121\001\180\001\181\001\001\001\050\001\ +\136\001\125\001\070\001\136\001\129\001\008\001\080\001\075\001\ +\076\001\077\001\078\001\049\001\050\001\034\004\084\001\052\001\ +\070\001\038\004\142\001\022\001\069\001\069\001\120\001\083\001\ +\001\001\123\001\069\001\050\001\149\000\014\001\156\001\008\001\ +\070\001\116\001\001\001\115\002\069\001\054\001\076\001\090\001\ +\190\003\008\001\090\001\167\001\001\001\113\001\084\001\125\001\ +\116\001\098\001\178\001\008\001\095\001\076\001\101\001\022\001\ +\001\001\090\001\099\003\113\001\000\000\082\004\116\001\008\001\ +\171\002\186\000\085\001\086\001\087\001\088\001\089\001\090\001\ +\091\001\092\001\093\001\094\001\125\001\022\001\030\001\030\003\ +\214\002\073\001\074\001\029\001\069\001\126\001\172\001\136\001\ +\050\001\054\004\021\001\083\001\160\001\161\001\162\001\163\001\ +\137\001\171\001\119\004\167\001\221\000\104\001\105\001\049\001\ +\050\001\248\001\002\001\175\001\176\001\177\001\244\002\009\003\ +\012\003\167\001\182\001\153\001\247\001\024\002\165\001\232\001\ +\085\004\040\001\084\001\161\001\252\001\001\001\254\001\104\001\ +\105\001\050\001\076\001\178\001\008\001\008\002\009\002\144\002\ +\145\002\104\001\105\001\060\001\125\001\157\001\025\001\134\001\ +\135\001\028\001\022\001\104\001\105\001\235\001\069\001\136\001\ +\127\001\221\001\217\002\164\002\126\001\148\001\244\001\104\001\ +\105\001\184\002\050\001\084\001\003\001\157\001\175\002\137\001\ +\236\001\243\001\226\002\054\001\228\002\038\001\167\003\231\002\ +\185\002\170\003\074\001\075\001\248\001\249\001\103\001\023\001\ +\014\001\253\001\029\001\052\001\076\001\137\001\055\001\126\001\ +\022\002\003\001\137\001\138\001\084\001\003\000\028\001\253\001\ +\040\001\040\001\137\001\154\003\076\002\028\001\049\001\050\001\ +\222\002\033\003\034\003\035\003\036\003\037\003\253\002\136\001\ +\055\001\007\002\146\002\147\002\148\002\149\002\150\002\151\002\ +\152\002\153\002\154\002\155\002\104\001\105\001\040\001\054\001\ +\030\001\076\001\134\001\135\001\070\001\043\000\028\001\021\001\ +\161\001\084\001\040\001\050\001\050\000\055\001\052\000\021\001\ +\088\001\127\003\084\001\093\002\150\003\153\003\127\001\003\001\ +\098\002\129\001\098\001\099\001\100\001\101\001\102\001\103\001\ +\054\001\013\001\040\001\073\000\168\001\076\001\110\001\028\001\ +\069\001\021\001\022\001\022\001\024\001\024\001\134\001\135\001\ +\031\001\032\001\033\001\034\001\020\002\076\002\092\000\212\000\ +\213\000\088\001\040\001\084\001\040\001\208\002\088\001\170\002\ +\022\001\054\001\024\001\106\003\073\001\173\002\177\002\102\001\ +\103\001\088\001\221\002\055\001\102\001\103\001\175\001\176\001\ +\177\001\119\000\126\002\127\002\067\001\182\001\118\003\102\001\ +\103\001\058\004\088\001\123\003\143\001\017\001\126\001\110\001\ +\126\002\127\002\024\001\079\001\085\003\076\001\038\001\095\002\ +\102\001\103\001\088\001\069\001\146\000\030\001\021\001\022\001\ +\110\001\024\001\221\002\090\001\098\001\099\001\100\001\101\001\ +\102\001\103\001\232\003\167\002\030\001\169\002\025\001\095\002\ +\110\001\040\001\162\003\125\001\030\001\222\002\102\002\088\001\ +\050\001\167\002\013\001\169\002\013\001\021\001\022\001\166\002\ +\024\001\125\001\126\001\127\001\025\001\102\001\103\001\168\001\ +\249\001\044\003\198\002\199\002\136\001\137\001\014\001\126\001\ +\040\001\040\001\142\001\088\001\202\000\203\000\137\001\023\001\ +\198\002\199\002\208\000\209\000\216\002\127\001\212\000\213\000\ +\101\001\102\001\103\001\217\000\218\000\225\002\151\004\227\002\ +\040\001\110\001\224\000\167\001\232\002\127\001\021\001\022\001\ +\224\003\024\001\015\001\239\002\024\001\219\002\031\001\032\001\ +\033\001\034\001\246\002\142\001\003\001\249\002\021\001\136\001\ +\136\001\040\001\254\002\144\001\127\001\069\001\252\000\253\000\ +\246\002\186\004\000\001\241\002\168\001\236\003\014\001\004\001\ +\254\002\215\003\216\003\217\003\218\003\219\003\012\001\023\001\ +\095\001\040\001\067\001\145\001\021\001\022\001\207\004\024\001\ +\142\001\029\003\098\001\099\001\100\001\101\001\002\001\014\001\ +\040\001\021\001\022\001\099\003\024\001\030\001\185\003\040\001\ +\044\003\126\001\021\001\047\003\048\003\049\003\030\001\255\002\ +\125\001\001\003\021\001\071\001\040\001\125\001\201\003\053\001\ +\054\001\047\003\048\003\049\003\040\001\069\001\030\001\061\001\ +\136\001\126\001\137\001\251\002\252\002\021\003\014\001\255\002\ +\021\001\001\003\181\003\127\001\136\001\075\001\076\001\077\001\ +\078\001\137\001\021\001\022\001\088\003\024\001\040\001\039\003\ +\136\001\093\003\098\001\099\001\100\001\021\003\136\001\136\001\ +\118\003\136\001\127\001\030\001\104\003\040\001\030\001\093\003\ +\014\001\001\000\002\000\003\000\004\000\113\003\136\001\039\003\ +\022\001\023\001\178\003\136\001\120\003\125\001\031\001\032\001\ +\033\001\034\001\120\001\125\001\040\001\123\001\055\001\136\001\ +\136\001\055\001\040\001\195\003\024\001\155\001\088\001\069\001\ +\136\001\141\003\134\001\143\003\168\001\145\003\137\001\147\003\ +\098\001\099\001\100\001\101\001\102\001\103\001\135\001\141\003\ +\017\001\143\003\067\001\145\003\110\001\147\003\148\001\069\001\ +\000\000\017\001\160\001\161\001\162\001\163\001\215\003\216\003\ +\217\003\218\003\219\003\031\001\032\001\033\001\034\001\127\001\ +\015\001\181\003\088\001\079\001\184\003\185\003\239\002\247\003\ +\031\001\032\001\033\001\034\001\098\001\099\001\100\001\101\001\ +\102\001\103\001\184\003\046\004\194\001\201\003\202\003\203\003\ +\110\001\199\001\224\003\136\001\105\001\040\001\126\001\067\001\ +\160\003\192\003\193\003\153\001\202\003\203\003\021\001\125\001\ +\040\001\002\001\222\003\031\001\032\001\033\001\034\001\221\001\ +\024\001\040\001\136\001\136\001\029\003\251\003\234\003\127\001\ +\160\003\237\003\137\001\015\001\014\004\023\001\236\001\167\003\ +\021\001\136\001\170\003\044\003\127\001\096\004\054\004\237\003\ +\136\001\100\004\136\001\136\001\079\001\088\001\040\001\136\001\ +\134\001\050\000\137\001\052\000\053\000\054\000\010\004\014\001\ +\247\003\100\001\101\001\102\001\103\001\092\004\014\001\122\004\ +\023\001\088\001\126\001\110\001\021\001\025\004\026\004\072\000\ +\091\004\022\001\137\001\079\001\032\004\100\001\101\001\102\001\ +\103\001\040\001\000\000\025\004\026\004\168\001\145\004\110\001\ +\055\001\017\001\046\004\002\001\125\001\155\001\088\001\125\001\ +\040\001\126\001\002\004\003\004\004\004\005\004\006\004\024\001\ +\098\001\099\001\100\001\101\001\102\001\103\001\069\001\120\003\ +\168\001\021\001\115\000\071\004\110\001\040\001\104\004\054\004\ +\134\001\014\001\002\004\003\004\004\004\005\004\006\004\040\001\ +\168\001\071\004\105\001\137\001\088\001\028\001\126\001\137\001\ +\092\004\093\004\014\001\098\001\099\001\100\001\101\001\093\002\ +\079\001\101\001\102\001\103\001\098\002\168\001\168\001\110\001\ +\090\001\109\004\110\001\111\004\112\004\136\001\114\004\115\004\ +\116\004\117\004\118\004\136\001\136\001\136\001\125\001\109\004\ +\136\001\111\004\112\004\136\001\181\003\136\001\040\001\088\001\ +\185\003\136\001\058\004\136\001\084\004\136\001\086\004\040\001\ +\014\001\136\001\099\001\100\001\101\001\102\001\103\001\125\001\ +\201\003\125\001\125\001\099\004\055\001\110\001\013\001\155\004\ +\126\001\202\000\203\000\037\001\084\004\050\001\086\004\208\000\ +\209\000\136\001\021\001\212\000\213\000\155\004\021\001\021\001\ +\021\001\218\000\021\001\099\004\126\001\168\001\088\001\224\000\ +\225\000\226\000\227\000\069\001\020\001\168\001\030\001\136\001\ +\233\000\099\001\100\001\101\001\102\001\103\001\239\000\240\000\ +\241\000\002\001\136\001\035\001\110\001\040\001\247\000\014\001\ +\125\001\136\001\042\001\252\000\253\000\125\001\046\001\211\004\ +\023\001\013\001\126\001\168\001\110\001\040\001\008\001\136\001\ +\056\001\168\001\216\002\059\001\002\001\211\004\125\001\151\004\ +\136\001\040\001\126\001\225\002\068\001\227\002\021\001\137\001\ +\025\001\030\001\232\002\126\001\125\001\040\001\030\001\040\001\ +\080\001\040\001\040\001\137\001\040\001\126\001\198\004\126\001\ +\041\001\089\001\136\001\249\002\092\001\136\001\069\001\040\001\ +\030\001\097\001\186\004\126\001\003\003\004\003\005\003\006\003\ +\136\001\136\001\136\001\136\001\136\001\040\001\198\004\111\001\ +\112\001\113\001\114\001\115\001\116\001\117\001\118\001\207\004\ +\137\001\050\001\002\001\098\001\099\001\100\001\101\001\136\001\ +\021\001\136\001\136\001\136\001\093\004\137\001\136\001\110\001\ +\006\001\040\003\041\003\042\003\043\003\141\001\136\001\136\001\ +\251\000\136\001\146\001\136\001\136\001\040\001\125\001\136\001\ +\136\001\136\001\154\001\136\001\136\001\027\001\030\001\126\001\ +\126\001\136\001\127\001\030\001\136\001\127\001\166\001\120\001\ +\015\001\055\001\123\001\126\001\000\000\021\001\040\001\050\000\ +\126\001\052\000\053\000\054\000\127\001\127\001\003\001\021\001\ +\127\001\126\001\088\003\057\001\058\001\126\001\092\003\055\001\ +\013\001\021\001\168\001\155\001\066\001\072\000\073\000\127\001\ +\021\001\022\001\104\003\024\001\074\001\075\001\126\001\021\001\ +\126\001\126\001\136\001\113\003\082\001\014\001\126\001\155\001\ +\126\001\040\001\002\001\040\001\003\001\091\001\023\001\093\001\ +\094\001\126\001\096\001\126\001\126\001\126\001\013\001\184\001\ +\126\001\001\003\055\001\223\002\072\001\089\004\021\001\040\001\ +\172\000\194\001\165\001\174\000\095\001\096\001\097\001\098\001\ +\099\001\119\001\030\003\177\002\085\002\009\002\250\003\224\002\ +\162\002\040\001\079\001\181\003\130\001\131\001\027\000\002\000\ +\194\001\218\001\029\000\186\000\069\001\024\003\184\002\153\003\ +\055\001\109\002\156\000\228\001\229\001\230\001\239\002\232\001\ +\150\001\070\000\193\003\232\001\237\001\192\003\157\003\088\001\ +\123\001\228\001\150\004\165\000\162\001\163\001\247\001\165\001\ +\079\001\098\001\099\001\100\001\101\001\102\001\103\001\223\000\ +\125\001\126\001\127\001\084\001\069\001\110\001\154\003\008\002\ +\009\002\201\001\011\002\136\001\137\001\194\002\112\001\252\001\ +\222\003\142\001\195\003\099\003\125\001\175\004\255\255\202\000\ +\203\000\026\002\255\255\255\255\234\003\208\000\209\000\136\001\ +\255\255\212\000\213\000\255\255\255\255\255\255\125\001\126\001\ +\127\001\255\255\167\001\026\002\255\255\224\000\225\000\226\000\ +\227\000\136\001\137\001\255\255\255\255\255\255\233\000\142\001\ +\255\255\255\255\255\255\255\255\010\004\240\000\241\000\013\004\ +\014\004\255\255\255\255\014\001\247\000\255\255\255\255\255\255\ +\255\255\252\000\253\000\255\255\023\001\255\255\255\255\255\255\ +\167\001\255\255\032\004\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\255\255\255\255\093\002\040\001\255\255\003\001\ +\046\004\098\002\255\255\255\255\255\255\255\255\025\001\255\255\ +\255\255\013\001\255\255\255\255\255\255\255\255\255\255\255\255\ +\011\002\021\001\255\255\255\255\024\001\067\004\255\255\255\255\ +\255\255\255\255\069\001\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\255\255\255\255\040\001\255\255\255\255\255\255\ +\255\255\255\255\255\255\255\255\255\255\255\255\255\255\144\002\ +\145\002\255\255\255\255\055\001\255\255\255\255\255\255\098\001\ +\099\001\100\001\104\004\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\255\255\164\002\114\004\115\004\116\004\117\004\ +\118\004\170\002\255\255\079\001\255\255\255\255\175\002\255\255\ +\177\002\178\002\125\001\255\255\255\255\255\255\255\255\184\002\ +\185\002\255\255\255\255\255\255\255\255\136\001\255\255\255\255\ +\255\255\255\255\255\255\255\255\255\255\120\001\255\255\255\255\ +\123\001\001\001\255\255\255\255\255\255\050\000\255\255\052\000\ +\053\000\054\000\255\255\255\255\255\255\019\001\014\001\255\255\ +\255\255\125\001\126\001\127\001\221\002\222\002\022\001\023\001\ +\255\255\255\255\255\255\072\000\136\001\137\001\231\002\255\255\ +\255\255\255\255\142\001\255\255\255\255\255\255\044\001\255\255\ +\040\001\255\255\192\004\255\255\050\001\255\255\255\255\255\255\ +\255\255\255\255\255\255\255\255\253\002\255\255\060\001\061\001\ +\062\001\063\001\064\001\167\001\255\255\255\255\255\255\255\255\ +\255\255\255\255\255\255\255\255\255\255\069\001\255\255\194\001\ +\017\003\255\255\255\255\255\255\255\255\255\255\084\001\085\001\ +\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ +\088\001\255\255\255\255\255\255\255\255\255\255\255\255\218\001\ +\255\255\255\255\098\001\099\001\100\001\101\001\102\001\103\001\ +\104\001\228\001\229\001\230\001\255\255\232\001\110\001\255\255\ +\255\255\255\255\237\001\255\255\255\255\255\255\124\001\255\255\ +\255\255\255\255\128\001\255\255\247\001\125\001\132\001\226\002\ +\255\255\228\002\255\255\255\255\231\002\139\001\140\001\255\255\ +\136\001\143\001\255\255\255\255\085\003\008\002\009\002\255\255\ +\011\002\151\001\255\255\092\003\255\255\255\255\255\255\157\001\ +\158\001\159\001\160\001\255\255\255\255\202\000\203\000\026\002\ +\255\255\255\255\255\255\255\255\007\003\255\255\255\255\212\000\ +\213\000\255\255\006\001\255\255\255\255\118\003\017\003\255\255\ +\255\255\122\003\255\255\255\255\225\000\226\000\227\000\255\255\ +\255\255\255\255\255\255\255\255\233\000\255\255\255\255\027\001\ +\255\255\255\255\255\255\240\000\241\000\255\255\255\255\255\255\ +\255\255\255\255\247\000\255\255\255\255\255\255\255\255\252\000\ +\253\000\255\255\255\255\255\255\255\255\255\255\255\255\160\003\ +\255\255\255\255\255\255\164\003\255\255\057\001\058\001\255\255\ +\255\255\255\255\093\002\255\255\255\255\255\255\066\001\098\002\ +\255\255\255\255\255\255\255\255\025\001\255\255\074\001\075\001\ +\255\255\255\255\255\255\255\255\255\255\255\255\082\001\255\255\ +\255\255\255\255\255\255\255\255\255\255\255\255\255\255\091\001\ +\255\255\093\001\094\001\255\255\096\001\255\255\255\255\255\255\ +\255\255\255\255\255\255\255\255\255\255\255\255\215\003\216\003\ +\217\003\218\003\219\003\118\003\255\255\144\002\145\002\122\003\ +\123\003\255\255\227\003\119\001\255\255\255\255\231\003\255\255\ +\255\255\255\255\255\255\236\003\255\255\255\255\130\001\131\001\ +\255\255\164\002\255\255\255\255\255\255\246\003\247\003\170\002\ +\255\255\255\255\255\255\255\255\175\002\255\255\177\002\178\002\ +\255\255\255\255\150\001\255\255\255\255\184\002\185\002\162\003\ +\255\255\164\003\255\255\255\255\013\004\014\004\162\001\163\001\ +\255\255\165\001\255\255\120\001\255\255\255\255\123\001\255\255\ +\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ +\255\255\034\004\255\255\003\000\255\255\038\004\255\255\255\255\ +\255\255\255\255\221\002\222\002\255\255\255\255\255\255\255\255\ +\227\002\255\255\255\255\052\004\255\255\054\004\255\255\255\255\ +\014\001\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ +\255\255\023\001\067\004\255\255\255\255\224\003\255\255\255\255\ +\227\003\255\255\253\002\043\000\231\003\255\255\255\255\255\255\ +\255\255\082\004\040\001\255\255\085\004\255\255\053\000\054\000\ +\255\255\255\255\091\004\255\255\255\255\255\255\017\003\255\255\ +\255\255\098\004\255\255\255\255\255\255\255\255\255\255\104\004\ +\255\255\072\000\255\255\255\255\255\255\255\255\255\255\069\001\ +\255\255\255\255\255\255\255\255\255\255\255\255\119\004\255\255\ +\255\255\255\255\255\255\255\255\092\000\255\255\255\255\228\001\ +\229\001\230\001\088\001\232\001\255\255\255\255\255\255\034\004\ +\237\001\255\255\255\255\038\004\098\001\099\001\100\001\101\001\ +\102\001\103\001\247\001\255\255\023\001\024\001\025\001\119\000\ +\110\001\028\001\255\255\255\255\255\255\255\255\255\255\255\255\ +\255\255\131\000\085\003\008\002\009\002\040\001\255\255\125\001\ +\255\255\092\003\255\255\255\255\255\255\255\255\175\004\255\255\ +\255\255\255\255\136\001\255\255\255\255\026\002\255\255\082\004\ +\255\255\255\255\255\255\255\255\255\255\156\000\255\255\192\004\ +\255\255\255\255\255\255\255\255\255\255\255\255\255\255\122\003\ +\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\177\000\178\000\255\255\088\001\255\255\255\255\ +\255\255\255\255\255\255\255\255\119\004\255\255\255\255\098\001\ +\099\001\100\001\101\001\102\001\103\001\255\255\255\255\255\255\ +\255\255\255\255\255\255\110\001\255\255\255\255\208\000\255\255\ +\255\255\164\003\255\255\255\255\089\002\255\255\255\255\255\255\ +\093\002\255\255\125\001\255\255\255\255\098\002\224\000\255\255\ +\255\255\255\255\225\000\226\000\227\000\136\001\255\255\255\255\ +\255\255\255\255\233\000\255\255\255\255\006\001\255\255\238\000\ +\239\000\240\000\241\000\255\255\255\255\255\255\255\255\255\255\ +\247\000\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\027\001\255\255\215\003\216\003\217\003\218\003\ +\219\003\255\255\255\255\255\255\012\001\255\255\255\255\255\255\ +\227\003\255\255\255\255\255\255\231\003\255\255\255\255\255\255\ +\255\255\236\003\025\001\255\255\255\255\255\255\255\255\255\255\ +\057\001\058\001\255\255\255\255\247\003\170\002\255\255\255\255\ +\255\255\066\001\255\255\255\255\177\002\178\002\255\255\255\255\ +\255\255\074\001\075\001\255\255\255\255\053\001\054\001\255\255\ +\255\255\082\001\013\004\014\004\255\255\061\001\255\255\255\255\ +\063\001\255\255\091\001\069\001\093\001\094\001\255\255\096\001\ +\255\255\255\255\255\255\075\001\076\001\077\001\078\001\034\004\ +\255\255\255\255\255\255\038\004\255\255\255\255\255\255\255\255\ +\221\002\222\002\255\255\255\255\255\255\255\255\119\001\255\255\ +\003\001\255\255\100\001\054\004\255\255\255\255\104\001\102\001\ +\255\255\130\001\131\001\255\255\255\255\255\255\255\255\255\255\ +\067\004\255\255\255\255\022\001\023\001\255\255\025\001\255\255\ +\253\002\028\001\255\255\255\255\255\255\150\001\255\255\082\004\ +\255\255\255\255\085\004\255\255\255\255\040\001\255\255\255\255\ +\091\004\162\001\163\001\255\255\165\001\255\255\255\255\098\004\ +\255\255\052\001\255\255\054\001\255\255\104\004\149\001\255\255\ +\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ +\160\001\161\001\162\001\163\001\119\004\255\255\255\255\255\255\ +\255\255\255\255\255\255\255\255\174\001\255\255\081\001\255\255\ +\255\255\179\001\180\001\181\001\255\255\183\001\184\001\255\255\ +\255\255\255\255\255\255\255\255\255\255\255\255\192\001\098\001\ +\099\001\100\001\101\001\102\001\103\001\104\001\105\001\106\001\ +\107\001\108\001\255\255\110\001\255\255\255\255\255\255\255\255\ +\255\255\255\255\255\255\255\255\255\255\255\255\255\255\092\003\ +\255\255\255\255\255\255\218\001\175\004\221\001\255\255\255\255\ +\255\255\255\255\255\255\226\001\255\255\228\001\229\001\230\001\ +\255\255\232\001\020\001\255\255\236\001\192\004\237\001\255\255\ +\255\255\255\255\149\001\255\255\255\255\255\255\255\255\255\255\ +\247\001\035\001\255\255\255\255\255\255\255\255\255\255\255\255\ \042\001\255\255\255\255\255\255\046\001\255\255\255\255\255\255\ -\255\255\217\002\218\002\255\255\255\255\255\255\056\001\255\255\ -\068\000\059\001\255\255\227\002\255\255\255\255\255\255\255\255\ -\255\255\255\255\068\001\255\255\255\255\255\255\255\255\188\004\ +\255\255\008\002\009\002\255\255\255\255\255\255\056\001\255\255\ +\255\255\059\001\255\255\255\255\255\255\255\255\255\255\255\255\ +\157\003\255\255\068\001\026\002\255\255\255\255\255\255\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\080\001\255\255\ -\255\255\249\002\255\255\255\255\255\255\255\255\255\255\089\001\ -\255\255\255\255\092\001\255\255\255\255\014\001\255\255\097\001\ -\255\255\255\255\255\255\255\255\190\001\013\003\023\001\255\255\ +\255\255\255\255\255\255\255\255\255\255\044\002\255\255\089\001\ +\255\255\255\255\092\001\255\255\255\255\255\255\255\255\097\001\ +\255\255\255\255\255\255\006\000\255\255\255\255\255\255\255\255\ \255\255\255\255\255\255\255\255\255\255\111\001\112\001\113\001\ -\114\001\115\001\116\001\117\001\118\001\255\255\255\255\040\001\ -\255\255\255\255\255\255\255\255\214\001\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\255\255\224\001\225\001\ -\226\001\255\255\228\001\141\001\255\255\255\255\255\255\233\001\ -\146\001\255\255\255\255\255\255\069\001\255\255\255\255\255\255\ -\154\001\243\001\255\255\255\255\222\002\255\255\224\002\255\255\ -\255\255\227\002\255\255\255\255\166\001\255\255\255\255\088\001\ -\255\255\081\003\004\002\005\002\255\255\007\002\255\255\255\255\ -\088\003\098\001\099\001\100\001\101\001\102\001\103\001\255\255\ -\255\255\255\255\198\000\199\000\022\002\110\001\255\255\255\255\ -\255\255\003\003\255\255\255\255\208\000\209\000\255\255\006\001\ -\255\255\255\255\114\003\013\003\125\001\255\255\118\003\255\255\ -\255\255\221\000\222\000\223\000\255\255\255\255\255\255\136\001\ -\255\255\229\000\255\255\255\255\027\001\255\255\255\255\255\255\ -\236\000\237\000\255\255\255\255\255\255\255\255\255\255\243\000\ -\255\255\255\255\255\255\255\255\248\000\249\000\255\255\255\255\ -\255\255\255\255\255\255\255\255\156\003\255\255\255\255\255\255\ -\160\003\255\255\057\001\058\001\255\255\255\255\255\255\089\002\ -\255\255\255\255\255\255\066\001\094\002\255\255\255\255\255\255\ -\255\255\021\001\255\255\074\001\075\001\255\255\255\255\255\255\ -\255\255\255\255\255\255\082\001\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\091\001\255\255\093\001\094\001\ -\255\255\096\001\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\211\003\212\003\213\003\214\003\215\003\ -\114\003\255\255\140\002\141\002\118\003\119\003\255\255\223\003\ -\119\001\255\255\255\255\227\003\255\255\255\255\255\255\255\255\ -\232\003\255\255\255\255\130\001\131\001\255\255\160\002\255\255\ -\255\255\255\255\242\003\243\003\166\002\255\255\255\255\255\255\ -\255\255\171\002\255\255\173\002\174\002\255\255\255\255\150\001\ -\255\255\255\255\180\002\181\002\158\003\255\255\160\003\255\255\ -\255\255\009\004\010\004\162\001\163\001\255\255\165\001\255\255\ -\116\001\255\255\255\255\119\001\001\001\255\255\255\255\255\255\ -\255\255\255\255\255\255\008\001\255\255\255\255\030\004\255\255\ -\255\255\014\001\034\004\255\255\255\255\255\255\255\255\217\002\ -\218\002\022\001\023\001\255\255\255\255\223\002\255\255\255\255\ -\048\004\255\255\050\004\255\255\255\255\014\001\255\255\255\255\ -\255\255\255\255\255\255\040\001\255\255\255\255\023\001\063\004\ -\255\255\255\255\220\003\255\255\255\255\223\003\255\255\249\002\ -\255\255\227\003\255\255\255\255\255\255\255\255\078\004\040\001\ -\255\255\081\004\255\255\049\000\050\000\255\255\255\255\087\004\ -\069\001\255\255\255\255\013\003\255\255\255\255\094\004\255\255\ -\255\255\255\255\255\255\255\255\100\004\255\255\068\000\255\255\ -\255\255\255\255\255\255\088\001\069\001\255\255\255\255\255\255\ -\255\255\255\255\255\255\115\004\255\255\098\001\099\001\100\001\ -\101\001\102\001\103\001\104\001\224\001\225\001\226\001\255\255\ -\228\001\110\001\255\255\255\255\030\004\233\001\255\255\255\255\ -\034\004\098\001\099\001\100\001\101\001\255\255\255\255\243\001\ -\125\001\255\255\255\255\255\255\255\255\110\001\255\255\255\255\ -\255\255\255\255\255\255\136\001\255\255\255\255\255\255\081\003\ -\004\002\005\002\255\255\255\255\125\001\255\255\088\003\255\255\ -\255\255\255\255\255\255\171\004\255\255\255\255\255\255\136\001\ -\255\255\255\255\022\002\255\255\078\004\255\255\255\255\255\255\ -\255\255\255\255\152\000\255\255\188\004\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\118\003\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\255\255\255\255\173\000\ -\174\000\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\115\004\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\255\255\160\003\255\255\ -\255\255\085\002\255\255\255\255\255\255\089\002\255\255\255\255\ -\255\255\255\255\094\002\255\255\255\255\255\255\255\255\221\000\ -\222\000\223\000\255\255\255\255\255\255\255\255\255\255\229\000\ -\255\255\255\255\255\255\255\255\234\000\235\000\236\000\237\000\ -\255\255\255\255\255\255\255\255\255\255\243\000\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\211\003\212\003\213\003\214\003\215\003\255\255\255\255\ -\255\255\255\255\020\001\255\255\255\255\223\003\255\255\005\000\ -\255\255\227\003\255\255\255\255\255\255\255\255\232\003\021\001\ -\255\255\035\001\255\255\255\255\255\255\255\255\255\255\255\255\ -\042\001\243\003\166\002\255\255\046\001\255\255\255\255\255\255\ -\255\255\173\002\174\002\255\255\127\000\255\255\056\001\255\255\ -\255\255\059\001\255\255\001\001\255\255\255\255\255\255\009\004\ -\010\004\255\255\068\001\049\000\050\000\059\001\255\255\255\255\ -\014\001\255\255\255\255\255\255\255\255\255\255\080\001\255\255\ -\022\001\023\001\255\255\255\255\030\004\255\255\068\000\089\001\ -\034\004\255\255\092\001\255\255\255\255\217\002\218\002\097\001\ -\255\255\255\255\040\001\255\255\255\255\255\255\255\255\255\255\ -\050\004\255\255\255\255\255\255\098\001\111\001\112\001\113\001\ -\114\001\115\001\116\001\117\001\118\001\063\004\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\249\002\255\255\069\001\ -\255\255\204\000\255\255\255\255\078\004\255\255\255\255\081\004\ -\255\255\255\255\255\255\141\001\255\255\087\004\255\255\255\255\ -\146\001\220\000\088\001\255\255\094\004\255\255\255\255\255\255\ -\154\001\255\255\100\004\145\001\098\001\099\001\100\001\101\001\ -\102\001\103\001\104\001\255\255\166\001\255\255\255\255\255\255\ -\110\001\115\004\152\000\153\000\154\000\155\000\156\000\255\255\ -\255\255\255\255\160\000\161\000\162\000\255\255\164\000\125\001\ -\255\255\255\255\255\255\255\255\255\255\255\255\172\000\173\000\ -\255\255\255\255\136\001\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\049\000\050\000\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\088\003\255\255\255\255\255\255\ -\214\001\171\004\255\255\068\000\255\255\255\255\255\255\255\255\ -\222\001\255\255\224\001\225\001\226\001\219\000\228\001\221\000\ -\222\000\223\000\188\004\233\001\255\255\255\255\065\001\229\000\ -\255\255\255\255\255\255\255\255\255\255\243\001\236\000\237\000\ -\255\255\255\255\255\255\255\255\255\255\243\000\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\255\255\004\002\005\002\ -\255\255\255\255\255\255\049\000\050\000\096\001\255\255\255\255\ -\255\255\100\001\255\255\255\255\255\255\153\003\255\255\255\255\ -\022\002\255\255\255\255\255\255\255\255\255\255\068\000\021\001\ -\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\040\002\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\155\000\156\000\255\255\255\255\255\255\160\000\ -\255\255\162\000\255\255\164\000\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\172\000\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\255\255\076\002\211\003\ -\212\003\213\003\214\003\215\003\255\255\255\255\255\255\170\001\ -\255\255\255\255\255\255\255\255\175\001\176\001\177\001\255\255\ -\179\001\180\001\255\255\255\255\232\003\255\255\255\255\255\255\ -\255\255\188\001\255\255\105\002\255\255\255\255\255\255\255\255\ -\255\255\103\001\255\255\255\255\221\000\222\000\223\000\255\255\ -\255\255\255\255\255\255\255\255\229\000\255\255\255\255\255\255\ -\118\001\255\255\255\255\236\000\237\000\009\004\010\004\125\001\ -\174\000\255\255\243\000\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\136\001\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\145\001\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\166\002\255\255\255\255\255\255\021\001\255\255\255\255\173\002\ -\174\002\255\255\255\255\255\255\255\255\255\255\180\002\221\000\ -\222\000\223\000\255\255\063\004\255\255\255\255\039\001\229\000\ -\255\255\255\255\255\255\255\255\255\255\255\255\236\000\237\000\ -\255\255\255\255\255\255\255\255\255\255\243\000\255\255\255\255\ -\255\255\255\255\255\255\087\004\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\217\002\218\002\255\255\255\255\255\255\ -\100\004\255\255\255\255\255\255\255\255\219\001\255\255\255\255\ -\255\255\255\255\224\001\225\001\226\001\255\255\228\001\021\001\ -\255\255\255\255\255\255\233\001\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\249\002\255\255\243\001\255\255\255\255\ -\255\255\255\255\248\001\255\255\250\001\255\255\255\255\255\255\ -\255\255\255\255\008\003\255\255\255\255\255\255\004\002\005\002\ -\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\026\003\255\255\255\255\136\001\ -\022\002\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\255\255\137\002\255\255\ -\188\004\140\002\141\002\142\002\143\002\144\002\145\002\146\002\ -\147\002\148\002\149\002\150\002\151\002\152\002\153\002\154\002\ -\155\002\156\002\157\002\158\002\255\255\160\002\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\255\255\076\002\255\255\ -\171\002\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\181\002\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\097\002\255\255\255\255\100\002\255\255\ -\255\255\255\255\104\002\105\002\255\255\255\255\255\255\224\001\ -\225\001\226\001\255\255\228\001\255\255\255\255\255\255\255\255\ -\233\001\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\243\001\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\255\255\255\255\149\003\ -\150\003\255\255\255\255\004\002\005\002\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\022\002\255\255\255\255\ -\166\002\167\002\005\003\169\002\255\255\255\255\255\255\173\002\ -\174\002\255\255\224\001\225\001\226\001\255\255\228\001\255\255\ -\255\255\255\255\255\255\233\001\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\243\001\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\211\003\212\003\213\003\ -\214\003\215\003\255\255\255\255\255\255\255\255\004\002\005\002\ -\255\255\255\255\255\255\217\002\218\002\255\255\255\255\255\255\ -\255\255\255\255\232\003\255\255\255\255\255\255\255\255\255\255\ -\022\002\255\255\255\255\255\255\255\255\243\003\255\255\255\255\ -\097\002\255\255\255\255\100\002\255\255\255\255\081\003\104\002\ -\255\255\019\001\000\000\249\002\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\044\001\255\255\255\255\255\255\255\255\255\255\ -\050\001\023\003\255\255\255\255\026\003\255\255\255\255\255\255\ -\255\255\255\255\060\001\061\001\062\001\063\001\064\001\045\004\ -\255\255\255\255\048\004\255\255\050\004\255\255\052\004\255\255\ -\255\255\255\255\255\255\255\255\255\255\166\002\255\255\146\003\ -\255\255\255\255\084\001\085\001\173\002\174\002\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\081\004\255\255\255\255\255\255\255\255\ -\255\255\087\004\255\255\255\255\255\255\255\255\255\255\001\001\ -\094\004\255\255\255\255\255\255\098\004\255\255\008\001\186\003\ -\255\255\095\003\124\001\255\255\014\001\255\255\128\001\255\255\ -\217\002\218\002\132\001\255\255\022\001\023\001\255\255\255\255\ -\255\255\139\001\140\001\255\255\255\255\143\001\255\255\255\255\ -\166\002\255\255\255\255\255\255\255\255\151\001\040\001\173\002\ -\174\002\255\255\255\255\157\001\158\001\159\001\160\001\255\255\ -\249\002\255\255\255\255\255\255\146\004\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\150\003\255\255\255\255\069\001\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\171\004\023\003\255\255\ -\255\255\255\255\255\255\217\002\218\002\255\255\088\001\255\255\ -\174\003\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\098\001\099\001\100\001\101\001\102\001\103\001\104\001\255\255\ -\255\255\191\003\255\255\255\255\110\001\255\255\255\255\255\255\ -\255\255\255\255\255\255\249\002\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\125\001\255\255\211\003\212\003\213\003\ -\214\003\215\003\255\255\255\255\255\255\255\255\136\001\255\255\ -\255\255\000\000\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\232\003\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\242\003\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\002\001\003\001\004\001\255\255\ -\006\001\255\255\255\255\255\255\255\255\255\255\255\255\013\001\ -\255\255\015\001\255\255\017\001\255\255\255\255\020\001\255\255\ -\255\255\255\255\255\255\255\255\255\255\027\001\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\035\001\255\255\037\001\ -\038\001\255\255\255\255\041\001\042\001\255\255\255\255\255\255\ -\046\001\255\255\255\255\041\004\255\255\255\255\052\001\255\255\ -\054\001\055\001\056\001\057\001\058\001\059\001\255\255\053\004\ -\255\255\255\255\255\255\255\255\066\001\255\255\068\001\255\255\ -\255\255\255\255\255\255\255\255\074\001\075\001\255\255\255\255\ -\255\255\079\001\080\001\255\255\082\001\255\255\255\255\255\255\ -\255\255\255\255\255\255\089\001\255\255\091\001\092\001\093\001\ -\094\001\087\004\096\001\097\001\255\255\255\255\255\255\255\255\ -\255\255\255\255\211\003\212\003\213\003\214\003\215\003\255\255\ -\255\255\111\001\112\001\113\001\114\001\115\001\116\001\117\001\ -\118\001\119\001\255\255\255\255\255\255\255\255\255\255\232\003\ -\126\001\127\001\255\255\255\255\130\001\131\001\255\255\255\255\ -\134\001\135\001\136\001\137\001\138\001\255\255\255\255\141\001\ -\255\255\000\000\144\001\145\001\146\001\255\255\148\001\255\255\ -\150\001\255\255\255\255\255\255\154\001\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\162\001\163\001\255\255\165\001\ -\166\001\255\255\168\001\255\255\255\255\211\003\212\003\213\003\ -\214\003\215\003\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\232\003\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\053\004\255\255\255\255\255\255\ +\114\001\115\001\116\001\117\001\118\001\255\255\255\255\255\255\ +\255\255\080\002\215\003\216\003\217\003\218\003\219\003\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\001\001\002\001\003\001\004\001\087\004\006\001\ -\255\255\008\001\255\255\255\255\011\001\012\001\013\001\014\001\ -\015\001\016\001\017\001\255\255\019\001\020\001\021\001\022\001\ -\023\001\024\001\025\001\255\255\027\001\028\001\029\001\030\001\ +\255\255\255\255\255\255\141\001\255\255\255\255\255\255\236\003\ +\146\001\255\255\053\000\054\000\255\255\255\255\109\002\255\255\ +\154\001\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\255\255\255\255\166\001\072\000\255\255\255\255\ +\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ +\013\004\014\004\255\255\141\002\255\255\255\255\144\002\145\002\ +\146\002\147\002\148\002\149\002\150\002\151\002\152\002\153\002\ +\154\002\155\002\156\002\157\002\158\002\159\002\160\002\161\002\ +\162\002\255\255\164\002\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\255\255\170\002\255\255\175\002\053\000\054\000\ +\255\255\255\255\177\002\178\002\255\255\255\255\255\255\185\002\ +\255\255\184\002\255\255\255\255\255\255\255\255\067\004\255\255\ +\255\255\072\000\255\255\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ +\255\255\156\000\157\000\158\000\159\000\160\000\091\004\255\255\ +\216\002\164\000\165\000\166\000\255\255\168\000\221\002\222\002\ +\255\255\225\002\255\255\104\004\255\255\176\000\177\000\255\255\ +\232\002\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ +\255\255\249\002\255\255\255\255\053\000\054\000\253\002\255\255\ +\255\255\255\255\255\255\011\001\255\255\255\255\255\255\009\003\ +\255\255\255\255\255\255\019\001\255\255\012\003\255\255\072\000\ +\255\255\255\255\255\255\255\255\223\000\029\001\225\000\226\000\ +\227\000\255\255\255\255\255\255\255\255\255\255\233\000\030\003\ +\255\255\255\255\255\255\255\255\044\001\240\000\241\000\255\255\ +\255\255\049\001\050\001\178\000\247\000\255\255\255\255\255\255\ +\255\255\255\255\255\255\255\255\060\001\061\001\062\001\063\001\ +\064\001\255\255\255\255\192\004\255\255\069\001\070\001\255\255\ +\255\255\073\001\255\255\255\255\076\001\255\255\078\001\255\255\ +\255\255\255\255\255\255\255\255\084\001\085\001\025\001\255\255\ +\255\255\255\255\090\001\085\003\255\255\255\255\255\255\255\255\ +\088\003\255\255\225\000\226\000\227\000\255\255\255\255\255\255\ +\255\255\255\255\233\000\255\255\255\255\255\255\159\000\160\000\ +\104\003\240\000\241\000\164\000\255\255\166\000\255\255\168\000\ +\247\000\113\003\255\255\255\255\124\001\255\255\255\255\176\000\ +\128\001\255\255\255\255\255\255\132\001\255\255\255\255\255\255\ +\255\255\255\255\255\255\139\001\140\001\255\255\255\255\143\001\ +\255\255\255\255\255\255\255\255\255\255\255\255\255\255\151\001\ +\255\255\153\001\025\001\000\000\150\003\157\001\158\001\159\001\ +\160\001\161\001\153\003\154\003\255\255\255\255\255\255\255\255\ +\107\001\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ +\225\000\226\000\227\000\255\255\255\255\255\255\255\255\122\001\ +\233\000\255\255\255\255\255\255\255\255\255\255\129\001\240\000\ +\241\000\255\255\255\255\255\255\190\003\255\255\247\000\255\255\ +\255\255\140\001\255\255\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\149\001\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ +\215\003\216\003\217\003\218\003\219\003\255\255\222\003\255\255\ +\025\001\255\255\255\255\255\255\001\001\255\255\255\255\255\255\ +\255\255\255\255\234\003\008\001\255\255\236\003\255\255\255\255\ +\255\255\014\001\043\001\255\255\255\255\255\255\255\255\255\255\ +\247\003\022\001\023\001\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\010\004\040\001\255\255\255\255\255\255\255\255\ +\255\255\255\255\255\255\255\255\223\001\255\255\255\255\255\255\ +\255\255\228\001\229\001\230\001\255\255\232\001\255\255\255\255\ +\032\004\255\255\237\001\255\255\255\255\255\255\255\255\255\255\ +\069\001\255\255\255\255\255\255\247\001\255\255\255\255\255\255\ +\255\255\252\001\049\004\254\001\255\255\052\004\255\255\054\004\ +\255\255\056\004\255\255\088\001\255\255\008\002\009\002\255\255\ +\255\255\255\255\255\255\255\255\255\255\098\001\099\001\100\001\ +\101\001\102\001\103\001\104\001\255\255\255\255\255\255\026\002\ +\255\255\110\001\255\255\140\001\255\255\255\255\085\004\255\255\ +\255\255\255\255\255\255\255\255\091\004\228\001\229\001\230\001\ +\125\001\232\001\255\255\098\004\255\255\255\255\237\001\102\004\ +\255\255\255\255\255\255\136\001\255\255\255\255\255\255\255\255\ +\247\001\255\255\114\004\115\004\116\004\117\004\118\004\255\255\ +\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ +\255\255\008\002\009\002\255\255\255\255\080\002\255\255\255\255\ +\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\255\255\026\002\255\255\002\001\003\001\150\004\ +\255\255\006\001\101\002\255\255\255\255\104\002\255\255\255\255\ +\255\255\108\002\109\002\255\255\017\001\255\255\255\255\020\001\ +\255\255\255\255\255\255\228\001\229\001\230\001\027\001\232\001\ +\175\004\255\255\255\255\255\255\237\001\255\255\035\001\255\255\ +\255\255\255\255\255\255\040\001\255\255\042\001\247\001\255\255\ +\255\255\046\001\255\255\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\055\001\056\001\057\001\058\001\059\001\008\002\ +\009\002\255\255\255\255\255\255\255\255\066\001\255\255\068\001\ +\255\255\255\255\255\255\255\255\255\255\074\001\075\001\170\002\ +\171\002\026\002\173\002\080\001\255\255\082\001\177\002\178\002\ +\255\255\255\255\255\255\255\255\089\001\255\255\091\001\092\001\ +\093\001\094\001\255\255\096\001\097\001\255\255\255\255\255\255\ +\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\111\001\112\001\113\001\114\001\115\001\116\001\ +\117\001\118\001\119\001\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\221\002\222\002\255\255\130\001\131\001\255\255\ +\255\255\255\255\255\255\001\001\255\255\255\255\255\255\255\255\ +\141\001\255\255\008\001\170\002\255\255\146\001\255\255\255\255\ +\014\001\150\001\177\002\178\002\101\002\154\001\255\255\104\002\ +\022\001\023\001\253\002\108\002\255\255\162\001\163\001\255\255\ +\165\001\166\001\255\255\168\001\255\255\255\255\255\255\255\255\ +\255\255\255\255\040\001\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ +\027\003\255\255\255\255\030\003\255\255\255\255\221\002\222\002\ +\255\255\255\255\255\255\255\255\255\255\255\255\255\255\069\001\ +\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ +\255\255\170\002\088\001\255\255\255\255\255\255\253\002\255\255\ +\177\002\178\002\255\255\255\255\098\001\099\001\100\001\101\001\ +\102\001\103\001\104\001\255\255\255\255\255\255\255\255\255\255\ +\110\001\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\255\255\255\255\255\255\255\255\255\255\125\001\ +\099\003\006\001\255\255\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\136\001\255\255\221\002\222\002\255\255\020\001\ +\255\255\000\000\255\255\255\255\255\255\255\255\027\001\255\255\ +\255\255\255\255\255\255\255\255\255\255\255\255\035\001\255\255\ +\255\255\255\255\255\255\255\255\255\255\042\001\255\255\255\255\ +\255\255\046\001\255\255\255\255\253\002\255\255\255\255\255\255\ +\255\255\255\255\255\255\056\001\057\001\058\001\059\001\154\003\ +\255\255\255\255\255\255\255\255\255\255\066\001\255\255\068\001\ +\255\255\255\255\255\255\255\255\255\255\074\001\075\001\255\255\ +\255\255\255\255\027\003\080\001\255\255\082\001\255\255\178\003\ +\255\255\255\255\255\255\255\255\089\001\255\255\091\001\092\001\ +\093\001\094\001\255\255\096\001\097\001\255\255\255\255\255\255\ +\195\003\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\111\001\112\001\113\001\114\001\115\001\116\001\ +\117\001\118\001\119\001\255\255\215\003\216\003\217\003\218\003\ +\219\003\255\255\255\255\255\255\255\255\130\001\131\001\255\255\ +\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ +\141\001\236\003\255\255\255\255\255\255\146\001\255\255\255\255\ +\255\255\150\001\255\255\246\003\255\255\154\001\255\255\255\255\ +\255\255\255\255\255\255\255\255\255\255\162\001\163\001\255\255\ +\165\001\166\001\255\255\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ +\215\003\216\003\217\003\218\003\219\003\000\000\255\255\255\255\ +\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\045\004\255\255\255\255\236\003\255\255\255\255\ +\255\255\255\255\255\255\255\255\255\255\255\255\057\004\255\255\ +\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ +\091\004\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\255\255\255\255\255\255\255\255\215\003\216\003\ +\217\003\218\003\219\003\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\001\001\002\001\003\001\004\001\255\255\006\001\ +\255\255\008\001\255\255\236\003\011\001\012\001\013\001\014\001\ +\015\001\016\001\017\001\255\255\019\001\020\001\021\001\022\001\ +\023\001\024\001\025\001\255\255\027\001\028\001\029\001\030\001\ \031\001\032\001\033\001\034\001\035\001\255\255\037\001\038\001\ -\255\255\040\001\041\001\042\001\255\255\044\001\255\255\046\001\ +\255\255\040\001\041\001\042\001\091\004\044\001\255\255\046\001\ \255\255\255\255\049\001\050\001\255\255\052\001\255\255\054\001\ \055\001\056\001\057\001\058\001\059\001\060\001\061\001\062\001\ \063\001\064\001\255\255\066\001\255\255\068\001\069\001\070\001\ -\255\255\087\004\073\001\074\001\075\001\076\001\077\001\078\001\ -\079\001\080\001\000\000\082\001\255\255\084\001\085\001\255\255\ -\255\255\088\001\089\001\090\001\091\001\092\001\093\001\094\001\ +\255\255\255\255\073\001\074\001\075\001\076\001\077\001\078\001\ +\079\001\080\001\255\255\082\001\255\255\084\001\085\001\000\000\ +\057\004\088\001\089\001\090\001\091\001\092\001\093\001\094\001\ \255\255\096\001\097\001\098\001\099\001\100\001\101\001\102\001\ \103\001\104\001\105\001\106\001\255\255\255\255\109\001\110\001\ \111\001\112\001\113\001\114\001\115\001\116\001\117\001\118\001\ -\119\001\255\255\255\255\255\255\255\255\124\001\125\001\126\001\ +\119\001\255\255\091\004\255\255\255\255\124\001\125\001\126\001\ \127\001\128\001\255\255\130\001\131\001\132\001\255\255\134\001\ \135\001\136\001\137\001\138\001\139\001\140\001\141\001\142\001\ \143\001\144\001\145\001\146\001\255\255\148\001\255\255\150\001\ \151\001\255\255\153\001\154\001\155\001\255\255\157\001\158\001\ \159\001\160\001\161\001\162\001\163\001\164\001\165\001\166\001\ -\167\001\168\001\001\001\002\001\003\001\004\001\255\255\006\001\ -\255\255\008\001\255\255\255\255\011\001\012\001\013\001\014\001\ +\167\001\168\001\255\255\255\255\255\255\255\255\001\001\002\001\ +\003\001\004\001\255\255\006\001\255\255\008\001\255\255\255\255\ +\011\001\012\001\013\001\014\001\015\001\016\001\017\001\255\255\ +\019\001\020\001\021\001\022\001\023\001\024\001\025\001\255\255\ +\027\001\028\001\029\001\255\255\255\255\255\255\255\255\034\001\ +\035\001\255\255\037\001\038\001\255\255\040\001\041\001\042\001\ +\255\255\044\001\255\255\046\001\255\255\255\255\049\001\050\001\ +\255\255\052\001\255\255\054\001\055\001\056\001\057\001\058\001\ +\059\001\060\001\061\001\062\001\063\001\064\001\255\255\066\001\ +\255\255\068\001\069\001\070\001\255\255\255\255\073\001\074\001\ +\075\001\076\001\077\001\078\001\079\001\080\001\255\255\082\001\ +\000\000\084\001\085\001\255\255\255\255\088\001\089\001\090\001\ +\091\001\092\001\093\001\094\001\255\255\096\001\097\001\098\001\ +\099\001\100\001\101\001\102\001\103\001\104\001\105\001\106\001\ +\255\255\255\255\109\001\110\001\111\001\112\001\113\001\114\001\ +\115\001\116\001\117\001\118\001\119\001\120\001\121\001\255\255\ +\255\255\124\001\125\001\126\001\127\001\128\001\255\255\130\001\ +\131\001\132\001\255\255\134\001\135\001\136\001\137\001\138\001\ +\139\001\140\001\141\001\142\001\143\001\144\001\145\001\146\001\ +\255\255\148\001\255\255\150\001\151\001\255\255\153\001\154\001\ +\155\001\156\001\157\001\158\001\159\001\160\001\161\001\162\001\ +\163\001\164\001\165\001\166\001\167\001\168\001\255\255\255\255\ +\001\001\002\001\003\001\004\001\255\255\006\001\255\255\008\001\ +\255\255\255\255\011\001\012\001\013\001\014\001\015\001\016\001\ +\017\001\255\255\019\001\020\001\021\001\022\001\023\001\024\001\ +\025\001\255\255\027\001\028\001\029\001\255\255\031\001\032\001\ +\033\001\034\001\035\001\255\255\037\001\038\001\255\255\040\001\ +\041\001\042\001\255\255\044\001\255\255\046\001\255\255\255\255\ +\049\001\050\001\255\255\052\001\255\255\054\001\055\001\056\001\ +\057\001\058\001\059\001\060\001\061\001\062\001\063\001\064\001\ +\255\255\066\001\255\255\068\001\069\001\070\001\255\255\255\255\ +\073\001\074\001\075\001\076\001\077\001\078\001\079\001\080\001\ +\000\000\082\001\255\255\084\001\085\001\255\255\255\255\088\001\ +\089\001\090\001\091\001\092\001\093\001\094\001\255\255\096\001\ +\097\001\098\001\099\001\100\001\101\001\102\001\103\001\104\001\ +\105\001\106\001\255\255\255\255\109\001\110\001\111\001\112\001\ +\113\001\114\001\115\001\116\001\117\001\118\001\119\001\255\255\ +\255\255\255\255\255\255\124\001\125\001\126\001\127\001\128\001\ +\255\255\130\001\131\001\132\001\255\255\134\001\135\001\136\001\ +\137\001\138\001\139\001\140\001\141\001\142\001\143\001\144\001\ +\145\001\146\001\255\255\148\001\255\255\150\001\151\001\255\255\ +\153\001\154\001\155\001\255\255\157\001\158\001\159\001\160\001\ +\161\001\162\001\163\001\164\001\165\001\166\001\167\001\168\001\ +\255\255\001\001\002\001\003\001\004\001\255\255\006\001\255\255\ +\008\001\255\255\255\255\255\255\012\001\013\001\014\001\015\001\ +\016\001\017\001\255\255\019\001\020\001\021\001\022\001\023\001\ +\024\001\025\001\255\255\027\001\028\001\255\255\030\001\031\001\ +\032\001\033\001\034\001\035\001\255\255\037\001\038\001\255\255\ +\040\001\041\001\042\001\255\255\044\001\255\255\046\001\255\255\ +\255\255\049\001\050\001\255\255\052\001\255\255\054\001\055\001\ +\056\001\057\001\058\001\059\001\060\001\061\001\062\001\063\001\ +\064\001\255\255\066\001\255\255\068\001\069\001\255\255\255\255\ +\255\255\073\001\074\001\075\001\255\255\077\001\078\001\079\001\ +\080\001\000\000\082\001\083\001\084\001\085\001\255\255\255\255\ +\088\001\089\001\090\001\091\001\092\001\093\001\094\001\255\255\ +\096\001\097\001\098\001\099\001\100\001\101\001\102\001\103\001\ +\104\001\105\001\106\001\255\255\255\255\109\001\110\001\111\001\ +\112\001\113\001\114\001\115\001\116\001\117\001\118\001\119\001\ +\255\255\255\255\255\255\255\255\124\001\125\001\126\001\127\001\ +\128\001\255\255\130\001\131\001\132\001\255\255\134\001\135\001\ +\136\001\137\001\138\001\139\001\140\001\141\001\142\001\143\001\ +\144\001\145\001\146\001\255\255\148\001\255\255\150\001\151\001\ +\255\255\153\001\154\001\255\255\156\001\157\001\158\001\159\001\ +\160\001\161\001\162\001\163\001\164\001\165\001\166\001\255\255\ +\168\001\001\001\002\001\003\001\004\001\255\255\006\001\255\255\ +\008\001\255\255\255\255\011\001\012\001\013\001\014\001\015\001\ +\016\001\017\001\255\255\019\001\020\001\021\001\022\001\023\001\ +\024\001\025\001\255\255\027\001\028\001\029\001\255\255\255\255\ +\255\255\255\255\255\255\035\001\255\255\037\001\038\001\255\255\ +\040\001\041\001\042\001\255\255\044\001\255\255\046\001\255\255\ +\255\255\049\001\050\001\255\255\052\001\255\255\054\001\055\001\ +\056\001\057\001\058\001\059\001\060\001\061\001\062\001\063\001\ +\064\001\255\255\066\001\255\255\068\001\069\001\070\001\255\255\ +\255\255\073\001\074\001\075\001\076\001\077\001\078\001\079\001\ +\080\001\000\000\082\001\255\255\084\001\085\001\255\255\255\255\ +\088\001\089\001\090\001\091\001\092\001\093\001\094\001\255\255\ +\096\001\097\001\098\001\099\001\100\001\101\001\102\001\103\001\ +\104\001\105\001\106\001\255\255\255\255\109\001\110\001\111\001\ +\112\001\113\001\114\001\115\001\116\001\117\001\118\001\119\001\ +\255\255\255\255\255\255\255\255\124\001\125\001\126\001\127\001\ +\128\001\255\255\130\001\131\001\132\001\255\255\134\001\135\001\ +\136\001\137\001\138\001\139\001\140\001\141\001\142\001\143\001\ +\144\001\145\001\146\001\255\255\148\001\255\255\150\001\151\001\ +\255\255\153\001\154\001\155\001\255\255\157\001\158\001\159\001\ +\160\001\161\001\162\001\163\001\164\001\165\001\166\001\167\001\ +\168\001\255\255\001\001\002\001\003\001\004\001\255\255\006\001\ +\255\255\008\001\255\255\255\255\255\255\012\001\013\001\014\001\ \015\001\016\001\017\001\255\255\019\001\020\001\021\001\022\001\ -\023\001\024\001\025\001\255\255\027\001\028\001\029\001\255\255\ -\255\255\255\255\255\255\034\001\035\001\255\255\037\001\038\001\ +\023\001\024\001\025\001\255\255\027\001\028\001\255\255\255\255\ +\031\001\032\001\033\001\034\001\035\001\255\255\037\001\038\001\ \255\255\040\001\041\001\042\001\255\255\044\001\255\255\046\001\ \255\255\255\255\049\001\050\001\255\255\052\001\255\255\054\001\ \055\001\056\001\057\001\058\001\059\001\060\001\061\001\062\001\ -\063\001\064\001\255\255\066\001\255\255\068\001\069\001\070\001\ -\255\255\255\255\073\001\074\001\075\001\076\001\077\001\078\001\ +\063\001\064\001\255\255\066\001\255\255\068\001\069\001\255\255\ +\255\255\255\255\073\001\074\001\075\001\255\255\077\001\078\001\ \079\001\080\001\000\000\082\001\255\255\084\001\085\001\255\255\ \255\255\088\001\089\001\090\001\091\001\092\001\093\001\094\001\ \255\255\096\001\097\001\098\001\099\001\100\001\101\001\102\001\ \103\001\104\001\105\001\106\001\255\255\255\255\109\001\110\001\ \111\001\112\001\113\001\114\001\115\001\116\001\117\001\118\001\ -\119\001\120\001\121\001\255\255\255\255\124\001\125\001\126\001\ +\119\001\255\255\255\255\255\255\255\255\124\001\125\001\126\001\ \127\001\128\001\255\255\130\001\131\001\132\001\255\255\134\001\ \135\001\136\001\137\001\138\001\139\001\140\001\141\001\142\001\ \143\001\144\001\145\001\146\001\255\255\148\001\255\255\150\001\ -\151\001\255\255\153\001\154\001\155\001\156\001\157\001\158\001\ +\151\001\255\255\153\001\154\001\255\255\255\255\157\001\158\001\ +\159\001\160\001\161\001\162\001\163\001\164\001\165\001\166\001\ +\255\255\168\001\001\001\002\001\003\001\004\001\255\255\006\001\ +\255\255\008\001\255\255\255\255\255\255\012\001\013\001\255\255\ +\015\001\016\001\255\255\255\255\019\001\020\001\021\001\022\001\ +\023\001\024\001\025\001\255\255\027\001\028\001\255\255\255\255\ +\031\001\032\001\033\001\034\001\035\001\255\255\037\001\038\001\ +\255\255\040\001\041\001\042\001\255\255\044\001\255\255\046\001\ +\255\255\255\255\049\001\050\001\255\255\052\001\255\255\054\001\ +\055\001\056\001\057\001\058\001\059\001\060\001\061\001\062\001\ +\063\001\064\001\255\255\066\001\255\255\068\001\069\001\255\255\ +\255\255\255\255\073\001\074\001\075\001\255\255\077\001\078\001\ +\079\001\080\001\000\000\082\001\083\001\084\001\085\001\255\255\ +\255\255\088\001\089\001\090\001\091\001\092\001\093\001\094\001\ +\255\255\096\001\097\001\098\001\099\001\100\001\101\001\102\001\ +\103\001\104\001\105\001\106\001\255\255\255\255\109\001\110\001\ +\111\001\112\001\113\001\114\001\115\001\116\001\117\001\118\001\ +\119\001\255\255\255\255\255\255\255\255\124\001\125\001\126\001\ +\127\001\128\001\255\255\130\001\131\001\132\001\255\255\134\001\ +\135\001\136\001\137\001\138\001\139\001\140\001\141\001\255\255\ +\143\001\144\001\145\001\146\001\255\255\148\001\255\255\150\001\ +\151\001\255\255\153\001\154\001\255\255\255\255\157\001\158\001\ \159\001\160\001\161\001\162\001\163\001\164\001\165\001\166\001\ -\167\001\168\001\255\255\001\001\002\001\003\001\004\001\255\255\ -\006\001\255\255\008\001\255\255\255\255\011\001\012\001\013\001\ -\014\001\015\001\016\001\017\001\255\255\019\001\020\001\021\001\ -\022\001\023\001\024\001\025\001\255\255\027\001\028\001\029\001\ -\255\255\031\001\032\001\033\001\034\001\035\001\255\255\037\001\ +\255\255\168\001\255\255\001\001\002\001\003\001\004\001\255\255\ +\006\001\255\255\008\001\255\255\255\255\255\255\012\001\013\001\ +\014\001\015\001\016\001\017\001\255\255\019\001\020\001\255\255\ +\022\001\023\001\024\001\025\001\255\255\027\001\028\001\255\255\ +\255\255\255\255\255\255\255\255\255\255\035\001\255\255\037\001\ \038\001\255\255\040\001\041\001\042\001\255\255\044\001\255\255\ \046\001\255\255\255\255\049\001\050\001\255\255\052\001\255\255\ \054\001\055\001\056\001\057\001\058\001\059\001\060\001\061\001\ \062\001\063\001\064\001\255\255\066\001\255\255\068\001\069\001\ -\070\001\255\255\255\255\073\001\074\001\075\001\076\001\077\001\ +\255\255\255\255\255\255\073\001\074\001\075\001\255\255\077\001\ \078\001\079\001\080\001\000\000\082\001\255\255\084\001\085\001\ \255\255\255\255\088\001\089\001\090\001\091\001\092\001\093\001\ \094\001\255\255\096\001\097\001\098\001\099\001\100\001\101\001\ @@ -4303,19 +4441,19 @@ let yycheck = "\002\000\ \126\001\127\001\128\001\255\255\130\001\131\001\132\001\255\255\ \134\001\135\001\136\001\137\001\138\001\139\001\140\001\141\001\ \142\001\143\001\144\001\145\001\146\001\255\255\148\001\255\255\ -\150\001\151\001\255\255\153\001\154\001\155\001\255\255\157\001\ +\150\001\151\001\255\255\153\001\154\001\255\255\255\255\157\001\ \158\001\159\001\160\001\161\001\162\001\163\001\164\001\165\001\ -\166\001\167\001\168\001\001\001\002\001\003\001\004\001\255\255\ +\166\001\255\255\168\001\001\001\002\001\003\001\004\001\255\255\ \006\001\255\255\008\001\255\255\255\255\255\255\012\001\013\001\ -\014\001\015\001\016\001\017\001\255\255\019\001\020\001\021\001\ +\014\001\015\001\016\001\017\001\255\255\019\001\020\001\255\255\ \022\001\023\001\024\001\025\001\255\255\027\001\028\001\255\255\ -\030\001\031\001\032\001\033\001\034\001\035\001\255\255\037\001\ +\255\255\255\255\255\255\255\255\255\255\035\001\255\255\037\001\ \038\001\255\255\040\001\041\001\042\001\255\255\044\001\255\255\ \046\001\255\255\255\255\049\001\050\001\255\255\052\001\255\255\ \054\001\055\001\056\001\057\001\058\001\059\001\060\001\061\001\ \062\001\063\001\064\001\255\255\066\001\255\255\068\001\069\001\ \255\255\255\255\255\255\073\001\074\001\075\001\255\255\077\001\ -\078\001\079\001\080\001\000\000\082\001\083\001\084\001\085\001\ +\078\001\079\001\080\001\000\000\082\001\255\255\084\001\085\001\ \255\255\255\255\088\001\089\001\090\001\091\001\092\001\093\001\ \094\001\255\255\096\001\097\001\098\001\099\001\100\001\101\001\ \102\001\103\001\104\001\105\001\106\001\255\255\255\255\109\001\ @@ -4324,18 +4462,18 @@ let yycheck = "\002\000\ \126\001\127\001\128\001\255\255\130\001\131\001\132\001\255\255\ \134\001\135\001\136\001\137\001\138\001\139\001\140\001\141\001\ \142\001\143\001\144\001\145\001\146\001\255\255\148\001\255\255\ -\150\001\151\001\255\255\153\001\154\001\255\255\156\001\157\001\ +\150\001\151\001\255\255\153\001\154\001\255\255\255\255\157\001\ \158\001\159\001\160\001\161\001\162\001\163\001\164\001\165\001\ \166\001\255\255\168\001\255\255\001\001\002\001\003\001\004\001\ -\255\255\006\001\255\255\008\001\255\255\255\255\011\001\012\001\ +\255\255\006\001\255\255\008\001\255\255\255\255\255\255\012\001\ \013\001\014\001\015\001\016\001\017\001\255\255\019\001\020\001\ -\021\001\022\001\023\001\024\001\025\001\255\255\027\001\028\001\ -\029\001\255\255\255\255\255\255\255\255\255\255\035\001\255\255\ +\255\255\022\001\023\001\024\001\025\001\255\255\027\001\028\001\ +\255\255\255\255\255\255\255\255\255\255\255\255\035\001\255\255\ \037\001\038\001\255\255\040\001\041\001\042\001\255\255\044\001\ \255\255\046\001\255\255\255\255\049\001\050\001\255\255\052\001\ \255\255\054\001\055\001\056\001\057\001\058\001\059\001\060\001\ \061\001\062\001\063\001\064\001\255\255\066\001\255\255\068\001\ -\069\001\070\001\255\255\255\255\073\001\074\001\075\001\076\001\ +\069\001\255\255\255\255\255\255\073\001\074\001\075\001\255\255\ \077\001\078\001\079\001\080\001\000\000\082\001\255\255\084\001\ \085\001\255\255\255\255\088\001\089\001\090\001\091\001\092\001\ \093\001\094\001\255\255\096\001\097\001\098\001\099\001\100\001\ @@ -4345,13 +4483,13 @@ let yycheck = "\002\000\ \125\001\126\001\127\001\128\001\255\255\130\001\131\001\132\001\ \255\255\134\001\135\001\136\001\137\001\138\001\139\001\140\001\ \141\001\142\001\143\001\144\001\145\001\146\001\255\255\148\001\ -\255\255\150\001\151\001\255\255\153\001\154\001\155\001\255\255\ +\255\255\150\001\151\001\255\255\153\001\154\001\255\255\255\255\ \157\001\158\001\159\001\160\001\161\001\162\001\163\001\164\001\ -\165\001\166\001\167\001\168\001\001\001\002\001\003\001\004\001\ +\165\001\166\001\255\255\168\001\001\001\002\001\003\001\004\001\ \255\255\006\001\255\255\008\001\255\255\255\255\255\255\012\001\ \013\001\014\001\015\001\016\001\017\001\255\255\019\001\020\001\ -\021\001\022\001\023\001\024\001\025\001\255\255\027\001\028\001\ -\255\255\255\255\031\001\032\001\033\001\034\001\035\001\255\255\ +\255\255\022\001\023\001\024\001\025\001\255\255\027\001\028\001\ +\255\255\255\255\255\255\255\255\255\255\255\255\035\001\255\255\ \037\001\038\001\255\255\040\001\041\001\042\001\255\255\044\001\ \255\255\046\001\255\255\255\255\049\001\050\001\255\255\052\001\ \255\255\054\001\055\001\056\001\057\001\058\001\059\001\060\001\ @@ -4370,15 +4508,15 @@ let yycheck = "\002\000\ \157\001\158\001\159\001\160\001\161\001\162\001\163\001\164\001\ \165\001\166\001\255\255\168\001\255\255\001\001\002\001\003\001\ \004\001\255\255\006\001\255\255\008\001\255\255\255\255\255\255\ -\012\001\013\001\255\255\015\001\016\001\255\255\255\255\019\001\ -\020\001\021\001\022\001\023\001\024\001\025\001\255\255\027\001\ -\028\001\255\255\255\255\031\001\032\001\033\001\034\001\035\001\ +\012\001\013\001\014\001\015\001\016\001\017\001\255\255\019\001\ +\020\001\255\255\022\001\023\001\024\001\025\001\255\255\027\001\ +\028\001\255\255\255\255\255\255\255\255\255\255\255\255\035\001\ \255\255\037\001\038\001\255\255\040\001\041\001\042\001\255\255\ \044\001\255\255\046\001\255\255\255\255\049\001\050\001\255\255\ \052\001\255\255\054\001\055\001\056\001\057\001\058\001\059\001\ \060\001\061\001\062\001\063\001\064\001\255\255\066\001\255\255\ \068\001\069\001\255\255\255\255\255\255\073\001\074\001\075\001\ -\255\255\077\001\078\001\079\001\080\001\000\000\082\001\083\001\ +\255\255\077\001\078\001\079\001\080\001\000\000\082\001\255\255\ \084\001\085\001\255\255\255\255\088\001\089\001\090\001\091\001\ \092\001\093\001\094\001\255\255\096\001\097\001\098\001\099\001\ \100\001\101\001\102\001\103\001\104\001\105\001\106\001\255\255\ @@ -4386,12 +4524,12 @@ let yycheck = "\002\000\ \116\001\117\001\118\001\119\001\255\255\255\255\255\255\255\255\ \124\001\125\001\126\001\127\001\128\001\255\255\130\001\131\001\ \132\001\255\255\134\001\135\001\136\001\137\001\138\001\139\001\ -\140\001\141\001\255\255\143\001\144\001\145\001\146\001\255\255\ +\140\001\141\001\142\001\143\001\144\001\145\001\146\001\255\255\ \148\001\255\255\150\001\151\001\255\255\153\001\154\001\255\255\ \255\255\157\001\158\001\159\001\160\001\161\001\162\001\163\001\ \164\001\165\001\166\001\255\255\168\001\001\001\002\001\003\001\ \004\001\255\255\006\001\255\255\008\001\255\255\255\255\255\255\ -\012\001\013\001\014\001\015\001\016\001\017\001\255\255\019\001\ +\012\001\013\001\255\255\015\001\016\001\255\255\255\255\019\001\ \020\001\255\255\022\001\023\001\024\001\025\001\255\255\027\001\ \028\001\255\255\255\255\255\255\255\255\255\255\255\255\035\001\ \255\255\037\001\038\001\255\255\040\001\041\001\042\001\255\255\ @@ -4407,12 +4545,12 @@ let yycheck = "\002\000\ \116\001\117\001\118\001\119\001\255\255\255\255\255\255\255\255\ \124\001\125\001\126\001\127\001\128\001\255\255\130\001\131\001\ \132\001\255\255\134\001\135\001\136\001\137\001\138\001\139\001\ -\140\001\141\001\142\001\143\001\144\001\145\001\146\001\255\255\ +\140\001\141\001\255\255\143\001\144\001\145\001\146\001\255\255\ \148\001\255\255\150\001\151\001\255\255\153\001\154\001\255\255\ \255\255\157\001\158\001\159\001\160\001\161\001\162\001\163\001\ \164\001\165\001\166\001\255\255\168\001\255\255\001\001\002\001\ \003\001\004\001\255\255\006\001\255\255\008\001\255\255\255\255\ -\255\255\012\001\013\001\014\001\015\001\016\001\017\001\255\255\ +\255\255\012\001\013\001\255\255\015\001\016\001\255\255\255\255\ \019\001\020\001\255\255\022\001\023\001\024\001\025\001\255\255\ \027\001\028\001\255\255\255\255\255\255\255\255\255\255\255\255\ \035\001\255\255\037\001\038\001\255\255\040\001\041\001\042\001\ @@ -4428,12 +4566,12 @@ let yycheck = "\002\000\ \115\001\116\001\117\001\118\001\119\001\255\255\255\255\255\255\ \255\255\124\001\125\001\126\001\127\001\128\001\255\255\130\001\ \131\001\132\001\255\255\134\001\135\001\136\001\137\001\138\001\ -\139\001\140\001\141\001\142\001\143\001\144\001\145\001\146\001\ +\139\001\140\001\141\001\255\255\143\001\144\001\145\001\146\001\ \255\255\148\001\255\255\150\001\151\001\255\255\153\001\154\001\ \255\255\255\255\157\001\158\001\159\001\160\001\161\001\162\001\ \163\001\164\001\165\001\166\001\255\255\168\001\001\001\002\001\ \003\001\004\001\255\255\006\001\255\255\008\001\255\255\255\255\ -\255\255\012\001\013\001\014\001\015\001\016\001\017\001\255\255\ +\255\255\012\001\013\001\255\255\015\001\016\001\255\255\255\255\ \019\001\020\001\255\255\022\001\023\001\024\001\025\001\255\255\ \027\001\028\001\255\255\255\255\255\255\255\255\255\255\255\255\ \035\001\255\255\037\001\038\001\255\255\040\001\041\001\042\001\ @@ -4449,12 +4587,12 @@ let yycheck = "\002\000\ \115\001\116\001\117\001\118\001\119\001\255\255\255\255\255\255\ \255\255\124\001\125\001\126\001\127\001\128\001\255\255\130\001\ \131\001\132\001\255\255\134\001\135\001\136\001\137\001\138\001\ -\139\001\140\001\141\001\142\001\143\001\144\001\145\001\146\001\ +\139\001\140\001\141\001\255\255\143\001\144\001\145\001\146\001\ \255\255\148\001\255\255\150\001\151\001\255\255\153\001\154\001\ \255\255\255\255\157\001\158\001\159\001\160\001\161\001\162\001\ \163\001\164\001\165\001\166\001\255\255\168\001\255\255\001\001\ \002\001\003\001\004\001\255\255\006\001\255\255\008\001\255\255\ -\255\255\255\255\012\001\013\001\014\001\015\001\016\001\017\001\ +\255\255\255\255\012\001\013\001\255\255\015\001\016\001\255\255\ \255\255\019\001\020\001\255\255\022\001\023\001\024\001\025\001\ \255\255\027\001\028\001\255\255\255\255\255\255\255\255\255\255\ \255\255\035\001\255\255\037\001\038\001\255\255\040\001\041\001\ @@ -4470,12 +4608,12 @@ let yycheck = "\002\000\ \114\001\115\001\116\001\117\001\118\001\119\001\255\255\255\255\ \255\255\255\255\124\001\125\001\126\001\127\001\128\001\255\255\ \130\001\131\001\132\001\255\255\134\001\135\001\136\001\137\001\ -\138\001\139\001\140\001\141\001\142\001\143\001\144\001\145\001\ +\138\001\139\001\140\001\141\001\255\255\143\001\144\001\145\001\ \146\001\255\255\148\001\255\255\150\001\151\001\255\255\153\001\ \154\001\255\255\255\255\157\001\158\001\159\001\160\001\161\001\ \162\001\163\001\164\001\165\001\166\001\255\255\168\001\001\001\ \002\001\003\001\004\001\255\255\006\001\255\255\008\001\255\255\ -\255\255\255\255\012\001\013\001\014\001\015\001\016\001\017\001\ +\255\255\255\255\012\001\013\001\255\255\015\001\016\001\255\255\ \255\255\019\001\020\001\255\255\022\001\023\001\024\001\025\001\ \255\255\027\001\028\001\255\255\255\255\255\255\255\255\255\255\ \255\255\035\001\255\255\037\001\038\001\255\255\040\001\041\001\ @@ -4491,129 +4629,129 @@ let yycheck = "\002\000\ \114\001\115\001\116\001\117\001\118\001\119\001\255\255\255\255\ \255\255\255\255\124\001\125\001\126\001\127\001\128\001\255\255\ \130\001\131\001\132\001\255\255\134\001\135\001\136\001\137\001\ -\138\001\139\001\140\001\141\001\142\001\143\001\144\001\145\001\ +\138\001\139\001\140\001\141\001\255\255\143\001\144\001\145\001\ \146\001\255\255\148\001\255\255\150\001\151\001\255\255\153\001\ \154\001\255\255\255\255\157\001\158\001\159\001\160\001\161\001\ \162\001\163\001\164\001\165\001\166\001\255\255\168\001\255\255\ \001\001\002\001\003\001\004\001\255\255\006\001\255\255\008\001\ -\255\255\255\255\255\255\012\001\013\001\255\255\015\001\016\001\ -\255\255\255\255\019\001\020\001\255\255\022\001\023\001\024\001\ +\255\255\255\255\255\255\255\255\013\001\014\001\015\001\255\255\ +\017\001\255\255\255\255\020\001\255\255\022\001\023\001\024\001\ \025\001\255\255\027\001\028\001\255\255\255\255\255\255\255\255\ \255\255\255\255\035\001\255\255\037\001\038\001\255\255\040\001\ -\041\001\042\001\255\255\044\001\255\255\046\001\255\255\255\255\ -\049\001\050\001\255\255\052\001\255\255\054\001\055\001\056\001\ -\057\001\058\001\059\001\060\001\061\001\062\001\063\001\064\001\ +\041\001\042\001\255\255\255\255\255\255\046\001\255\255\255\255\ +\255\255\255\255\255\255\052\001\255\255\054\001\055\001\056\001\ +\057\001\058\001\059\001\255\255\255\255\255\255\255\255\255\255\ \255\255\066\001\255\255\068\001\069\001\255\255\255\255\255\255\ -\073\001\074\001\075\001\255\255\077\001\078\001\079\001\080\001\ -\000\000\082\001\255\255\084\001\085\001\255\255\255\255\088\001\ -\089\001\090\001\091\001\092\001\093\001\094\001\255\255\096\001\ +\255\255\074\001\075\001\255\255\255\255\255\255\079\001\080\001\ +\000\000\082\001\255\255\255\255\255\255\255\255\255\255\088\001\ +\089\001\255\255\091\001\092\001\093\001\094\001\255\255\096\001\ \097\001\098\001\099\001\100\001\101\001\102\001\103\001\104\001\ -\105\001\106\001\255\255\255\255\109\001\110\001\111\001\112\001\ +\105\001\255\255\255\255\255\255\255\255\110\001\111\001\112\001\ \113\001\114\001\115\001\116\001\117\001\118\001\119\001\255\255\ -\255\255\255\255\255\255\124\001\125\001\126\001\127\001\128\001\ -\255\255\130\001\131\001\132\001\255\255\134\001\135\001\136\001\ -\137\001\138\001\139\001\140\001\141\001\255\255\143\001\144\001\ -\145\001\146\001\255\255\148\001\255\255\150\001\151\001\255\255\ -\153\001\154\001\255\255\255\255\157\001\158\001\159\001\160\001\ -\161\001\162\001\163\001\164\001\165\001\166\001\255\255\168\001\ +\255\255\255\255\255\255\255\255\125\001\126\001\127\001\255\255\ +\255\255\130\001\131\001\255\255\255\255\134\001\135\001\136\001\ +\137\001\138\001\255\255\255\255\141\001\142\001\255\255\144\001\ +\145\001\146\001\255\255\148\001\255\255\150\001\255\255\255\255\ +\255\255\154\001\255\255\255\255\255\255\255\255\255\255\255\255\ +\255\255\162\001\163\001\255\255\165\001\166\001\255\255\168\001\ \001\001\002\001\003\001\004\001\255\255\006\001\255\255\008\001\ -\255\255\255\255\255\255\012\001\013\001\255\255\015\001\016\001\ -\255\255\255\255\019\001\020\001\255\255\022\001\023\001\024\001\ +\255\255\255\255\255\255\255\255\013\001\014\001\015\001\255\255\ +\017\001\255\255\255\255\020\001\255\255\022\001\023\001\024\001\ \025\001\255\255\027\001\028\001\255\255\255\255\255\255\255\255\ \255\255\255\255\035\001\255\255\037\001\038\001\255\255\040\001\ -\041\001\042\001\255\255\044\001\255\255\046\001\255\255\255\255\ -\049\001\050\001\255\255\052\001\255\255\054\001\055\001\056\001\ -\057\001\058\001\059\001\060\001\061\001\062\001\063\001\064\001\ -\255\255\066\001\255\255\068\001\069\001\255\255\255\255\255\255\ -\073\001\074\001\075\001\255\255\077\001\078\001\079\001\080\001\ -\000\000\082\001\255\255\084\001\085\001\255\255\255\255\088\001\ -\089\001\090\001\091\001\092\001\093\001\094\001\255\255\096\001\ +\041\001\042\001\255\255\255\255\255\255\046\001\255\255\255\255\ +\255\255\255\255\255\255\052\001\255\255\054\001\055\001\056\001\ +\057\001\058\001\059\001\255\255\255\255\255\255\255\255\255\255\ +\255\255\066\001\255\255\068\001\255\255\255\255\255\255\255\255\ +\255\255\074\001\075\001\255\255\255\255\255\255\079\001\080\001\ +\000\000\082\001\255\255\255\255\255\255\255\255\255\255\088\001\ +\089\001\255\255\091\001\092\001\093\001\094\001\255\255\096\001\ \097\001\098\001\099\001\100\001\101\001\102\001\103\001\104\001\ -\105\001\106\001\255\255\255\255\109\001\110\001\111\001\112\001\ +\105\001\255\255\255\255\255\255\255\255\110\001\111\001\112\001\ \113\001\114\001\115\001\116\001\117\001\118\001\119\001\255\255\ -\255\255\255\255\255\255\124\001\125\001\126\001\127\001\128\001\ -\255\255\130\001\131\001\132\001\255\255\134\001\135\001\136\001\ -\137\001\138\001\139\001\140\001\141\001\255\255\143\001\144\001\ -\145\001\146\001\255\255\148\001\255\255\150\001\151\001\255\255\ -\153\001\154\001\255\255\255\255\157\001\158\001\159\001\160\001\ -\161\001\162\001\163\001\164\001\165\001\166\001\255\255\168\001\ +\255\255\255\255\255\255\255\255\125\001\126\001\127\001\255\255\ +\255\255\130\001\131\001\255\255\255\255\134\001\135\001\136\001\ +\137\001\138\001\255\255\255\255\141\001\142\001\255\255\144\001\ +\145\001\146\001\255\255\148\001\255\255\150\001\255\255\255\255\ +\255\255\154\001\255\255\255\255\255\255\255\255\255\255\255\255\ +\255\255\162\001\163\001\255\255\165\001\166\001\255\255\168\001\ \255\255\001\001\002\001\003\001\004\001\255\255\006\001\255\255\ -\008\001\255\255\255\255\255\255\012\001\013\001\255\255\015\001\ -\016\001\255\255\255\255\019\001\020\001\255\255\022\001\023\001\ +\008\001\255\255\255\255\255\255\255\255\013\001\014\001\015\001\ +\255\255\017\001\255\255\255\255\020\001\255\255\022\001\023\001\ \024\001\025\001\255\255\027\001\028\001\255\255\255\255\255\255\ \255\255\255\255\255\255\035\001\255\255\037\001\038\001\255\255\ -\040\001\041\001\042\001\255\255\044\001\255\255\046\001\255\255\ -\255\255\049\001\050\001\255\255\052\001\255\255\054\001\055\001\ -\056\001\057\001\058\001\059\001\060\001\061\001\062\001\063\001\ -\064\001\255\255\066\001\255\255\068\001\069\001\255\255\255\255\ -\255\255\073\001\074\001\075\001\255\255\077\001\078\001\079\001\ -\080\001\000\000\082\001\255\255\084\001\085\001\255\255\255\255\ -\088\001\089\001\090\001\091\001\092\001\093\001\094\001\255\255\ +\040\001\041\001\042\001\255\255\255\255\255\255\046\001\255\255\ +\255\255\255\255\255\255\255\255\052\001\255\255\054\001\055\001\ +\056\001\057\001\058\001\059\001\255\255\255\255\255\255\255\255\ +\255\255\255\255\066\001\255\255\068\001\255\255\255\255\255\255\ +\255\255\255\255\074\001\075\001\255\255\255\255\255\255\079\001\ +\080\001\000\000\082\001\255\255\255\255\255\255\255\255\255\255\ +\088\001\089\001\255\255\091\001\092\001\093\001\094\001\255\255\ \096\001\097\001\098\001\099\001\100\001\101\001\102\001\103\001\ -\104\001\105\001\106\001\255\255\255\255\109\001\110\001\111\001\ +\104\001\105\001\255\255\255\255\255\255\255\255\110\001\111\001\ \112\001\113\001\114\001\115\001\116\001\117\001\118\001\119\001\ -\255\255\255\255\255\255\255\255\124\001\125\001\126\001\127\001\ -\128\001\255\255\130\001\131\001\132\001\255\255\134\001\135\001\ -\136\001\137\001\138\001\139\001\140\001\141\001\255\255\143\001\ -\144\001\145\001\146\001\255\255\148\001\255\255\150\001\151\001\ -\255\255\153\001\154\001\255\255\255\255\157\001\158\001\159\001\ -\160\001\161\001\162\001\163\001\164\001\165\001\166\001\255\255\ +\255\255\255\255\255\255\255\255\255\255\125\001\126\001\127\001\ +\255\255\255\255\130\001\131\001\255\255\255\255\134\001\135\001\ +\136\001\137\001\138\001\255\255\255\255\141\001\142\001\255\255\ +\144\001\145\001\146\001\255\255\148\001\255\255\150\001\255\255\ +\255\255\255\255\154\001\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\162\001\163\001\255\255\165\001\166\001\255\255\ \168\001\001\001\002\001\003\001\004\001\255\255\006\001\255\255\ -\008\001\255\255\255\255\255\255\012\001\013\001\255\255\015\001\ -\016\001\255\255\255\255\019\001\020\001\255\255\022\001\023\001\ +\008\001\255\255\255\255\255\255\255\255\013\001\014\001\015\001\ +\255\255\017\001\255\255\255\255\020\001\255\255\022\001\023\001\ \024\001\025\001\255\255\027\001\028\001\255\255\255\255\255\255\ \255\255\255\255\255\255\035\001\255\255\037\001\038\001\255\255\ -\040\001\041\001\042\001\255\255\044\001\255\255\046\001\255\255\ -\255\255\049\001\050\001\255\255\052\001\255\255\054\001\055\001\ -\056\001\057\001\058\001\059\001\060\001\061\001\062\001\063\001\ -\064\001\255\255\066\001\255\255\068\001\069\001\255\255\255\255\ -\255\255\073\001\074\001\075\001\255\255\077\001\078\001\079\001\ -\080\001\000\000\082\001\255\255\084\001\085\001\255\255\255\255\ -\088\001\089\001\090\001\091\001\092\001\093\001\094\001\255\255\ +\040\001\041\001\042\001\255\255\255\255\255\255\046\001\255\255\ +\255\255\255\255\255\255\255\255\052\001\255\255\054\001\055\001\ +\056\001\057\001\058\001\059\001\255\255\255\255\255\255\255\255\ +\255\255\255\255\066\001\255\255\068\001\255\255\255\255\255\255\ +\255\255\255\255\074\001\075\001\255\255\255\255\255\255\079\001\ +\080\001\000\000\082\001\255\255\255\255\255\255\255\255\255\255\ +\088\001\089\001\255\255\091\001\092\001\093\001\094\001\255\255\ \096\001\097\001\098\001\099\001\100\001\101\001\102\001\103\001\ -\104\001\105\001\106\001\255\255\255\255\109\001\110\001\111\001\ +\104\001\255\255\255\255\255\255\255\255\255\255\110\001\111\001\ \112\001\113\001\114\001\115\001\116\001\117\001\118\001\119\001\ -\255\255\255\255\255\255\255\255\124\001\125\001\126\001\127\001\ -\128\001\255\255\130\001\131\001\132\001\255\255\134\001\135\001\ -\136\001\137\001\138\001\139\001\140\001\141\001\255\255\143\001\ -\144\001\145\001\146\001\255\255\148\001\255\255\150\001\151\001\ -\255\255\153\001\154\001\255\255\255\255\157\001\158\001\159\001\ -\160\001\161\001\162\001\163\001\164\001\165\001\166\001\255\255\ +\255\255\255\255\255\255\255\255\255\255\125\001\126\001\127\001\ +\255\255\255\255\130\001\131\001\255\255\255\255\134\001\135\001\ +\136\001\137\001\138\001\255\255\255\255\141\001\142\001\255\255\ +\144\001\145\001\146\001\255\255\148\001\255\255\150\001\255\255\ +\255\255\255\255\154\001\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\162\001\163\001\255\255\165\001\166\001\255\255\ \168\001\255\255\001\001\002\001\003\001\004\001\255\255\006\001\ -\255\255\008\001\255\255\255\255\255\255\012\001\013\001\255\255\ -\015\001\016\001\255\255\255\255\019\001\020\001\255\255\022\001\ +\255\255\008\001\255\255\255\255\255\255\255\255\013\001\014\001\ +\015\001\255\255\017\001\255\255\255\255\020\001\255\255\022\001\ \023\001\024\001\025\001\255\255\027\001\028\001\255\255\255\255\ \255\255\255\255\255\255\255\255\035\001\255\255\037\001\038\001\ -\255\255\040\001\041\001\042\001\255\255\044\001\255\255\046\001\ -\255\255\255\255\049\001\050\001\255\255\052\001\255\255\054\001\ -\055\001\056\001\057\001\058\001\059\001\060\001\061\001\062\001\ -\063\001\064\001\255\255\066\001\255\255\068\001\069\001\255\255\ -\255\255\255\255\073\001\074\001\075\001\255\255\077\001\078\001\ -\079\001\080\001\000\000\082\001\255\255\084\001\085\001\255\255\ -\255\255\088\001\089\001\090\001\091\001\092\001\093\001\094\001\ -\255\255\096\001\097\001\098\001\099\001\100\001\101\001\102\001\ -\103\001\104\001\105\001\106\001\255\255\255\255\109\001\110\001\ -\111\001\112\001\113\001\114\001\115\001\116\001\117\001\118\001\ -\119\001\255\255\255\255\255\255\255\255\124\001\125\001\126\001\ -\127\001\128\001\255\255\130\001\131\001\132\001\255\255\134\001\ -\135\001\136\001\137\001\138\001\139\001\140\001\141\001\255\255\ -\143\001\144\001\145\001\146\001\255\255\148\001\255\255\150\001\ -\151\001\255\255\153\001\154\001\255\255\255\255\157\001\158\001\ -\159\001\160\001\161\001\162\001\163\001\164\001\165\001\166\001\ +\255\255\040\001\041\001\042\001\255\255\255\255\255\255\046\001\ +\255\255\255\255\255\255\255\255\255\255\052\001\255\255\054\001\ +\055\001\056\001\057\001\058\001\059\001\255\255\255\255\255\255\ +\255\255\255\255\255\255\066\001\255\255\068\001\255\255\255\255\ +\255\255\255\255\255\255\074\001\075\001\255\255\255\255\255\255\ +\079\001\080\001\000\000\082\001\255\255\255\255\255\255\255\255\ +\255\255\088\001\089\001\255\255\091\001\092\001\093\001\094\001\ +\255\255\096\001\097\001\098\001\099\001\100\001\101\001\102\001\ +\103\001\104\001\255\255\255\255\255\255\255\255\255\255\110\001\ +\111\001\112\001\113\001\114\001\115\001\116\001\117\001\118\001\ +\119\001\255\255\255\255\255\255\255\255\255\255\125\001\126\001\ +\127\001\255\255\255\255\130\001\131\001\255\255\255\255\134\001\ +\135\001\136\001\137\001\138\001\255\255\255\255\141\001\142\001\ +\255\255\144\001\145\001\146\001\255\255\148\001\255\255\150\001\ +\255\255\255\255\255\255\154\001\255\255\255\255\255\255\255\255\ +\255\255\255\255\255\255\162\001\163\001\255\255\165\001\166\001\ \255\255\168\001\001\001\002\001\003\001\004\001\255\255\006\001\ -\255\255\008\001\255\255\255\255\255\255\255\255\013\001\014\001\ +\255\255\255\255\255\255\255\255\255\255\255\255\013\001\014\001\ \015\001\255\255\017\001\255\255\255\255\020\001\255\255\022\001\ \023\001\024\001\025\001\255\255\027\001\028\001\255\255\255\255\ \255\255\255\255\255\255\255\255\035\001\255\255\037\001\038\001\ \255\255\040\001\041\001\042\001\255\255\255\255\255\255\046\001\ \255\255\255\255\255\255\255\255\255\255\052\001\255\255\054\001\ \055\001\056\001\057\001\058\001\059\001\255\255\255\255\255\255\ -\255\255\255\255\255\255\066\001\255\255\068\001\069\001\255\255\ +\255\255\255\255\255\255\066\001\255\255\068\001\255\255\255\255\ \255\255\255\255\255\255\074\001\075\001\255\255\255\255\255\255\ \079\001\080\001\000\000\082\001\255\255\255\255\255\255\255\255\ \255\255\088\001\089\001\255\255\091\001\092\001\093\001\094\001\ \255\255\096\001\097\001\098\001\099\001\100\001\101\001\102\001\ -\103\001\104\001\105\001\255\255\255\255\255\255\255\255\110\001\ +\103\001\104\001\255\255\255\255\255\255\255\255\255\255\110\001\ \111\001\112\001\113\001\114\001\115\001\116\001\117\001\118\001\ \119\001\255\255\255\255\255\255\255\255\255\255\125\001\126\001\ \127\001\255\255\255\255\130\001\131\001\255\255\255\255\134\001\ @@ -4621,8 +4759,8 @@ let yycheck = "\002\000\ \255\255\144\001\145\001\146\001\255\255\148\001\255\255\150\001\ \255\255\255\255\255\255\154\001\255\255\255\255\255\255\255\255\ \255\255\255\255\255\255\162\001\163\001\255\255\165\001\166\001\ -\255\255\168\001\255\255\001\001\002\001\003\001\004\001\255\255\ -\006\001\255\255\008\001\255\255\255\255\255\255\255\255\013\001\ +\255\255\168\001\255\255\255\255\002\001\003\001\004\001\255\255\ +\006\001\255\255\255\255\255\255\255\255\255\255\255\255\013\001\ \014\001\015\001\255\255\017\001\255\255\255\255\020\001\255\255\ \022\001\023\001\024\001\025\001\255\255\027\001\028\001\255\255\ \255\255\255\255\255\255\255\255\255\255\035\001\255\255\037\001\ @@ -4634,7 +4772,7 @@ let yycheck = "\002\000\ \255\255\079\001\080\001\000\000\082\001\255\255\255\255\255\255\ \255\255\255\255\088\001\089\001\255\255\091\001\092\001\093\001\ \094\001\255\255\096\001\097\001\098\001\099\001\100\001\101\001\ -\102\001\103\001\104\001\105\001\255\255\255\255\255\255\255\255\ +\102\001\103\001\255\255\255\255\255\255\255\255\255\255\255\255\ \110\001\111\001\112\001\113\001\114\001\115\001\116\001\117\001\ \118\001\119\001\255\255\255\255\255\255\255\255\255\255\125\001\ \126\001\127\001\255\255\255\255\130\001\131\001\255\255\255\255\ @@ -4642,10 +4780,10 @@ let yycheck = "\002\000\ \142\001\255\255\144\001\145\001\146\001\255\255\148\001\255\255\ \150\001\255\255\255\255\255\255\154\001\255\255\255\255\255\255\ \255\255\255\255\255\255\255\255\162\001\163\001\255\255\165\001\ -\166\001\255\255\168\001\001\001\002\001\003\001\004\001\255\255\ -\006\001\255\255\008\001\255\255\255\255\255\255\255\255\013\001\ +\166\001\255\255\168\001\255\255\002\001\003\001\004\001\255\255\ +\006\001\255\255\255\255\255\255\255\255\255\255\255\255\013\001\ \014\001\015\001\255\255\017\001\255\255\255\255\020\001\255\255\ -\022\001\023\001\024\001\025\001\255\255\027\001\028\001\255\255\ +\255\255\023\001\024\001\025\001\255\255\027\001\028\001\255\255\ \255\255\255\255\255\255\255\255\255\255\035\001\255\255\037\001\ \038\001\255\255\040\001\041\001\042\001\255\255\255\255\255\255\ \046\001\255\255\255\255\255\255\255\255\255\255\052\001\255\255\ @@ -4655,7 +4793,7 @@ let yycheck = "\002\000\ \255\255\079\001\080\001\000\000\082\001\255\255\255\255\255\255\ \255\255\255\255\088\001\089\001\255\255\091\001\092\001\093\001\ \094\001\255\255\096\001\097\001\098\001\099\001\100\001\101\001\ -\102\001\103\001\104\001\105\001\255\255\255\255\255\255\255\255\ +\102\001\103\001\255\255\255\255\255\255\255\255\255\255\255\255\ \110\001\111\001\112\001\113\001\114\001\115\001\116\001\117\001\ \118\001\119\001\255\255\255\255\255\255\255\255\255\255\125\001\ \126\001\127\001\255\255\255\255\130\001\131\001\255\255\255\255\ @@ -4663,10 +4801,10 @@ let yycheck = "\002\000\ \142\001\255\255\144\001\145\001\146\001\255\255\148\001\255\255\ \150\001\255\255\255\255\255\255\154\001\255\255\255\255\255\255\ \255\255\255\255\255\255\255\255\162\001\163\001\255\255\165\001\ -\166\001\255\255\168\001\255\255\001\001\002\001\003\001\004\001\ -\255\255\006\001\255\255\008\001\255\255\255\255\255\255\255\255\ +\166\001\255\255\168\001\255\255\255\255\002\001\003\001\004\001\ +\255\255\006\001\255\255\255\255\255\255\255\255\255\255\255\255\ \013\001\014\001\015\001\255\255\017\001\255\255\255\255\020\001\ -\255\255\022\001\023\001\024\001\025\001\255\255\027\001\028\001\ +\255\255\255\255\023\001\024\001\025\001\255\255\027\001\028\001\ \255\255\255\255\255\255\255\255\255\255\255\255\035\001\255\255\ \037\001\038\001\255\255\040\001\041\001\042\001\255\255\255\255\ \255\255\046\001\255\255\255\255\255\255\255\255\255\255\052\001\ @@ -4676,7 +4814,7 @@ let yycheck = "\002\000\ \255\255\255\255\079\001\080\001\000\000\082\001\255\255\255\255\ \255\255\255\255\255\255\088\001\089\001\255\255\091\001\092\001\ \093\001\094\001\255\255\096\001\097\001\098\001\099\001\100\001\ -\101\001\102\001\103\001\104\001\255\255\255\255\255\255\255\255\ +\101\001\102\001\103\001\255\255\255\255\255\255\255\255\255\255\ \255\255\110\001\111\001\112\001\113\001\114\001\115\001\116\001\ \117\001\118\001\119\001\255\255\255\255\255\255\255\255\255\255\ \125\001\126\001\127\001\255\255\255\255\130\001\131\001\255\255\ @@ -4684,10 +4822,10 @@ let yycheck = "\002\000\ \141\001\142\001\255\255\144\001\145\001\146\001\255\255\148\001\ \255\255\150\001\255\255\255\255\255\255\154\001\255\255\255\255\ \255\255\255\255\255\255\255\255\255\255\162\001\163\001\255\255\ -\165\001\166\001\255\255\168\001\001\001\002\001\003\001\004\001\ -\255\255\006\001\255\255\008\001\255\255\255\255\255\255\255\255\ +\165\001\166\001\255\255\168\001\255\255\002\001\003\001\004\001\ +\255\255\006\001\255\255\255\255\255\255\255\255\255\255\255\255\ \013\001\014\001\015\001\255\255\017\001\255\255\255\255\020\001\ -\255\255\022\001\023\001\024\001\025\001\255\255\027\001\028\001\ +\255\255\255\255\023\001\024\001\025\001\255\255\027\001\028\001\ \255\255\255\255\255\255\255\255\255\255\255\255\035\001\255\255\ \037\001\038\001\255\255\040\001\041\001\042\001\255\255\255\255\ \255\255\046\001\255\255\255\255\255\255\255\255\255\255\052\001\ @@ -4695,9 +4833,9 @@ let yycheck = "\002\000\ \255\255\255\255\255\255\255\255\255\255\066\001\255\255\068\001\ \255\255\255\255\255\255\255\255\255\255\074\001\075\001\255\255\ \255\255\255\255\079\001\080\001\000\000\082\001\255\255\255\255\ -\255\255\255\255\255\255\088\001\089\001\255\255\091\001\092\001\ +\255\255\255\255\255\255\255\255\089\001\255\255\091\001\092\001\ \093\001\094\001\255\255\096\001\097\001\098\001\099\001\100\001\ -\101\001\102\001\103\001\104\001\255\255\255\255\255\255\255\255\ +\101\001\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ \255\255\110\001\111\001\112\001\113\001\114\001\115\001\116\001\ \117\001\118\001\119\001\255\255\255\255\255\255\255\255\255\255\ \125\001\126\001\127\001\255\255\255\255\130\001\131\001\255\255\ @@ -4705,10 +4843,10 @@ let yycheck = "\002\000\ \141\001\142\001\255\255\144\001\145\001\146\001\255\255\148\001\ \255\255\150\001\255\255\255\255\255\255\154\001\255\255\255\255\ \255\255\255\255\255\255\255\255\255\255\162\001\163\001\255\255\ -\165\001\166\001\255\255\168\001\255\255\001\001\002\001\003\001\ +\165\001\166\001\255\255\168\001\255\255\255\255\002\001\003\001\ \004\001\255\255\006\001\255\255\255\255\255\255\255\255\255\255\ \255\255\013\001\014\001\015\001\255\255\017\001\255\255\255\255\ -\020\001\255\255\022\001\023\001\024\001\025\001\255\255\027\001\ +\020\001\255\255\255\255\023\001\024\001\025\001\255\255\027\001\ \028\001\255\255\255\255\255\255\255\255\255\255\255\255\035\001\ \255\255\037\001\038\001\255\255\040\001\041\001\042\001\255\255\ \255\255\255\255\046\001\255\255\255\255\255\255\255\255\255\255\ @@ -4716,9 +4854,9 @@ let yycheck = "\002\000\ \255\255\255\255\255\255\255\255\255\255\255\255\066\001\255\255\ \068\001\255\255\255\255\255\255\255\255\255\255\074\001\075\001\ \255\255\255\255\255\255\079\001\080\001\000\000\082\001\255\255\ -\255\255\255\255\255\255\255\255\088\001\089\001\255\255\091\001\ +\255\255\255\255\255\255\255\255\255\255\089\001\255\255\091\001\ \092\001\093\001\094\001\255\255\096\001\097\001\098\001\099\001\ -\100\001\101\001\102\001\103\001\104\001\255\255\255\255\255\255\ +\100\001\101\001\255\255\255\255\255\255\255\255\255\255\255\255\ \255\255\255\255\110\001\111\001\112\001\113\001\114\001\115\001\ \116\001\117\001\118\001\119\001\255\255\255\255\255\255\255\255\ \255\255\125\001\126\001\127\001\255\255\255\255\130\001\131\001\ @@ -4729,7 +4867,7 @@ let yycheck = "\002\000\ \255\255\165\001\166\001\255\255\168\001\255\255\002\001\003\001\ \004\001\255\255\006\001\255\255\255\255\255\255\255\255\255\255\ \255\255\013\001\014\001\015\001\255\255\017\001\255\255\255\255\ -\020\001\255\255\022\001\023\001\024\001\025\001\255\255\027\001\ +\020\001\255\255\255\255\023\001\024\001\025\001\255\255\027\001\ \028\001\255\255\255\255\255\255\255\255\255\255\255\255\035\001\ \255\255\037\001\038\001\255\255\040\001\041\001\042\001\255\255\ \255\255\255\255\046\001\255\255\255\255\255\255\255\255\255\255\ @@ -4737,10 +4875,10 @@ let yycheck = "\002\000\ \255\255\255\255\255\255\255\255\255\255\255\255\066\001\255\255\ \068\001\255\255\255\255\255\255\255\255\255\255\074\001\075\001\ \255\255\255\255\255\255\079\001\080\001\000\000\082\001\255\255\ -\255\255\255\255\255\255\255\255\088\001\089\001\255\255\091\001\ +\255\255\255\255\255\255\255\255\255\255\089\001\255\255\091\001\ \092\001\093\001\094\001\255\255\096\001\097\001\098\001\099\001\ -\100\001\101\001\102\001\103\001\255\255\255\255\255\255\255\255\ -\255\255\255\255\110\001\111\001\112\001\113\001\114\001\115\001\ +\100\001\101\001\255\255\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\255\255\111\001\112\001\113\001\114\001\115\001\ \116\001\117\001\118\001\119\001\255\255\255\255\255\255\255\255\ \255\255\125\001\126\001\127\001\255\255\255\255\130\001\131\001\ \255\255\255\255\134\001\135\001\136\001\137\001\138\001\255\255\ @@ -4758,10 +4896,10 @@ let yycheck = "\002\000\ \059\001\255\255\255\255\255\255\255\255\255\255\255\255\066\001\ \255\255\068\001\255\255\255\255\255\255\255\255\255\255\074\001\ \075\001\255\255\255\255\255\255\079\001\080\001\000\000\082\001\ -\255\255\255\255\255\255\255\255\255\255\088\001\089\001\255\255\ +\255\255\255\255\255\255\255\255\255\255\255\255\089\001\255\255\ \091\001\092\001\093\001\094\001\255\255\096\001\097\001\098\001\ -\099\001\100\001\101\001\102\001\103\001\255\255\255\255\255\255\ -\255\255\255\255\255\255\110\001\111\001\112\001\113\001\114\001\ +\099\001\100\001\255\255\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\255\255\255\255\111\001\112\001\113\001\114\001\ \115\001\116\001\117\001\118\001\119\001\255\255\255\255\255\255\ \255\255\255\255\125\001\126\001\127\001\255\255\255\255\130\001\ \131\001\255\255\255\255\134\001\135\001\136\001\137\001\138\001\ @@ -4779,10 +4917,10 @@ let yycheck = "\002\000\ \059\001\255\255\255\255\255\255\255\255\255\255\255\255\066\001\ \255\255\068\001\255\255\255\255\255\255\255\255\255\255\074\001\ \075\001\255\255\255\255\255\255\079\001\080\001\000\000\082\001\ -\255\255\255\255\255\255\255\255\255\255\088\001\089\001\255\255\ +\255\255\255\255\255\255\255\255\255\255\255\255\089\001\255\255\ \091\001\092\001\093\001\094\001\255\255\096\001\097\001\098\001\ -\099\001\100\001\101\001\102\001\103\001\255\255\255\255\255\255\ -\255\255\255\255\255\255\110\001\111\001\112\001\113\001\114\001\ +\099\001\100\001\255\255\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\255\255\255\255\111\001\112\001\113\001\114\001\ \115\001\116\001\117\001\118\001\119\001\255\255\255\255\255\255\ \255\255\255\255\125\001\126\001\127\001\255\255\255\255\130\001\ \131\001\255\255\255\255\134\001\135\001\136\001\137\001\138\001\ @@ -4794,7 +4932,7 @@ let yycheck = "\002\000\ \255\255\255\255\255\255\013\001\014\001\015\001\255\255\017\001\ \255\255\255\255\020\001\255\255\255\255\023\001\024\001\025\001\ \255\255\027\001\028\001\255\255\255\255\255\255\255\255\255\255\ -\255\255\035\001\255\255\037\001\038\001\255\255\040\001\041\001\ +\255\255\035\001\255\255\037\001\038\001\255\255\255\255\041\001\ \042\001\255\255\255\255\255\255\046\001\255\255\255\255\255\255\ \255\255\255\255\052\001\255\255\054\001\055\001\056\001\057\001\ \058\001\059\001\255\255\255\255\255\255\255\255\255\255\255\255\ @@ -4802,8 +4940,8 @@ let yycheck = "\002\000\ \074\001\075\001\255\255\255\255\255\255\079\001\080\001\000\000\ \082\001\255\255\255\255\255\255\255\255\255\255\255\255\089\001\ \255\255\091\001\092\001\093\001\094\001\255\255\096\001\097\001\ -\098\001\099\001\100\001\101\001\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\110\001\111\001\112\001\113\001\ +\098\001\099\001\255\255\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\255\255\255\255\255\255\111\001\112\001\113\001\ \114\001\115\001\116\001\117\001\118\001\119\001\255\255\255\255\ \255\255\255\255\255\255\125\001\126\001\127\001\255\255\255\255\ \130\001\131\001\255\255\255\255\134\001\135\001\136\001\137\001\ @@ -4815,7 +4953,7 @@ let yycheck = "\002\000\ \255\255\255\255\255\255\013\001\014\001\015\001\255\255\017\001\ \255\255\255\255\020\001\255\255\255\255\023\001\024\001\025\001\ \255\255\027\001\028\001\255\255\255\255\255\255\255\255\255\255\ -\255\255\035\001\255\255\037\001\038\001\255\255\040\001\041\001\ +\255\255\035\001\255\255\037\001\038\001\255\255\255\255\041\001\ \042\001\255\255\255\255\255\255\046\001\255\255\255\255\255\255\ \255\255\255\255\052\001\255\255\054\001\055\001\056\001\057\001\ \058\001\059\001\255\255\255\255\255\255\255\255\255\255\255\255\ @@ -4823,8 +4961,8 @@ let yycheck = "\002\000\ \074\001\075\001\255\255\255\255\255\255\079\001\080\001\000\000\ \082\001\255\255\255\255\255\255\255\255\255\255\255\255\089\001\ \255\255\091\001\092\001\093\001\094\001\255\255\096\001\097\001\ -\098\001\099\001\100\001\101\001\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\110\001\111\001\112\001\113\001\ +\098\001\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\255\255\255\255\255\255\111\001\112\001\113\001\ \114\001\115\001\116\001\117\001\118\001\119\001\255\255\255\255\ \255\255\255\255\255\255\125\001\126\001\127\001\255\255\255\255\ \130\001\131\001\255\255\255\255\134\001\135\001\136\001\137\001\ @@ -4834,9 +4972,9 @@ let yycheck = "\002\000\ \162\001\163\001\255\255\165\001\166\001\255\255\168\001\255\255\ \255\255\002\001\003\001\004\001\255\255\006\001\255\255\255\255\ \255\255\255\255\255\255\255\255\013\001\014\001\015\001\255\255\ -\017\001\255\255\255\255\020\001\255\255\255\255\023\001\024\001\ +\017\001\255\255\255\255\020\001\255\255\255\255\255\255\024\001\ \025\001\255\255\027\001\028\001\255\255\255\255\255\255\255\255\ -\255\255\255\255\035\001\255\255\037\001\038\001\255\255\040\001\ +\255\255\255\255\035\001\255\255\037\001\038\001\255\255\255\255\ \041\001\042\001\255\255\255\255\255\255\046\001\255\255\255\255\ \255\255\255\255\255\255\052\001\255\255\054\001\055\001\056\001\ \057\001\058\001\059\001\255\255\255\255\255\255\255\255\255\255\ @@ -4844,7 +4982,7 @@ let yycheck = "\002\000\ \255\255\074\001\075\001\255\255\255\255\255\255\079\001\080\001\ \000\000\082\001\255\255\255\255\255\255\255\255\255\255\255\255\ \089\001\255\255\091\001\092\001\093\001\094\001\255\255\096\001\ -\097\001\098\001\099\001\100\001\101\001\255\255\255\255\255\255\ +\097\001\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\111\001\112\001\ \113\001\114\001\115\001\116\001\117\001\118\001\119\001\255\255\ \255\255\255\255\255\255\255\255\125\001\126\001\127\001\255\255\ @@ -4855,9 +4993,9 @@ let yycheck = "\002\000\ \255\255\162\001\163\001\255\255\165\001\166\001\255\255\168\001\ \255\255\002\001\003\001\004\001\255\255\006\001\255\255\255\255\ \255\255\255\255\255\255\255\255\013\001\014\001\015\001\255\255\ -\017\001\255\255\255\255\020\001\255\255\255\255\023\001\024\001\ +\017\001\255\255\255\255\020\001\255\255\255\255\255\255\024\001\ \025\001\255\255\027\001\028\001\255\255\255\255\255\255\255\255\ -\255\255\255\255\035\001\255\255\037\001\038\001\255\255\040\001\ +\255\255\255\255\035\001\255\255\037\001\038\001\255\255\255\255\ \041\001\042\001\255\255\255\255\255\255\046\001\255\255\255\255\ \255\255\255\255\255\255\052\001\255\255\054\001\055\001\056\001\ \057\001\058\001\059\001\255\255\255\255\255\255\255\255\255\255\ @@ -4865,7 +5003,7 @@ let yycheck = "\002\000\ \255\255\074\001\075\001\255\255\255\255\255\255\079\001\080\001\ \000\000\082\001\255\255\255\255\255\255\255\255\255\255\255\255\ \089\001\255\255\091\001\092\001\093\001\094\001\255\255\096\001\ -\097\001\098\001\099\001\100\001\255\255\255\255\255\255\255\255\ +\097\001\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\111\001\112\001\ \113\001\114\001\115\001\116\001\117\001\118\001\119\001\255\255\ \255\255\255\255\255\255\255\255\125\001\126\001\127\001\255\255\ @@ -4876,17 +5014,17 @@ let yycheck = "\002\000\ \255\255\162\001\163\001\255\255\165\001\166\001\255\255\168\001\ \255\255\255\255\002\001\003\001\004\001\255\255\006\001\255\255\ \255\255\255\255\255\255\255\255\255\255\013\001\014\001\015\001\ -\255\255\017\001\255\255\255\255\020\001\255\255\255\255\023\001\ -\024\001\025\001\255\255\027\001\028\001\255\255\255\255\255\255\ +\255\255\017\001\255\255\255\255\020\001\255\255\255\255\255\255\ +\255\255\025\001\255\255\027\001\028\001\255\255\255\255\255\255\ \255\255\255\255\255\255\035\001\255\255\037\001\038\001\255\255\ -\040\001\041\001\042\001\255\255\255\255\255\255\046\001\255\255\ +\255\255\041\001\042\001\255\255\255\255\255\255\046\001\255\255\ \255\255\255\255\255\255\255\255\052\001\255\255\054\001\055\001\ \056\001\057\001\058\001\059\001\255\255\255\255\255\255\255\255\ \255\255\255\255\066\001\255\255\068\001\255\255\255\255\255\255\ \255\255\255\255\074\001\075\001\255\255\255\255\255\255\079\001\ \080\001\000\000\082\001\255\255\255\255\255\255\255\255\255\255\ \255\255\089\001\255\255\091\001\092\001\093\001\094\001\255\255\ -\096\001\097\001\098\001\099\001\100\001\255\255\255\255\255\255\ +\096\001\097\001\255\255\255\255\255\255\255\255\255\255\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\111\001\ \112\001\113\001\114\001\115\001\116\001\117\001\118\001\119\001\ \255\255\255\255\255\255\255\255\255\255\125\001\126\001\127\001\ @@ -4897,17 +5035,17 @@ let yycheck = "\002\000\ \255\255\255\255\162\001\163\001\255\255\165\001\166\001\255\255\ \168\001\255\255\002\001\003\001\004\001\255\255\006\001\255\255\ \255\255\255\255\255\255\255\255\255\255\013\001\014\001\015\001\ -\255\255\017\001\255\255\255\255\020\001\255\255\255\255\023\001\ -\024\001\025\001\255\255\027\001\028\001\255\255\255\255\255\255\ +\255\255\017\001\255\255\255\255\020\001\255\255\255\255\255\255\ +\255\255\255\255\255\255\027\001\028\001\255\255\255\255\255\255\ \255\255\255\255\255\255\035\001\255\255\037\001\038\001\255\255\ \255\255\041\001\042\001\255\255\255\255\255\255\046\001\255\255\ \255\255\255\255\255\255\255\255\052\001\255\255\054\001\055\001\ \056\001\057\001\058\001\059\001\255\255\255\255\255\255\255\255\ -\255\255\255\255\066\001\255\255\068\001\255\255\255\255\255\255\ +\255\255\255\255\066\001\000\000\068\001\255\255\255\255\255\255\ \255\255\255\255\074\001\075\001\255\255\255\255\255\255\079\001\ -\080\001\000\000\082\001\255\255\255\255\255\255\255\255\255\255\ +\080\001\255\255\082\001\255\255\255\255\255\255\255\255\255\255\ \255\255\089\001\255\255\091\001\092\001\093\001\094\001\255\255\ -\096\001\097\001\098\001\099\001\255\255\255\255\255\255\255\255\ +\096\001\097\001\255\255\255\255\255\255\255\255\255\255\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\111\001\ \112\001\113\001\114\001\115\001\116\001\117\001\118\001\119\001\ \255\255\255\255\255\255\255\255\255\255\125\001\126\001\127\001\ @@ -4919,35 +5057,14 @@ let yycheck = "\002\000\ \168\001\255\255\255\255\002\001\003\001\004\001\255\255\006\001\ \255\255\255\255\255\255\255\255\255\255\255\255\013\001\014\001\ \015\001\255\255\017\001\255\255\255\255\020\001\255\255\255\255\ -\023\001\024\001\025\001\255\255\027\001\028\001\255\255\255\255\ -\255\255\255\255\255\255\255\255\035\001\255\255\037\001\038\001\ -\255\255\255\255\041\001\042\001\255\255\255\255\255\255\046\001\ -\255\255\255\255\255\255\255\255\255\255\052\001\255\255\054\001\ -\055\001\056\001\057\001\058\001\059\001\255\255\255\255\255\255\ -\255\255\255\255\255\255\066\001\255\255\068\001\255\255\255\255\ -\255\255\255\255\255\255\074\001\075\001\255\255\255\255\255\255\ -\079\001\080\001\000\000\082\001\255\255\255\255\255\255\255\255\ -\255\255\255\255\089\001\255\255\091\001\092\001\093\001\094\001\ -\255\255\096\001\097\001\098\001\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\111\001\112\001\113\001\114\001\115\001\116\001\117\001\118\001\ -\119\001\255\255\255\255\255\255\255\255\255\255\125\001\126\001\ -\127\001\255\255\255\255\130\001\131\001\255\255\255\255\134\001\ -\135\001\136\001\137\001\138\001\255\255\255\255\141\001\142\001\ -\255\255\144\001\145\001\146\001\255\255\148\001\255\255\150\001\ -\255\255\255\255\255\255\154\001\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\162\001\163\001\255\255\165\001\166\001\ -\255\255\168\001\255\255\002\001\003\001\004\001\255\255\006\001\ -\255\255\255\255\255\255\255\255\255\255\255\255\013\001\014\001\ -\015\001\255\255\017\001\255\255\255\255\020\001\255\255\255\255\ -\255\255\024\001\025\001\255\255\027\001\028\001\255\255\255\255\ +\255\255\255\255\255\255\255\255\027\001\028\001\255\255\255\255\ \255\255\255\255\255\255\255\255\035\001\255\255\037\001\038\001\ \255\255\255\255\041\001\042\001\255\255\255\255\255\255\046\001\ \255\255\255\255\255\255\255\255\255\255\052\001\255\255\054\001\ \055\001\056\001\057\001\058\001\059\001\255\255\255\255\255\255\ -\255\255\255\255\255\255\066\001\255\255\068\001\255\255\255\255\ +\255\255\255\255\000\000\066\001\255\255\068\001\255\255\255\255\ \255\255\255\255\255\255\074\001\075\001\255\255\255\255\255\255\ -\079\001\080\001\000\000\082\001\255\255\255\255\255\255\255\255\ +\079\001\080\001\255\255\082\001\255\255\255\255\255\255\255\255\ \255\255\255\255\089\001\255\255\091\001\092\001\093\001\094\001\ \255\255\096\001\097\001\255\255\255\255\255\255\255\255\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ @@ -4956,80 +5073,15 @@ let yycheck = "\002\000\ \127\001\255\255\255\255\130\001\131\001\255\255\255\255\134\001\ \135\001\136\001\137\001\138\001\255\255\255\255\141\001\142\001\ \255\255\144\001\145\001\146\001\255\255\148\001\255\255\150\001\ -\255\255\255\255\255\255\154\001\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\162\001\163\001\255\255\165\001\166\001\ -\255\255\168\001\255\255\255\255\002\001\003\001\004\001\255\255\ -\006\001\255\255\255\255\255\255\255\255\255\255\255\255\013\001\ -\014\001\015\001\255\255\017\001\255\255\255\255\020\001\255\255\ -\255\255\255\255\024\001\025\001\255\255\027\001\028\001\255\255\ -\255\255\255\255\255\255\255\255\255\255\035\001\255\255\037\001\ -\038\001\255\255\255\255\041\001\042\001\255\255\255\255\255\255\ -\046\001\255\255\255\255\255\255\255\255\255\255\052\001\255\255\ -\054\001\055\001\056\001\057\001\058\001\059\001\255\255\255\255\ -\255\255\255\255\255\255\255\255\066\001\255\255\068\001\255\255\ -\255\255\255\255\255\255\255\255\074\001\075\001\255\255\255\255\ -\255\255\079\001\080\001\000\000\082\001\255\255\255\255\255\255\ -\255\255\255\255\255\255\089\001\255\255\091\001\092\001\093\001\ -\094\001\255\255\096\001\097\001\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\111\001\112\001\113\001\114\001\115\001\116\001\117\001\ -\118\001\119\001\255\255\255\255\255\255\255\255\255\255\125\001\ -\126\001\127\001\255\255\255\255\130\001\131\001\255\255\255\255\ -\134\001\135\001\136\001\137\001\138\001\255\255\255\255\141\001\ -\142\001\255\255\144\001\145\001\146\001\255\255\148\001\255\255\ -\150\001\255\255\255\255\255\255\154\001\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\162\001\163\001\255\255\165\001\ -\166\001\255\255\168\001\255\255\002\001\003\001\004\001\255\255\ -\006\001\255\255\255\255\255\255\255\255\255\255\255\255\013\001\ -\014\001\015\001\255\255\017\001\255\255\255\255\020\001\255\255\ -\255\255\255\255\255\255\025\001\255\255\027\001\028\001\255\255\ -\255\255\255\255\255\255\255\255\255\255\035\001\255\255\037\001\ -\038\001\255\255\255\255\041\001\042\001\255\255\255\255\255\255\ -\046\001\255\255\255\255\255\255\255\255\255\255\052\001\255\255\ -\054\001\055\001\056\001\057\001\058\001\059\001\255\255\255\255\ -\255\255\255\255\255\255\255\255\066\001\255\255\068\001\255\255\ -\255\255\255\255\255\255\255\255\074\001\075\001\255\255\255\255\ -\255\255\079\001\080\001\000\000\082\001\255\255\255\255\255\255\ -\255\255\255\255\255\255\089\001\255\255\091\001\092\001\093\001\ -\094\001\255\255\096\001\097\001\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\111\001\112\001\113\001\114\001\115\001\116\001\117\001\ -\118\001\119\001\255\255\255\255\255\255\255\255\255\255\125\001\ -\126\001\127\001\255\255\255\255\130\001\131\001\255\255\255\255\ -\134\001\135\001\136\001\137\001\138\001\255\255\255\255\141\001\ -\142\001\255\255\144\001\145\001\146\001\255\255\148\001\255\255\ -\150\001\255\255\255\255\255\255\154\001\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\162\001\163\001\255\255\165\001\ -\166\001\255\255\168\001\255\255\255\255\002\001\003\001\004\001\ -\255\255\006\001\255\255\255\255\255\255\255\255\255\255\255\255\ -\013\001\014\001\015\001\255\255\017\001\255\255\255\255\020\001\ -\255\255\255\255\255\255\255\255\255\255\255\255\027\001\028\001\ -\255\255\255\255\255\255\255\255\255\255\255\255\035\001\255\255\ -\037\001\038\001\255\255\255\255\041\001\042\001\255\255\255\255\ -\255\255\046\001\255\255\255\255\255\255\255\255\255\255\052\001\ -\255\255\054\001\055\001\056\001\057\001\058\001\059\001\255\255\ -\255\255\255\255\255\255\255\255\255\255\066\001\255\255\068\001\ -\255\255\255\255\255\255\255\255\255\255\074\001\075\001\255\255\ -\255\255\255\255\079\001\080\001\000\000\082\001\255\255\255\255\ -\255\255\255\255\255\255\255\255\089\001\255\255\091\001\092\001\ -\093\001\094\001\255\255\096\001\097\001\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\111\001\112\001\113\001\114\001\115\001\116\001\ -\117\001\118\001\119\001\255\255\255\255\255\255\255\255\255\255\ -\125\001\126\001\127\001\255\255\255\255\130\001\131\001\255\255\ -\255\255\134\001\135\001\136\001\137\001\138\001\255\255\255\255\ -\141\001\142\001\255\255\144\001\145\001\146\001\255\255\148\001\ -\255\255\150\001\255\255\255\255\255\255\154\001\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\162\001\163\001\255\255\ -\165\001\166\001\255\255\168\001\255\255\002\001\003\001\004\001\ -\255\255\006\001\255\255\255\255\255\255\255\255\255\255\255\255\ -\013\001\014\001\015\001\255\255\017\001\255\255\255\255\020\001\ -\255\255\255\255\255\255\255\255\255\255\255\255\027\001\028\001\ +\255\255\255\255\255\255\154\001\255\255\002\001\003\001\004\001\ +\255\255\006\001\255\255\162\001\163\001\255\255\165\001\166\001\ +\013\001\168\001\015\001\255\255\017\001\255\255\255\255\020\001\ +\255\255\255\255\255\255\255\255\255\255\255\255\027\001\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\035\001\255\255\ \037\001\038\001\255\255\255\255\041\001\042\001\255\255\255\255\ \255\255\046\001\255\255\255\255\255\255\255\255\255\255\052\001\ \255\255\054\001\055\001\056\001\057\001\058\001\059\001\255\255\ -\255\255\255\255\255\255\255\255\255\255\066\001\000\000\068\001\ +\255\255\255\255\255\255\255\255\000\000\066\001\255\255\068\001\ \255\255\255\255\255\255\255\255\255\255\074\001\075\001\255\255\ \255\255\255\255\079\001\080\001\255\255\082\001\255\255\255\255\ \255\255\255\255\255\255\255\255\089\001\255\255\091\001\092\001\ @@ -5037,19 +5089,38 @@ let yycheck = "\002\000\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ \255\255\255\255\111\001\112\001\113\001\114\001\115\001\116\001\ \117\001\118\001\119\001\255\255\255\255\255\255\255\255\255\255\ -\125\001\126\001\127\001\255\255\255\255\130\001\131\001\255\255\ +\255\255\126\001\127\001\255\255\255\255\130\001\131\001\255\255\ \255\255\134\001\135\001\136\001\137\001\138\001\255\255\255\255\ -\141\001\142\001\255\255\144\001\145\001\146\001\255\255\148\001\ +\141\001\255\255\255\255\144\001\145\001\146\001\255\255\148\001\ \255\255\150\001\255\255\255\255\255\255\154\001\255\255\255\255\ \255\255\255\255\255\255\255\255\255\255\162\001\163\001\255\255\ -\165\001\166\001\255\255\168\001\255\255\255\255\002\001\003\001\ -\004\001\255\255\006\001\255\255\255\255\255\255\255\255\255\255\ -\255\255\013\001\255\255\015\001\255\255\017\001\255\255\255\255\ +\165\001\166\001\255\255\168\001\002\001\003\001\004\001\255\255\ +\006\001\255\255\255\255\255\255\255\255\255\255\255\255\013\001\ +\255\255\015\001\255\255\017\001\255\255\255\255\020\001\255\255\ +\255\255\255\255\255\255\255\255\255\255\027\001\255\255\255\255\ +\255\255\255\255\255\255\255\255\255\255\035\001\255\255\037\001\ +\038\001\255\255\255\255\041\001\042\001\255\255\255\255\255\255\ +\046\001\255\255\255\255\255\255\255\255\255\255\052\001\255\255\ +\255\255\055\001\056\001\057\001\058\001\059\001\255\255\255\255\ +\255\255\255\255\255\255\000\000\066\001\255\255\068\001\255\255\ +\255\255\255\255\255\255\255\255\074\001\075\001\255\255\255\255\ +\255\255\079\001\080\001\255\255\082\001\255\255\255\255\255\255\ +\255\255\255\255\255\255\089\001\255\255\091\001\092\001\093\001\ +\094\001\255\255\096\001\097\001\255\255\255\255\255\255\255\255\ +\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ +\255\255\111\001\112\001\113\001\114\001\115\001\116\001\117\001\ +\118\001\119\001\255\255\255\255\255\255\255\255\255\255\255\255\ +\126\001\127\001\255\255\255\255\130\001\131\001\255\255\255\255\ +\134\001\135\001\136\001\137\001\138\001\255\255\255\255\141\001\ +\255\255\255\255\144\001\145\001\146\001\255\255\148\001\255\255\ +\150\001\255\255\255\255\255\255\154\001\255\255\002\001\003\001\ +\004\001\255\255\006\001\255\255\162\001\163\001\255\255\165\001\ +\166\001\013\001\168\001\015\001\255\255\017\001\255\255\255\255\ \020\001\255\255\255\255\255\255\255\255\255\255\255\255\027\001\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\035\001\ \255\255\037\001\038\001\255\255\255\255\041\001\042\001\255\255\ \255\255\255\255\046\001\255\255\255\255\255\255\255\255\255\255\ -\052\001\255\255\255\255\055\001\056\001\057\001\058\001\059\001\ +\255\255\255\255\255\255\055\001\056\001\057\001\058\001\059\001\ \255\255\255\255\255\255\255\255\255\255\000\000\066\001\255\255\ \068\001\255\255\255\255\255\255\255\255\255\255\074\001\075\001\ \255\255\255\255\255\255\079\001\080\001\255\255\082\001\255\255\ @@ -5062,29 +5133,29 @@ let yycheck = "\002\000\ \255\255\255\255\134\001\135\001\136\001\137\001\138\001\255\255\ \255\255\141\001\255\255\255\255\144\001\145\001\146\001\255\255\ \148\001\255\255\150\001\255\255\255\255\255\255\154\001\255\255\ -\002\001\003\001\004\001\255\255\006\001\255\255\162\001\163\001\ -\255\255\165\001\166\001\013\001\168\001\015\001\255\255\017\001\ -\255\255\255\255\020\001\255\255\255\255\255\255\255\255\255\255\ -\255\255\027\001\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\035\001\255\255\037\001\038\001\255\255\255\255\041\001\ -\042\001\255\255\255\255\255\255\046\001\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\055\001\056\001\057\001\ -\058\001\059\001\255\255\255\255\255\255\255\255\255\255\000\000\ -\066\001\255\255\068\001\255\255\255\255\255\255\255\255\255\255\ -\074\001\075\001\255\255\255\255\255\255\079\001\080\001\255\255\ -\082\001\255\255\255\255\255\255\255\255\255\255\255\255\089\001\ -\255\255\091\001\092\001\093\001\094\001\255\255\096\001\097\001\ +\255\255\255\255\255\255\255\255\255\255\255\255\162\001\163\001\ +\255\255\165\001\166\001\255\255\168\001\002\001\003\001\004\001\ +\255\255\006\001\255\255\255\255\255\255\255\255\255\255\255\255\ +\013\001\255\255\015\001\255\255\255\255\255\255\255\255\020\001\ +\255\255\255\255\255\255\255\255\255\255\255\255\027\001\255\255\ +\255\255\255\255\255\255\255\255\255\255\255\255\035\001\255\255\ +\037\001\038\001\255\255\255\255\255\255\042\001\255\255\255\255\ +\255\255\046\001\255\255\255\255\255\255\255\255\255\255\000\000\ +\255\255\255\255\055\001\056\001\057\001\058\001\059\001\255\255\ +\255\255\255\255\255\255\255\255\255\255\066\001\255\255\068\001\ +\255\255\255\255\255\255\255\255\255\255\074\001\075\001\255\255\ +\255\255\255\255\079\001\080\001\255\255\082\001\255\255\255\255\ +\255\255\255\255\255\255\255\255\089\001\255\255\091\001\092\001\ +\093\001\094\001\255\255\096\001\097\001\255\255\255\255\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\111\001\112\001\113\001\ -\114\001\115\001\116\001\117\001\118\001\119\001\255\255\255\255\ -\255\255\255\255\255\255\255\255\126\001\127\001\255\255\255\255\ -\130\001\131\001\255\255\255\255\134\001\135\001\136\001\137\001\ -\138\001\255\255\255\255\141\001\255\255\255\255\144\001\145\001\ -\146\001\255\255\148\001\255\255\150\001\255\255\255\255\255\255\ -\154\001\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\162\001\163\001\255\255\165\001\166\001\255\255\168\001\002\001\ -\003\001\004\001\255\255\006\001\255\255\255\255\255\255\255\255\ -\255\255\255\255\013\001\255\255\015\001\255\255\255\255\255\255\ +\255\255\255\255\111\001\112\001\113\001\114\001\115\001\116\001\ +\117\001\118\001\119\001\255\255\255\255\255\255\255\255\255\255\ +\255\255\126\001\127\001\255\255\255\255\130\001\131\001\255\255\ +\255\255\134\001\135\001\136\001\137\001\138\001\255\255\255\255\ +\141\001\255\255\255\255\144\001\255\255\146\001\255\255\148\001\ +\255\255\150\001\255\255\255\255\255\255\154\001\255\255\002\001\ +\003\001\004\001\255\255\006\001\255\255\162\001\163\001\255\255\ +\165\001\166\001\013\001\168\001\015\001\255\255\255\255\255\255\ \255\255\020\001\255\255\255\255\255\255\255\255\255\255\255\255\ \027\001\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ \035\001\255\255\037\001\038\001\255\255\255\255\255\255\042\001\ @@ -5165,8 +5236,8 @@ let yycheck = "\002\000\ \255\255\020\001\255\255\255\255\255\255\255\255\255\255\255\255\ \027\001\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ \035\001\255\255\037\001\038\001\255\255\255\255\255\255\042\001\ -\255\255\255\255\255\255\046\001\255\255\255\255\255\255\255\255\ -\255\255\000\000\255\255\255\255\055\001\056\001\057\001\058\001\ +\255\255\255\255\255\255\046\001\000\000\255\255\255\255\255\255\ +\255\255\255\255\255\255\255\255\055\001\056\001\057\001\058\001\ \059\001\255\255\255\255\255\255\255\255\255\255\255\255\066\001\ \255\255\068\001\255\255\255\255\255\255\255\255\255\255\074\001\ \075\001\255\255\255\255\255\255\079\001\080\001\255\255\082\001\ @@ -5197,29 +5268,29 @@ let yycheck = "\002\000\ \255\255\255\255\255\255\255\255\255\255\126\001\127\001\255\255\ \255\255\130\001\131\001\255\255\255\255\134\001\135\001\136\001\ \137\001\138\001\255\255\255\255\141\001\255\255\255\255\144\001\ -\255\255\146\001\255\255\148\001\255\255\150\001\255\255\255\255\ -\255\255\154\001\255\255\002\001\003\001\004\001\255\255\006\001\ -\255\255\162\001\163\001\255\255\165\001\166\001\013\001\168\001\ -\015\001\255\255\255\255\255\255\255\255\020\001\255\255\255\255\ -\255\255\255\255\255\255\255\255\027\001\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\035\001\255\255\037\001\038\001\ -\255\255\255\255\255\255\042\001\255\255\255\255\255\255\046\001\ -\000\000\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\055\001\056\001\057\001\058\001\059\001\255\255\255\255\255\255\ -\255\255\255\255\255\255\066\001\255\255\068\001\255\255\255\255\ -\255\255\255\255\255\255\074\001\075\001\255\255\255\255\255\255\ -\079\001\080\001\255\255\082\001\255\255\255\255\255\255\255\255\ -\255\255\255\255\089\001\255\255\091\001\092\001\093\001\094\001\ -\255\255\096\001\097\001\255\255\255\255\255\255\255\255\255\255\ +\255\255\146\001\255\255\148\001\255\255\150\001\002\001\003\001\ +\004\001\154\001\006\001\255\255\255\255\255\255\255\255\255\255\ +\255\255\162\001\163\001\015\001\165\001\166\001\255\255\168\001\ +\020\001\255\255\255\255\255\255\255\255\255\255\255\255\027\001\ +\255\255\255\255\255\255\255\255\255\255\255\255\255\255\035\001\ +\255\255\037\001\038\001\255\255\255\255\255\255\042\001\255\255\ +\255\255\255\255\046\001\000\000\255\255\255\255\255\255\255\255\ +\255\255\255\255\255\255\055\001\056\001\057\001\058\001\059\001\ +\255\255\255\255\255\255\255\255\255\255\255\255\066\001\255\255\ +\068\001\255\255\255\255\255\255\255\255\255\255\074\001\075\001\ +\255\255\255\255\255\255\079\001\080\001\255\255\082\001\255\255\ +\255\255\255\255\255\255\255\255\255\255\089\001\255\255\091\001\ +\092\001\093\001\094\001\255\255\096\001\097\001\255\255\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\111\001\112\001\113\001\114\001\115\001\116\001\117\001\118\001\ -\119\001\255\255\255\255\255\255\255\255\255\255\255\255\126\001\ -\127\001\255\255\255\255\130\001\131\001\255\255\255\255\134\001\ -\135\001\136\001\137\001\138\001\255\255\255\255\141\001\255\255\ -\255\255\144\001\255\255\146\001\255\255\148\001\255\255\150\001\ -\002\001\003\001\004\001\154\001\006\001\255\255\255\255\255\255\ -\255\255\255\255\255\255\162\001\163\001\015\001\165\001\166\001\ -\255\255\168\001\020\001\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\255\255\111\001\112\001\113\001\114\001\115\001\ +\116\001\117\001\118\001\119\001\255\255\255\255\255\255\255\255\ +\255\255\255\255\126\001\127\001\255\255\255\255\130\001\131\001\ +\255\255\255\255\134\001\135\001\136\001\137\001\138\001\255\255\ +\255\255\141\001\255\255\255\255\144\001\255\255\146\001\255\255\ +\148\001\255\255\150\001\255\255\255\255\255\255\154\001\255\255\ +\002\001\003\001\004\001\255\255\006\001\255\255\162\001\163\001\ +\255\255\165\001\166\001\255\255\168\001\015\001\255\255\255\255\ +\255\255\255\255\020\001\255\255\255\255\255\255\255\255\255\255\ \255\255\027\001\255\255\255\255\255\255\255\255\255\255\255\255\ \255\255\035\001\255\255\037\001\038\001\255\255\255\255\255\255\ \042\001\255\255\255\255\255\255\046\001\000\000\255\255\255\255\ @@ -5235,33 +5306,33 @@ let yycheck = "\002\000\ \255\255\255\255\255\255\255\255\126\001\127\001\255\255\255\255\ \130\001\131\001\255\255\255\255\134\001\135\001\136\001\137\001\ \138\001\255\255\255\255\141\001\255\255\255\255\144\001\255\255\ -\146\001\255\255\148\001\255\255\150\001\255\255\255\255\255\255\ -\154\001\255\255\002\001\003\001\004\001\255\255\006\001\255\255\ -\162\001\163\001\255\255\165\001\166\001\255\255\168\001\015\001\ -\255\255\255\255\255\255\255\255\020\001\255\255\255\255\255\255\ -\255\255\255\255\255\255\027\001\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\035\001\255\255\037\001\038\001\255\255\ -\255\255\255\255\042\001\255\255\255\255\255\255\046\001\000\000\ -\255\255\255\255\255\255\255\255\255\255\255\255\255\255\055\001\ -\056\001\057\001\058\001\059\001\255\255\255\255\255\255\255\255\ -\255\255\255\255\066\001\255\255\068\001\255\255\255\255\255\255\ -\255\255\255\255\074\001\075\001\255\255\255\255\255\255\079\001\ -\080\001\255\255\082\001\255\255\255\255\255\255\255\255\255\255\ -\255\255\089\001\255\255\091\001\092\001\093\001\094\001\255\255\ -\096\001\097\001\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\255\255\255\255\111\001\ -\112\001\113\001\114\001\115\001\116\001\117\001\118\001\119\001\ -\255\255\255\255\255\255\255\255\255\255\255\255\126\001\127\001\ -\255\255\255\255\130\001\131\001\255\255\255\255\134\001\135\001\ -\136\001\137\001\138\001\255\255\255\255\141\001\255\255\255\255\ -\144\001\255\255\146\001\255\255\148\001\255\255\150\001\002\001\ -\003\001\004\001\154\001\006\001\255\255\255\255\255\255\255\255\ -\255\255\255\255\162\001\163\001\015\001\165\001\166\001\255\255\ -\168\001\020\001\255\255\255\255\255\255\255\255\255\255\255\255\ +\146\001\255\255\148\001\255\255\150\001\002\001\003\001\004\001\ +\154\001\006\001\255\255\255\255\255\255\255\255\255\255\255\255\ +\162\001\163\001\015\001\165\001\166\001\255\255\168\001\020\001\ +\255\255\255\255\255\255\255\255\255\255\255\255\027\001\255\255\ +\255\255\255\255\255\255\255\255\255\255\255\255\035\001\255\255\ +\037\001\038\001\255\255\255\255\255\255\042\001\255\255\255\255\ +\255\255\046\001\255\255\255\255\255\255\255\255\255\255\000\000\ +\255\255\255\255\055\001\056\001\057\001\058\001\059\001\255\255\ +\255\255\255\255\255\255\255\255\255\255\066\001\255\255\068\001\ +\255\255\255\255\255\255\255\255\255\255\074\001\075\001\255\255\ +\255\255\255\255\079\001\080\001\255\255\082\001\255\255\255\255\ +\255\255\255\255\255\255\255\255\089\001\255\255\091\001\092\001\ +\093\001\094\001\255\255\096\001\097\001\255\255\255\255\255\255\ +\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\111\001\112\001\113\001\114\001\115\001\116\001\ +\117\001\118\001\119\001\255\255\255\255\255\255\255\255\255\255\ +\255\255\126\001\127\001\255\255\255\255\130\001\131\001\255\255\ +\255\255\134\001\135\001\136\001\137\001\138\001\255\255\255\255\ +\141\001\255\255\255\255\144\001\255\255\146\001\255\255\148\001\ +\255\255\150\001\255\255\255\255\255\255\154\001\255\255\002\001\ +\003\001\004\001\255\255\006\001\255\255\162\001\163\001\255\255\ +\165\001\166\001\013\001\168\001\015\001\255\255\255\255\255\255\ +\255\255\020\001\255\255\255\255\255\255\255\255\255\255\255\255\ \027\001\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\035\001\255\255\037\001\038\001\255\255\255\255\255\255\042\001\ -\255\255\255\255\255\255\046\001\255\255\255\255\255\255\255\255\ -\255\255\000\000\255\255\255\255\055\001\056\001\057\001\058\001\ +\035\001\255\255\255\255\038\001\255\255\255\255\255\255\042\001\ +\255\255\255\255\255\255\046\001\255\255\000\000\255\255\255\255\ +\255\255\255\255\255\255\255\255\055\001\056\001\057\001\058\001\ \059\001\255\255\255\255\255\255\255\255\255\255\255\255\066\001\ \255\255\068\001\255\255\255\255\255\255\255\255\255\255\074\001\ \075\001\255\255\255\255\255\255\079\001\080\001\255\255\082\001\ @@ -5277,9 +5348,9 @@ let yycheck = "\002\000\ \255\255\002\001\003\001\004\001\255\255\006\001\255\255\162\001\ \163\001\255\255\165\001\166\001\013\001\168\001\015\001\255\255\ \255\255\255\255\255\255\020\001\255\255\255\255\255\255\255\255\ -\255\255\255\255\027\001\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\035\001\255\255\255\255\038\001\255\255\255\255\ -\255\255\042\001\255\255\255\255\255\255\046\001\255\255\000\000\ +\255\255\255\255\027\001\255\255\255\255\255\255\000\000\255\255\ +\255\255\255\255\035\001\255\255\037\001\038\001\255\255\255\255\ +\255\255\042\001\255\255\255\255\255\255\046\001\255\255\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\055\001\056\001\ \057\001\058\001\059\001\255\255\255\255\255\255\255\255\255\255\ \255\255\066\001\255\255\068\001\255\255\255\255\255\255\255\255\ @@ -5291,127 +5362,105 @@ let yycheck = "\002\000\ \113\001\114\001\115\001\116\001\117\001\118\001\119\001\255\255\ \255\255\255\255\255\255\255\255\255\255\126\001\127\001\255\255\ \255\255\130\001\131\001\255\255\255\255\134\001\135\001\136\001\ -\137\001\138\001\255\255\255\255\141\001\255\255\255\255\144\001\ -\255\255\146\001\255\255\148\001\255\255\150\001\255\255\255\255\ -\255\255\154\001\255\255\002\001\003\001\004\001\255\255\006\001\ -\255\255\162\001\163\001\255\255\165\001\166\001\013\001\168\001\ -\015\001\255\255\255\255\255\255\255\255\020\001\255\255\255\255\ -\255\255\255\255\255\255\000\000\027\001\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\035\001\255\255\037\001\038\001\ -\255\255\255\255\255\255\042\001\255\255\255\255\255\255\046\001\ +\255\255\255\255\255\255\255\255\141\001\255\255\255\255\144\001\ +\255\255\146\001\255\255\148\001\255\255\150\001\255\255\002\001\ +\003\001\154\001\255\255\006\001\255\255\255\255\255\255\255\255\ +\255\255\162\001\163\001\014\001\165\001\166\001\017\001\168\001\ +\255\255\020\001\255\255\255\255\255\255\255\255\255\255\255\255\ +\027\001\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ +\035\001\255\255\255\255\255\255\000\000\040\001\255\255\042\001\ +\255\255\255\255\255\255\046\001\255\255\255\255\255\255\255\255\ +\255\255\255\255\255\255\255\255\055\001\056\001\057\001\058\001\ +\059\001\255\255\255\255\255\255\255\255\255\255\255\255\066\001\ +\255\255\068\001\255\255\255\255\255\255\255\255\255\255\074\001\ +\075\001\255\255\255\255\255\255\255\255\080\001\255\255\082\001\ +\255\255\255\255\255\255\255\255\255\255\255\255\089\001\255\255\ +\091\001\092\001\093\001\094\001\255\255\096\001\097\001\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\055\001\056\001\057\001\058\001\059\001\255\255\255\255\255\255\ -\255\255\255\255\255\255\066\001\255\255\068\001\255\255\255\255\ -\255\255\255\255\255\255\074\001\075\001\255\255\255\255\255\255\ -\079\001\080\001\255\255\082\001\255\255\255\255\255\255\255\255\ -\255\255\255\255\089\001\255\255\091\001\092\001\093\001\094\001\ -\255\255\096\001\097\001\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\255\255\255\255\111\001\112\001\113\001\114\001\ +\115\001\116\001\117\001\118\001\119\001\255\255\255\255\255\255\ +\255\255\255\255\125\001\126\001\127\001\255\255\255\255\130\001\ +\131\001\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ +\002\001\255\255\141\001\255\255\006\001\255\255\255\255\146\001\ +\255\255\255\255\255\255\150\001\255\255\255\255\255\255\154\001\ +\255\255\255\255\020\001\021\001\255\255\255\255\255\255\162\001\ +\163\001\027\001\165\001\166\001\030\001\168\001\255\255\255\255\ +\255\255\035\001\255\255\255\255\255\255\255\255\040\001\255\255\ +\042\001\255\255\255\255\255\255\046\001\255\255\255\255\255\255\ +\255\255\255\255\000\000\255\255\255\255\255\255\056\001\057\001\ +\058\001\059\001\255\255\255\255\255\255\255\255\255\255\255\255\ +\066\001\255\255\068\001\255\255\255\255\255\255\255\255\255\255\ +\074\001\075\001\255\255\255\255\255\255\255\255\080\001\255\255\ +\082\001\255\255\255\255\255\255\255\255\255\255\255\255\089\001\ +\255\255\091\001\092\001\093\001\094\001\255\255\096\001\097\001\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\111\001\112\001\113\001\114\001\115\001\116\001\117\001\118\001\ -\119\001\255\255\255\255\255\255\255\255\255\255\255\255\126\001\ -\127\001\255\255\255\255\130\001\131\001\255\255\255\255\134\001\ -\135\001\136\001\255\255\255\255\255\255\255\255\141\001\255\255\ -\255\255\144\001\255\255\146\001\255\255\148\001\255\255\150\001\ -\255\255\002\001\003\001\154\001\255\255\006\001\255\255\255\255\ -\255\255\255\255\255\255\162\001\163\001\014\001\165\001\166\001\ -\017\001\168\001\255\255\020\001\255\255\255\255\255\255\255\255\ -\000\000\255\255\027\001\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\035\001\255\255\255\255\255\255\255\255\040\001\ -\255\255\042\001\255\255\255\255\255\255\046\001\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\255\255\055\001\056\001\ -\057\001\058\001\059\001\255\255\255\255\255\255\255\255\255\255\ -\255\255\066\001\255\255\068\001\255\255\255\255\255\255\255\255\ -\255\255\074\001\075\001\255\255\255\255\255\255\255\255\080\001\ -\255\255\082\001\255\255\255\255\255\255\255\255\255\255\255\255\ -\089\001\255\255\091\001\092\001\093\001\094\001\255\255\096\001\ -\097\001\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\255\255\111\001\112\001\ -\113\001\114\001\115\001\116\001\117\001\118\001\119\001\255\255\ -\255\255\255\255\255\255\255\255\125\001\126\001\127\001\255\255\ -\255\255\130\001\131\001\255\255\255\255\002\001\003\001\255\255\ -\255\255\006\001\255\255\255\255\141\001\255\255\255\255\255\255\ -\255\255\146\001\255\255\255\255\017\001\150\001\255\255\020\001\ -\255\255\154\001\255\255\255\255\255\255\255\255\027\001\255\255\ -\255\255\162\001\163\001\255\255\165\001\166\001\035\001\168\001\ -\255\255\255\255\255\255\040\001\255\255\042\001\255\255\255\255\ -\255\255\046\001\255\255\255\255\255\255\255\255\000\000\255\255\ -\255\255\255\255\055\001\056\001\057\001\058\001\059\001\255\255\ -\255\255\255\255\255\255\255\255\255\255\066\001\255\255\068\001\ -\255\255\255\255\255\255\255\255\255\255\074\001\075\001\255\255\ -\255\255\255\255\255\255\080\001\255\255\082\001\255\255\255\255\ -\255\255\255\255\255\255\255\255\089\001\255\255\091\001\092\001\ -\093\001\094\001\255\255\096\001\097\001\255\255\255\255\255\255\ +\255\255\255\255\255\255\255\255\255\255\111\001\112\001\113\001\ +\114\001\115\001\116\001\117\001\118\001\119\001\255\255\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\111\001\112\001\113\001\114\001\115\001\116\001\ -\117\001\118\001\119\001\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\130\001\131\001\255\255\ +\130\001\131\001\255\255\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\255\255\141\001\255\255\255\255\255\255\255\255\ +\146\001\255\255\255\255\255\255\150\001\255\255\255\255\255\255\ +\154\001\255\255\255\255\255\255\255\255\255\255\255\255\003\001\ +\162\001\163\001\006\001\165\001\166\001\255\255\168\001\011\001\ +\000\000\013\001\255\255\255\255\255\255\255\255\255\255\019\001\ +\020\001\021\001\022\001\255\255\024\001\255\255\255\255\027\001\ +\255\255\029\001\255\255\255\255\255\255\255\255\255\255\035\001\ +\255\255\255\255\255\255\255\255\040\001\255\255\042\001\255\255\ +\044\001\255\255\046\001\255\255\255\255\049\001\050\001\255\255\ +\255\255\255\255\255\255\055\001\056\001\057\001\058\001\059\001\ +\060\001\061\001\062\001\063\001\064\001\255\255\066\001\255\255\ +\068\001\069\001\070\001\255\255\255\255\073\001\074\001\075\001\ +\076\001\255\255\078\001\079\001\080\001\255\255\082\001\255\255\ +\084\001\085\001\255\255\255\255\255\255\089\001\090\001\091\001\ +\092\001\093\001\094\001\255\255\096\001\097\001\255\255\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\141\001\255\255\255\255\255\255\255\255\146\001\255\255\255\255\ -\255\255\150\001\002\001\255\255\255\255\154\001\006\001\255\255\ -\255\255\255\255\255\255\255\255\255\255\162\001\163\001\255\255\ -\165\001\166\001\255\255\168\001\020\001\021\001\255\255\255\255\ -\255\255\255\255\255\255\027\001\255\255\255\255\030\001\255\255\ -\255\255\255\255\255\255\035\001\255\255\255\255\255\255\255\255\ -\040\001\255\255\042\001\255\255\255\255\255\255\046\001\255\255\ +\255\255\255\255\255\255\111\001\112\001\113\001\114\001\115\001\ +\116\001\117\001\118\001\119\001\255\255\255\255\255\255\255\255\ +\124\001\125\001\126\001\127\001\128\001\255\255\130\001\131\001\ +\132\001\255\255\000\000\255\255\136\001\137\001\255\255\139\001\ +\140\001\141\001\142\001\143\001\255\255\145\001\146\001\255\255\ +\255\255\255\255\150\001\151\001\002\001\153\001\154\001\255\255\ +\006\001\157\001\158\001\159\001\160\001\161\001\162\001\163\001\ +\255\255\165\001\166\001\167\001\255\255\255\255\020\001\021\001\ +\255\255\255\255\255\255\255\255\255\255\027\001\255\255\255\255\ +\255\255\255\255\255\255\255\255\255\255\035\001\255\255\255\255\ +\255\255\255\255\040\001\255\255\042\001\255\255\255\255\255\255\ +\046\001\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\056\001\057\001\058\001\059\001\255\255\255\255\ +\255\255\255\255\255\255\255\255\066\001\255\255\068\001\255\255\ +\255\255\255\255\255\255\255\255\074\001\075\001\255\255\255\255\ +\255\255\255\255\080\001\255\255\082\001\255\255\255\255\255\255\ +\255\255\255\255\255\255\089\001\255\255\091\001\092\001\093\001\ +\094\001\255\255\096\001\097\001\255\255\255\255\255\255\255\255\ \255\255\255\255\255\255\255\255\000\000\255\255\255\255\255\255\ +\255\255\111\001\112\001\113\001\114\001\115\001\116\001\117\001\ +\118\001\119\001\002\001\255\255\255\255\255\255\006\001\255\255\ +\255\255\255\255\255\255\255\255\130\001\131\001\255\255\255\255\ +\255\255\255\255\255\255\255\255\020\001\255\255\255\255\141\001\ +\255\255\255\255\255\255\027\001\146\001\255\255\255\255\255\255\ +\150\001\255\255\255\255\035\001\154\001\255\255\255\255\255\255\ +\040\001\255\255\042\001\255\255\162\001\163\001\046\001\165\001\ +\166\001\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ \056\001\057\001\058\001\059\001\255\255\255\255\255\255\255\255\ \255\255\255\255\066\001\255\255\068\001\255\255\255\255\255\255\ \255\255\255\255\074\001\075\001\255\255\255\255\255\255\255\255\ \080\001\255\255\082\001\255\255\255\255\255\255\255\255\255\255\ -\255\255\089\001\255\255\091\001\092\001\093\001\094\001\255\255\ +\255\255\089\001\000\000\091\001\092\001\093\001\094\001\255\255\ \096\001\097\001\255\255\255\255\255\255\255\255\255\255\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\111\001\ \112\001\113\001\114\001\115\001\116\001\117\001\118\001\119\001\ -\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\130\001\131\001\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\141\001\255\255\255\255\ -\255\255\255\255\146\001\255\255\255\255\255\255\150\001\255\255\ -\255\255\255\255\154\001\255\255\255\255\255\255\255\255\255\255\ -\255\255\003\001\162\001\163\001\006\001\165\001\166\001\255\255\ -\168\001\011\001\000\000\013\001\255\255\255\255\255\255\255\255\ -\255\255\019\001\020\001\021\001\022\001\255\255\024\001\255\255\ -\255\255\027\001\255\255\029\001\255\255\255\255\255\255\255\255\ -\255\255\035\001\255\255\255\255\255\255\255\255\040\001\255\255\ -\042\001\255\255\044\001\255\255\046\001\255\255\255\255\049\001\ -\050\001\255\255\255\255\255\255\255\255\055\001\056\001\057\001\ -\058\001\059\001\060\001\061\001\062\001\063\001\064\001\255\255\ -\066\001\255\255\068\001\069\001\070\001\255\255\255\255\073\001\ -\074\001\075\001\076\001\255\255\078\001\079\001\080\001\255\255\ -\082\001\255\255\084\001\085\001\255\255\255\255\255\255\089\001\ -\090\001\091\001\092\001\093\001\094\001\255\255\096\001\097\001\ -\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\111\001\112\001\113\001\ -\114\001\115\001\116\001\117\001\118\001\119\001\255\255\255\255\ -\255\255\255\255\124\001\125\001\126\001\127\001\128\001\255\255\ -\130\001\131\001\132\001\255\255\000\000\255\255\136\001\137\001\ -\255\255\139\001\140\001\141\001\142\001\143\001\255\255\145\001\ -\146\001\255\255\255\255\255\255\150\001\151\001\002\001\153\001\ -\154\001\255\255\006\001\157\001\158\001\159\001\160\001\161\001\ -\162\001\163\001\255\255\165\001\166\001\167\001\255\255\255\255\ -\020\001\021\001\255\255\255\255\255\255\255\255\255\255\027\001\ -\255\255\255\255\255\255\255\255\255\255\255\255\255\255\035\001\ -\255\255\255\255\255\255\255\255\040\001\255\255\042\001\255\255\ -\255\255\255\255\046\001\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\056\001\057\001\058\001\059\001\ -\255\255\255\255\255\255\255\255\255\255\255\255\066\001\255\255\ -\068\001\255\255\255\255\255\255\255\255\255\255\074\001\075\001\ -\255\255\255\255\255\255\255\255\080\001\255\255\082\001\255\255\ -\255\255\255\255\255\255\255\255\255\255\089\001\255\255\091\001\ -\092\001\093\001\094\001\255\255\096\001\097\001\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\255\255\000\000\255\255\ -\255\255\255\255\255\255\111\001\112\001\113\001\114\001\115\001\ -\116\001\117\001\118\001\119\001\002\001\255\255\255\255\255\255\ -\006\001\255\255\255\255\255\255\255\255\255\255\130\001\131\001\ -\255\255\255\255\255\255\255\255\255\255\255\255\020\001\255\255\ -\255\255\141\001\255\255\255\255\255\255\027\001\146\001\255\255\ -\255\255\255\255\150\001\255\255\255\255\035\001\154\001\255\255\ -\255\255\255\255\040\001\255\255\042\001\255\255\162\001\163\001\ -\046\001\165\001\166\001\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\255\255\255\255\002\001\255\255\255\255\255\255\ +\006\001\255\255\130\001\131\001\255\255\255\255\255\255\013\001\ +\255\255\255\255\255\255\255\255\255\255\141\001\020\001\255\255\ +\255\255\255\255\146\001\255\255\255\255\027\001\150\001\255\255\ +\255\255\255\255\154\001\255\255\255\255\035\001\255\255\255\255\ +\255\255\255\255\162\001\163\001\042\001\165\001\166\001\255\255\ +\046\001\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ \255\255\255\255\056\001\057\001\058\001\059\001\255\255\255\255\ -\255\255\255\255\255\255\255\255\066\001\255\255\068\001\255\255\ +\255\255\255\255\255\255\255\255\066\001\255\255\068\001\000\000\ \255\255\255\255\255\255\255\255\074\001\075\001\255\255\255\255\ \255\255\255\255\080\001\255\255\082\001\255\255\255\255\255\255\ -\255\255\255\255\255\255\089\001\000\000\091\001\092\001\093\001\ +\255\255\255\255\255\255\089\001\255\255\091\001\092\001\093\001\ \094\001\255\255\096\001\097\001\255\255\255\255\255\255\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ \255\255\111\001\112\001\113\001\114\001\115\001\116\001\117\001\ @@ -5421,42 +5470,52 @@ let yycheck = "\002\000\ \020\001\255\255\255\255\255\255\146\001\255\255\255\255\027\001\ \150\001\255\255\255\255\255\255\154\001\255\255\255\255\035\001\ \255\255\255\255\255\255\255\255\162\001\163\001\042\001\165\001\ -\166\001\255\255\046\001\255\255\255\255\255\255\255\255\255\255\ +\166\001\255\255\046\001\255\255\000\000\255\255\255\255\255\255\ \255\255\255\255\255\255\255\255\056\001\057\001\058\001\059\001\ \255\255\255\255\255\255\255\255\255\255\255\255\066\001\255\255\ -\068\001\000\000\255\255\255\255\255\255\255\255\074\001\075\001\ +\068\001\255\255\255\255\255\255\255\255\255\255\074\001\075\001\ \255\255\255\255\255\255\255\255\080\001\255\255\082\001\255\255\ \255\255\255\255\255\255\255\255\255\255\089\001\255\255\091\001\ \092\001\093\001\094\001\255\255\096\001\097\001\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\111\001\112\001\113\001\114\001\115\001\ -\116\001\117\001\118\001\119\001\255\255\255\255\255\255\255\255\ -\002\001\255\255\255\255\255\255\006\001\255\255\130\001\131\001\ -\255\255\255\255\255\255\013\001\255\255\255\255\255\255\255\255\ -\255\255\141\001\020\001\255\255\255\255\255\255\146\001\255\255\ -\255\255\027\001\150\001\255\255\255\255\255\255\154\001\255\255\ -\255\255\035\001\255\255\255\255\255\255\255\255\162\001\163\001\ -\042\001\165\001\166\001\255\255\046\001\255\255\000\000\255\255\ -\255\255\255\255\255\255\255\255\255\255\255\255\056\001\057\001\ -\058\001\059\001\255\255\255\255\255\255\255\255\255\255\255\255\ -\066\001\255\255\068\001\255\255\255\255\255\255\255\255\255\255\ -\074\001\075\001\255\255\255\255\255\255\255\255\080\001\255\255\ -\082\001\255\255\255\255\255\255\255\255\255\255\255\255\089\001\ -\255\255\091\001\092\001\093\001\094\001\255\255\096\001\097\001\ -\255\255\255\255\255\255\255\255\255\255\255\255\002\001\255\255\ -\255\255\255\255\006\001\255\255\255\255\111\001\112\001\113\001\ -\114\001\115\001\116\001\117\001\118\001\119\001\255\255\255\255\ +\255\255\255\255\255\255\255\255\002\001\255\255\255\255\255\255\ +\006\001\255\255\255\255\111\001\112\001\113\001\114\001\115\001\ +\116\001\117\001\118\001\119\001\255\255\255\255\020\001\255\255\ +\255\255\255\255\255\255\255\255\255\255\027\001\130\001\131\001\ +\255\255\255\255\255\255\255\255\255\255\035\001\255\255\255\255\ +\255\255\141\001\255\255\255\255\042\001\255\255\146\001\255\255\ +\046\001\000\000\150\001\255\255\255\255\255\255\154\001\255\255\ +\255\255\255\255\056\001\057\001\058\001\059\001\162\001\163\001\ +\255\255\165\001\166\001\255\255\066\001\255\255\068\001\255\255\ +\255\255\255\255\255\255\255\255\074\001\075\001\255\255\255\255\ +\255\255\255\255\080\001\255\255\082\001\255\255\255\255\255\255\ +\255\255\255\255\255\255\089\001\255\255\091\001\092\001\093\001\ +\094\001\255\255\096\001\097\001\255\255\255\255\255\255\255\255\ +\255\255\002\001\255\255\255\255\255\255\006\001\255\255\255\255\ +\255\255\111\001\112\001\113\001\114\001\115\001\116\001\117\001\ +\118\001\119\001\255\255\020\001\255\255\255\255\255\255\255\255\ +\255\255\255\255\027\001\255\255\130\001\131\001\255\255\255\255\ +\255\255\255\255\035\001\255\255\255\255\255\255\255\255\141\001\ +\255\255\042\001\255\255\255\255\146\001\046\001\000\000\255\255\ +\150\001\255\255\255\255\255\255\154\001\255\255\255\255\056\001\ +\057\001\058\001\059\001\255\255\162\001\163\001\255\255\165\001\ +\166\001\066\001\255\255\068\001\255\255\255\255\255\255\255\255\ +\255\255\074\001\075\001\255\255\255\255\255\255\255\255\080\001\ +\255\255\082\001\255\255\255\255\255\255\255\255\255\255\255\255\ +\089\001\255\255\091\001\092\001\093\001\094\001\255\255\096\001\ +\097\001\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\006\001\255\255\255\255\255\255\111\001\112\001\ +\113\001\114\001\115\001\116\001\117\001\118\001\119\001\255\255\ \020\001\255\255\255\255\255\255\255\255\255\255\255\255\027\001\ -\130\001\131\001\255\255\255\255\255\255\255\255\255\255\035\001\ -\255\255\255\255\255\255\141\001\255\255\255\255\042\001\255\255\ -\146\001\255\255\046\001\000\000\150\001\255\255\255\255\255\255\ -\154\001\255\255\255\255\255\255\056\001\057\001\058\001\059\001\ -\162\001\163\001\255\255\165\001\166\001\255\255\066\001\255\255\ +\255\255\130\001\131\001\255\255\255\255\255\255\255\255\035\001\ +\255\255\255\255\255\255\255\255\141\001\255\255\042\001\255\255\ +\255\255\146\001\046\001\000\000\255\255\150\001\255\255\255\255\ +\255\255\154\001\255\255\255\255\056\001\057\001\058\001\059\001\ +\255\255\162\001\163\001\255\255\165\001\166\001\066\001\255\255\ \068\001\255\255\255\255\255\255\255\255\255\255\074\001\075\001\ \255\255\255\255\255\255\255\255\080\001\255\255\082\001\255\255\ \255\255\255\255\255\255\255\255\255\255\089\001\255\255\091\001\ \092\001\093\001\094\001\255\255\096\001\097\001\255\255\255\255\ -\255\255\255\255\255\255\002\001\255\255\255\255\255\255\006\001\ +\255\255\255\255\255\255\255\255\255\255\255\255\255\255\006\001\ \255\255\255\255\255\255\111\001\112\001\113\001\114\001\115\001\ \116\001\117\001\118\001\119\001\255\255\020\001\255\255\255\255\ \255\255\255\255\255\255\255\255\027\001\255\255\130\001\131\001\ @@ -5474,7 +5533,7 @@ let yycheck = "\002\000\ \119\001\255\255\020\001\255\255\255\255\255\255\255\255\255\255\ \255\255\027\001\255\255\130\001\131\001\255\255\255\255\255\255\ \255\255\035\001\255\255\255\255\255\255\255\255\141\001\255\255\ -\042\001\255\255\255\255\146\001\046\001\000\000\255\255\150\001\ +\042\001\255\255\255\255\146\001\046\001\255\255\255\255\150\001\ \255\255\255\255\255\255\154\001\255\255\255\255\056\001\057\001\ \058\001\059\001\255\255\162\001\163\001\255\255\165\001\166\001\ \066\001\255\255\068\001\255\255\255\255\255\255\255\255\255\255\ @@ -5487,7 +5546,7 @@ let yycheck = "\002\000\ \255\255\255\255\255\255\255\255\255\255\255\255\027\001\255\255\ \130\001\131\001\255\255\255\255\255\255\255\255\035\001\255\255\ \255\255\255\255\255\255\141\001\255\255\042\001\255\255\255\255\ -\146\001\046\001\000\000\255\255\150\001\255\255\255\255\255\255\ +\146\001\046\001\255\255\255\255\150\001\255\255\255\255\255\255\ \154\001\255\255\255\255\056\001\057\001\058\001\059\001\255\255\ \162\001\163\001\255\255\165\001\166\001\066\001\255\255\068\001\ \255\255\255\255\255\255\255\255\255\255\074\001\075\001\255\255\ @@ -5507,52 +5566,34 @@ let yycheck = "\002\000\ \080\001\255\255\082\001\255\255\255\255\255\255\255\255\255\255\ \255\255\089\001\255\255\091\001\092\001\093\001\094\001\255\255\ \096\001\097\001\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\006\001\255\255\255\255\255\255\111\001\ +\255\255\255\255\255\255\255\255\255\255\255\255\255\255\111\001\ \112\001\113\001\114\001\115\001\116\001\117\001\118\001\119\001\ -\255\255\020\001\255\255\255\255\255\255\255\255\255\255\255\255\ -\027\001\255\255\130\001\131\001\255\255\255\255\255\255\255\255\ -\035\001\255\255\255\255\255\255\255\255\141\001\255\255\042\001\ -\255\255\255\255\146\001\046\001\255\255\255\255\150\001\255\255\ -\255\255\255\255\154\001\255\255\255\255\056\001\057\001\058\001\ -\059\001\255\255\162\001\163\001\255\255\165\001\166\001\066\001\ -\255\255\068\001\255\255\255\255\255\255\255\255\255\255\074\001\ -\075\001\255\255\255\255\255\255\255\255\080\001\255\255\082\001\ -\255\255\255\255\255\255\255\255\255\255\255\255\089\001\255\255\ -\091\001\092\001\093\001\094\001\255\255\096\001\097\001\255\255\ -\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\006\001\255\255\255\255\255\255\111\001\112\001\113\001\114\001\ -\115\001\116\001\117\001\118\001\119\001\255\255\020\001\255\255\ -\255\255\255\255\255\255\255\255\255\255\027\001\255\255\130\001\ -\131\001\255\255\255\255\255\255\255\255\035\001\255\255\255\255\ -\255\255\255\255\141\001\255\255\042\001\255\255\255\255\146\001\ -\046\001\255\255\255\255\150\001\255\255\255\255\255\255\154\001\ -\255\255\255\255\056\001\057\001\058\001\059\001\255\255\162\001\ -\163\001\255\255\165\001\166\001\066\001\255\255\068\001\255\255\ -\255\255\255\255\255\255\255\255\074\001\075\001\255\255\255\255\ -\255\255\255\255\080\001\255\255\082\001\255\255\255\255\255\255\ -\255\255\255\255\255\255\089\001\255\255\091\001\092\001\093\001\ -\094\001\255\255\096\001\097\001\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\006\001\255\255\255\255\ -\255\255\111\001\112\001\113\001\114\001\115\001\116\001\117\001\ -\118\001\119\001\255\255\020\001\255\255\255\255\255\255\255\255\ -\255\255\255\255\027\001\255\255\130\001\131\001\255\255\255\255\ -\255\255\255\255\035\001\255\255\255\255\255\255\255\255\141\001\ -\255\255\042\001\255\255\255\255\146\001\046\001\255\255\255\255\ -\150\001\255\255\255\255\255\255\154\001\255\255\255\255\056\001\ -\057\001\058\001\059\001\255\255\162\001\163\001\255\255\165\001\ -\166\001\066\001\255\255\068\001\255\255\255\255\255\255\255\255\ -\255\255\074\001\075\001\255\255\255\255\255\255\255\255\080\001\ -\255\255\082\001\255\255\255\255\255\255\255\255\255\255\255\255\ -\089\001\255\255\091\001\092\001\093\001\094\001\255\255\096\001\ -\097\001\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\255\255\111\001\112\001\ -\113\001\114\001\115\001\116\001\117\001\118\001\119\001\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\130\001\131\001\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\141\001\255\255\255\255\255\255\ -\255\255\146\001\255\255\255\255\255\255\150\001\255\255\255\255\ -\255\255\154\001\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\162\001\163\001\255\255\165\001\166\001\003\001\255\255\ +\255\255\255\255\130\001\131\001\255\255\255\255\255\255\255\255\ +\255\255\255\255\255\255\255\255\255\255\141\001\255\255\255\255\ +\255\255\255\255\146\001\255\255\255\255\255\255\150\001\255\255\ +\255\255\255\255\154\001\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\162\001\163\001\255\255\165\001\166\001\003\001\ +\255\255\005\001\006\001\007\001\008\001\009\001\010\001\011\001\ +\012\001\255\255\255\255\255\255\016\001\255\255\018\001\019\001\ +\255\255\255\255\022\001\023\001\255\255\025\001\026\001\255\255\ +\028\001\029\001\255\255\255\255\255\255\255\255\255\255\255\255\ +\036\001\255\255\255\255\039\001\040\001\255\255\255\255\043\001\ +\044\001\045\001\255\255\047\001\048\001\049\001\050\001\051\001\ +\052\001\053\001\054\001\255\255\255\255\255\255\255\255\255\255\ +\060\001\061\001\062\001\063\001\064\001\065\001\255\255\255\255\ +\255\255\069\001\070\001\255\255\255\255\073\001\074\001\075\001\ +\076\001\077\001\078\001\255\255\080\001\081\001\255\255\255\255\ +\084\001\085\001\086\001\087\001\088\001\255\255\090\001\255\255\ +\255\255\255\255\255\255\255\255\255\255\255\255\098\001\099\001\ +\100\001\101\001\102\001\103\001\104\001\105\001\106\001\107\001\ +\108\001\109\001\110\001\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\255\255\255\255\255\255\255\255\122\001\255\255\ +\124\001\255\255\255\255\255\255\128\001\255\255\255\255\255\255\ +\132\001\133\001\255\255\255\255\255\255\255\255\255\255\139\001\ +\140\001\255\255\255\255\143\001\255\255\255\255\255\255\147\001\ +\255\255\149\001\255\255\151\001\152\001\153\001\255\255\255\255\ +\255\255\157\001\158\001\159\001\160\001\161\001\003\001\255\255\ \005\001\006\001\007\001\008\001\009\001\010\001\011\001\012\001\ \255\255\255\255\255\255\016\001\255\255\018\001\019\001\255\255\ \255\255\022\001\023\001\255\255\025\001\026\001\255\255\028\001\ @@ -5590,28 +5631,26 @@ let yycheck = "\002\000\ \255\255\255\255\255\255\255\255\122\001\255\255\124\001\255\255\ \255\255\255\255\128\001\255\255\255\255\255\255\132\001\133\001\ \255\255\255\255\255\255\255\255\255\255\139\001\140\001\255\255\ -\255\255\143\001\255\255\255\255\255\255\147\001\255\255\149\001\ -\255\255\151\001\152\001\153\001\255\255\255\255\255\255\157\001\ -\158\001\159\001\160\001\161\001\003\001\255\255\005\001\006\001\ -\007\001\008\001\009\001\010\001\011\001\012\001\255\255\255\255\ -\255\255\016\001\255\255\018\001\019\001\255\255\255\255\022\001\ -\023\001\255\255\025\001\026\001\255\255\028\001\029\001\255\255\ -\255\255\255\255\255\255\255\255\255\255\036\001\255\255\255\255\ -\039\001\040\001\255\255\255\255\043\001\044\001\045\001\255\255\ -\047\001\048\001\049\001\050\001\051\001\052\001\053\001\054\001\ +\255\255\143\001\255\255\255\255\003\001\147\001\255\255\149\001\ +\255\255\151\001\152\001\153\001\011\001\255\255\255\255\157\001\ +\158\001\159\001\160\001\161\001\019\001\255\255\255\255\022\001\ +\023\001\255\255\025\001\255\255\255\255\028\001\029\001\255\255\ +\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ +\255\255\040\001\255\255\255\255\255\255\044\001\255\255\255\255\ +\255\255\255\255\049\001\050\001\255\255\052\001\255\255\054\001\ \255\255\255\255\255\255\255\255\255\255\060\001\061\001\062\001\ -\063\001\064\001\065\001\255\255\255\255\255\255\069\001\070\001\ -\255\255\255\255\073\001\074\001\075\001\076\001\077\001\078\001\ -\255\255\080\001\081\001\255\255\255\255\084\001\085\001\086\001\ -\087\001\088\001\255\255\090\001\255\255\255\255\255\255\255\255\ +\063\001\064\001\255\255\255\255\255\255\255\255\069\001\255\255\ +\255\255\255\255\073\001\255\255\255\255\076\001\255\255\078\001\ +\255\255\255\255\081\001\255\255\255\255\084\001\085\001\255\255\ +\255\255\255\255\255\255\090\001\255\255\255\255\255\255\255\255\ \255\255\255\255\255\255\098\001\099\001\100\001\101\001\102\001\ -\103\001\104\001\105\001\106\001\107\001\108\001\109\001\110\001\ +\103\001\104\001\105\001\106\001\107\001\108\001\255\255\110\001\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\122\001\255\255\124\001\255\255\255\255\ -\255\255\128\001\255\255\255\255\255\255\132\001\133\001\255\255\ +\255\255\255\255\255\255\255\255\255\255\124\001\255\255\255\255\ +\255\255\128\001\255\255\255\255\255\255\132\001\255\255\255\255\ \255\255\255\255\255\255\255\255\139\001\140\001\255\255\255\255\ -\143\001\255\255\255\255\003\001\147\001\255\255\149\001\255\255\ -\151\001\152\001\153\001\011\001\255\255\255\255\157\001\158\001\ +\143\001\255\255\255\255\003\001\255\255\255\255\149\001\255\255\ +\151\001\255\255\153\001\011\001\255\255\255\255\157\001\158\001\ \159\001\160\001\161\001\019\001\255\255\255\255\022\001\023\001\ \255\255\025\001\255\255\255\255\028\001\029\001\255\255\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ @@ -5628,311 +5667,340 @@ let yycheck = "\002\000\ \255\255\255\255\255\255\255\255\124\001\255\255\255\255\255\255\ \128\001\255\255\255\255\255\255\132\001\255\255\255\255\255\255\ \255\255\255\255\255\255\139\001\140\001\255\255\255\255\143\001\ -\255\255\255\255\003\001\255\255\255\255\149\001\255\255\151\001\ -\255\255\153\001\011\001\255\255\255\255\157\001\158\001\159\001\ -\160\001\161\001\019\001\255\255\255\255\022\001\023\001\255\255\ -\025\001\255\255\255\255\028\001\029\001\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\255\255\255\255\040\001\ -\255\255\255\255\255\255\044\001\255\255\255\255\255\255\255\255\ -\049\001\050\001\255\255\052\001\255\255\054\001\255\255\255\255\ -\255\255\255\255\255\255\060\001\061\001\062\001\063\001\064\001\ -\255\255\255\255\255\255\255\255\069\001\255\255\255\255\255\255\ -\073\001\255\255\255\255\076\001\255\255\078\001\255\255\255\255\ -\081\001\255\255\255\255\084\001\085\001\255\255\255\255\255\255\ -\255\255\090\001\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\098\001\099\001\100\001\101\001\102\001\103\001\104\001\ -\105\001\106\001\107\001\108\001\255\255\110\001\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\124\001\255\255\255\255\255\255\128\001\ -\255\255\255\255\255\255\132\001\255\255\255\255\255\255\255\255\ -\255\255\255\255\139\001\140\001\255\255\255\255\143\001\255\255\ -\255\255\255\255\255\255\255\255\149\001\255\255\151\001\255\255\ -\153\001\255\255\255\255\255\255\157\001\158\001\159\001\160\001\ -\161\001\005\001\006\001\007\001\008\001\009\001\010\001\011\001\ -\012\001\255\255\255\255\255\255\016\001\255\255\018\001\019\001\ -\255\255\255\255\255\255\255\255\255\255\255\255\026\001\255\255\ -\255\255\029\001\255\255\255\255\255\255\255\255\255\255\255\255\ -\036\001\255\255\255\255\039\001\255\255\255\255\255\255\043\001\ -\044\001\045\001\255\255\047\001\048\001\049\001\050\001\051\001\ -\255\255\053\001\255\255\255\255\255\255\255\255\255\255\255\255\ -\060\001\061\001\062\001\063\001\064\001\065\001\255\255\255\255\ -\255\255\069\001\070\001\255\255\255\255\073\001\074\001\075\001\ -\076\001\077\001\078\001\255\255\080\001\081\001\255\255\255\255\ -\084\001\085\001\086\001\087\001\088\001\255\255\090\001\255\255\ +\255\255\255\255\255\255\255\255\255\255\149\001\255\255\151\001\ +\255\255\153\001\255\255\255\255\255\255\157\001\158\001\159\001\ +\160\001\161\001\005\001\006\001\007\001\008\001\009\001\010\001\ +\011\001\012\001\255\255\255\255\255\255\016\001\255\255\018\001\ +\019\001\255\255\255\255\255\255\255\255\255\255\255\255\026\001\ +\255\255\255\255\029\001\255\255\255\255\255\255\255\255\255\255\ +\255\255\036\001\255\255\255\255\039\001\255\255\255\255\255\255\ +\043\001\044\001\045\001\255\255\047\001\048\001\049\001\050\001\ +\051\001\255\255\053\001\255\255\255\255\255\255\255\255\255\255\ +\255\255\060\001\061\001\062\001\063\001\064\001\065\001\255\255\ +\255\255\255\255\069\001\070\001\255\255\255\255\073\001\074\001\ +\075\001\076\001\077\001\078\001\255\255\080\001\081\001\255\255\ +\255\255\084\001\085\001\086\001\087\001\088\001\255\255\090\001\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\255\255\106\001\255\255\ -\255\255\109\001\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\255\255\122\001\255\255\ -\124\001\255\255\255\255\255\255\128\001\255\255\255\255\255\255\ -\132\001\133\001\255\255\255\255\255\255\255\255\255\255\139\001\ -\140\001\255\255\255\255\143\001\255\255\255\255\255\255\147\001\ -\255\255\149\001\255\255\151\001\152\001\153\001\255\255\255\255\ -\255\255\157\001\158\001\159\001\160\001\161\001\005\001\006\001\ -\007\001\008\001\009\001\010\001\011\001\012\001\255\255\255\255\ -\255\255\016\001\255\255\018\001\019\001\255\255\255\255\255\255\ -\255\255\255\255\255\255\026\001\255\255\255\255\029\001\255\255\ -\255\255\255\255\255\255\255\255\255\255\036\001\255\255\255\255\ -\039\001\255\255\255\255\255\255\043\001\044\001\045\001\255\255\ -\047\001\048\001\049\001\050\001\051\001\255\255\053\001\255\255\ -\255\255\255\255\255\255\255\255\255\255\060\001\061\001\062\001\ -\063\001\064\001\065\001\255\255\255\255\255\255\069\001\070\001\ -\255\255\255\255\073\001\074\001\075\001\076\001\077\001\078\001\ -\255\255\080\001\081\001\255\255\255\255\084\001\085\001\086\001\ -\087\001\088\001\255\255\090\001\255\255\255\255\255\255\255\255\ +\255\255\255\255\255\255\255\255\255\255\255\255\255\255\106\001\ +\255\255\255\255\109\001\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\255\255\255\255\255\255\255\255\255\255\122\001\ +\255\255\124\001\255\255\255\255\255\255\128\001\255\255\255\255\ +\255\255\132\001\133\001\255\255\255\255\255\255\255\255\255\255\ +\139\001\140\001\255\255\255\255\143\001\255\255\255\255\255\255\ +\147\001\255\255\149\001\255\255\151\001\152\001\153\001\255\255\ +\255\255\255\255\157\001\158\001\159\001\160\001\161\001\005\001\ +\006\001\007\001\008\001\009\001\010\001\011\001\012\001\255\255\ +\255\255\255\255\016\001\255\255\018\001\019\001\255\255\255\255\ +\255\255\255\255\255\255\255\255\026\001\255\255\255\255\029\001\ +\255\255\255\255\255\255\255\255\255\255\255\255\036\001\255\255\ +\255\255\039\001\255\255\255\255\255\255\043\001\044\001\045\001\ +\255\255\047\001\048\001\049\001\050\001\051\001\255\255\053\001\ +\255\255\255\255\255\255\255\255\255\255\255\255\060\001\061\001\ +\062\001\063\001\064\001\065\001\255\255\255\255\255\255\069\001\ +\070\001\255\255\255\255\073\001\074\001\075\001\076\001\077\001\ +\078\001\255\255\080\001\081\001\255\255\255\255\084\001\085\001\ +\086\001\087\001\088\001\255\255\090\001\255\255\255\255\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\106\001\255\255\255\255\109\001\255\255\ +\255\255\255\255\255\255\255\255\106\001\255\255\255\255\109\001\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\122\001\255\255\124\001\255\255\255\255\ -\255\255\128\001\255\255\255\255\255\255\132\001\133\001\255\255\ -\255\255\255\255\255\255\255\255\139\001\140\001\255\255\255\255\ -\143\001\255\255\255\255\255\255\147\001\255\255\149\001\255\255\ -\151\001\152\001\153\001\255\255\255\255\255\255\157\001\158\001\ -\159\001\160\001\161\001\012\001\255\255\255\255\255\255\016\001\ -\255\255\255\255\019\001\255\255\021\001\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\030\001\031\001\032\001\ -\033\001\034\001\255\255\255\255\255\255\255\255\255\255\040\001\ -\255\255\255\255\255\255\044\001\255\255\255\255\255\255\255\255\ -\049\001\050\001\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\060\001\061\001\062\001\063\001\064\001\ -\255\255\255\255\255\255\255\255\069\001\255\255\255\255\255\255\ -\073\001\255\255\255\255\255\255\077\001\078\001\255\255\255\255\ -\255\255\255\255\255\255\084\001\085\001\255\255\255\255\255\255\ -\255\255\090\001\255\255\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\255\255\255\255\122\001\255\255\124\001\255\255\ +\255\255\255\255\128\001\255\255\255\255\255\255\132\001\133\001\ +\255\255\255\255\255\255\255\255\255\255\139\001\140\001\255\255\ +\255\255\143\001\255\255\255\255\255\255\147\001\255\255\149\001\ +\255\255\151\001\152\001\153\001\255\255\255\255\255\255\157\001\ +\158\001\159\001\160\001\161\001\012\001\255\255\255\255\255\255\ +\016\001\255\255\255\255\019\001\255\255\021\001\255\255\255\255\ +\255\255\255\255\255\255\255\255\255\255\255\255\030\001\031\001\ +\032\001\033\001\034\001\255\255\255\255\255\255\255\255\255\255\ +\040\001\255\255\255\255\255\255\044\001\255\255\255\255\255\255\ +\255\255\049\001\050\001\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\255\255\255\255\060\001\061\001\062\001\063\001\ +\064\001\255\255\255\255\255\255\255\255\069\001\255\255\255\255\ +\255\255\073\001\255\255\255\255\255\255\077\001\078\001\255\255\ +\255\255\255\255\255\255\255\255\084\001\085\001\255\255\255\255\ +\255\255\255\255\090\001\255\255\255\255\255\255\255\255\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\106\001\255\255\255\255\109\001\255\255\255\255\255\255\ +\255\255\255\255\106\001\255\255\255\255\109\001\255\255\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\124\001\255\255\126\001\255\255\128\001\ -\255\255\255\255\255\255\132\001\255\255\255\255\255\255\255\255\ -\137\001\255\255\139\001\140\001\255\255\255\255\143\001\255\255\ -\255\255\255\255\255\255\255\255\255\255\255\255\151\001\255\255\ -\153\001\255\255\255\255\156\001\157\001\158\001\159\001\160\001\ -\161\001\012\001\255\255\164\001\255\255\016\001\255\255\168\001\ -\019\001\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\030\001\031\001\032\001\033\001\034\001\ -\255\255\255\255\255\255\255\255\255\255\040\001\255\255\255\255\ -\255\255\044\001\255\255\255\255\255\255\255\255\049\001\050\001\ +\255\255\255\255\255\255\255\255\124\001\255\255\126\001\255\255\ +\128\001\255\255\255\255\255\255\132\001\255\255\255\255\255\255\ +\255\255\137\001\255\255\139\001\140\001\255\255\255\255\143\001\ +\255\255\255\255\255\255\255\255\255\255\255\255\255\255\151\001\ +\255\255\153\001\255\255\255\255\156\001\157\001\158\001\159\001\ +\160\001\161\001\012\001\255\255\164\001\255\255\016\001\255\255\ +\168\001\019\001\255\255\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\255\255\255\255\030\001\031\001\032\001\033\001\ +\034\001\255\255\255\255\255\255\255\255\255\255\040\001\255\255\ +\255\255\255\255\044\001\255\255\255\255\255\255\255\255\049\001\ +\050\001\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\060\001\061\001\062\001\063\001\064\001\255\255\ +\255\255\255\255\255\255\069\001\255\255\255\255\255\255\073\001\ +\255\255\255\255\255\255\077\001\078\001\255\255\255\255\255\255\ +\255\255\255\255\084\001\085\001\255\255\255\255\255\255\255\255\ +\090\001\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\060\001\061\001\062\001\063\001\064\001\255\255\255\255\ -\255\255\255\255\069\001\255\255\255\255\255\255\073\001\255\255\ -\255\255\255\255\077\001\078\001\255\255\255\255\255\255\255\255\ -\255\255\084\001\085\001\255\255\255\255\255\255\255\255\090\001\ +\106\001\255\255\255\255\109\001\255\255\255\255\255\255\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\255\255\255\255\106\001\ -\255\255\255\255\109\001\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\124\001\255\255\126\001\255\255\128\001\255\255\ +\255\255\255\255\132\001\255\255\255\255\255\255\255\255\137\001\ +\255\255\139\001\140\001\255\255\255\255\143\001\255\255\255\255\ +\255\255\255\255\255\255\255\255\255\255\151\001\255\255\153\001\ +\255\255\255\255\255\255\157\001\158\001\159\001\160\001\161\001\ +\012\001\255\255\164\001\255\255\016\001\255\255\168\001\019\001\ +\255\255\021\001\255\255\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\255\255\031\001\032\001\033\001\034\001\255\255\ +\255\255\255\255\255\255\255\255\040\001\255\255\255\255\255\255\ +\044\001\255\255\255\255\255\255\255\255\049\001\050\001\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\124\001\255\255\126\001\255\255\128\001\255\255\255\255\ -\255\255\132\001\255\255\255\255\255\255\255\255\137\001\255\255\ -\139\001\140\001\255\255\255\255\143\001\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\151\001\255\255\153\001\255\255\ -\255\255\255\255\157\001\158\001\159\001\160\001\161\001\012\001\ -\255\255\164\001\255\255\016\001\255\255\168\001\019\001\255\255\ -\021\001\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\031\001\032\001\033\001\034\001\255\255\255\255\ -\255\255\255\255\255\255\040\001\255\255\255\255\255\255\044\001\ -\255\255\255\255\255\255\255\255\049\001\050\001\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\255\255\255\255\060\001\ -\061\001\062\001\063\001\064\001\255\255\255\255\255\255\255\255\ -\069\001\255\255\255\255\255\255\073\001\255\255\255\255\255\255\ -\077\001\078\001\255\255\255\255\255\255\255\255\255\255\084\001\ -\085\001\255\255\255\255\255\255\255\255\090\001\255\255\255\255\ +\060\001\061\001\062\001\063\001\064\001\255\255\255\255\255\255\ +\255\255\069\001\255\255\255\255\255\255\073\001\255\255\255\255\ +\255\255\077\001\078\001\255\255\255\255\255\255\255\255\255\255\ +\084\001\085\001\255\255\255\255\255\255\255\255\090\001\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\106\001\255\255\255\255\ -\109\001\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\255\255\255\255\124\001\ -\255\255\126\001\255\255\128\001\255\255\255\255\255\255\132\001\ -\255\255\255\255\255\255\255\255\137\001\255\255\139\001\140\001\ -\255\255\255\255\143\001\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\151\001\255\255\153\001\255\255\255\255\255\255\ -\157\001\158\001\159\001\160\001\161\001\012\001\255\255\164\001\ -\255\255\016\001\255\255\168\001\019\001\255\255\021\001\255\255\ +\255\255\255\255\255\255\255\255\255\255\255\255\106\001\255\255\ +\255\255\109\001\255\255\255\255\255\255\255\255\255\255\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\031\001\032\001\033\001\034\001\255\255\255\255\255\255\255\255\ -\255\255\040\001\255\255\255\255\255\255\044\001\255\255\255\255\ -\255\255\255\255\049\001\050\001\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\060\001\061\001\062\001\ -\063\001\064\001\255\255\255\255\255\255\255\255\069\001\255\255\ -\255\255\255\255\073\001\255\255\255\255\255\255\077\001\078\001\ -\255\255\255\255\255\255\255\255\255\255\084\001\085\001\255\255\ -\255\255\255\255\255\255\090\001\255\255\255\255\255\255\255\255\ +\124\001\255\255\126\001\255\255\128\001\255\255\255\255\255\255\ +\132\001\255\255\255\255\255\255\255\255\137\001\255\255\139\001\ +\140\001\255\255\255\255\143\001\255\255\255\255\255\255\255\255\ +\255\255\255\255\255\255\151\001\255\255\153\001\255\255\255\255\ +\255\255\157\001\158\001\159\001\160\001\161\001\012\001\255\255\ +\164\001\255\255\016\001\255\255\168\001\019\001\255\255\021\001\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\106\001\255\255\255\255\109\001\255\255\ +\255\255\031\001\032\001\033\001\034\001\255\255\255\255\255\255\ +\255\255\255\255\040\001\255\255\255\255\255\255\044\001\255\255\ +\255\255\255\255\255\255\049\001\050\001\255\255\255\255\255\255\ +\255\255\255\255\255\255\255\255\255\255\255\255\060\001\061\001\ +\062\001\063\001\064\001\255\255\255\255\255\255\255\255\069\001\ +\255\255\255\255\255\255\073\001\255\255\255\255\255\255\077\001\ +\078\001\255\255\255\255\255\255\255\255\255\255\084\001\085\001\ +\255\255\255\255\255\255\255\255\090\001\255\255\255\255\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\124\001\255\255\126\001\ -\255\255\128\001\255\255\255\255\255\255\132\001\255\255\255\255\ -\255\255\255\255\137\001\255\255\139\001\140\001\255\255\255\255\ -\143\001\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\151\001\255\255\153\001\255\255\255\255\255\255\157\001\158\001\ -\159\001\160\001\161\001\012\001\255\255\164\001\255\255\016\001\ -\255\255\168\001\019\001\255\255\021\001\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\255\255\031\001\032\001\ -\033\001\034\001\255\255\255\255\255\255\255\255\255\255\040\001\ -\255\255\255\255\255\255\044\001\255\255\255\255\255\255\255\255\ -\049\001\050\001\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\060\001\061\001\062\001\063\001\064\001\ -\255\255\255\255\255\255\255\255\069\001\255\255\255\255\255\255\ -\073\001\255\255\255\255\255\255\077\001\078\001\255\255\255\255\ -\255\255\255\255\255\255\084\001\085\001\255\255\255\255\255\255\ -\255\255\090\001\255\255\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\255\255\255\255\106\001\255\255\255\255\109\001\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\106\001\255\255\255\255\109\001\255\255\255\255\255\255\ +\255\255\255\255\255\255\255\255\255\255\255\255\124\001\255\255\ +\126\001\255\255\128\001\255\255\255\255\255\255\132\001\255\255\ +\255\255\255\255\255\255\137\001\255\255\139\001\140\001\255\255\ +\255\255\143\001\255\255\255\255\255\255\255\255\255\255\255\255\ +\255\255\151\001\255\255\153\001\255\255\255\255\255\255\157\001\ +\158\001\159\001\160\001\161\001\012\001\255\255\164\001\255\255\ +\016\001\255\255\168\001\019\001\255\255\021\001\255\255\255\255\ +\255\255\255\255\255\255\255\255\255\255\255\255\255\255\031\001\ +\032\001\033\001\034\001\255\255\255\255\255\255\255\255\255\255\ +\040\001\255\255\255\255\255\255\044\001\255\255\255\255\255\255\ +\255\255\049\001\050\001\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\255\255\255\255\060\001\061\001\062\001\063\001\ +\064\001\255\255\255\255\255\255\255\255\069\001\255\255\255\255\ +\255\255\073\001\255\255\255\255\255\255\077\001\078\001\255\255\ +\255\255\255\255\255\255\255\255\084\001\085\001\255\255\255\255\ +\255\255\255\255\090\001\255\255\255\255\255\255\255\255\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\124\001\255\255\126\001\255\255\128\001\ -\255\255\255\255\255\255\132\001\255\255\255\255\255\255\255\255\ -\137\001\255\255\139\001\140\001\255\255\255\255\143\001\255\255\ -\255\255\255\255\255\255\255\255\255\255\255\255\151\001\255\255\ -\153\001\255\255\255\255\255\255\157\001\158\001\159\001\160\001\ -\161\001\012\001\255\255\164\001\255\255\016\001\255\255\168\001\ -\019\001\255\255\021\001\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\031\001\032\001\033\001\034\001\ -\255\255\255\255\255\255\255\255\255\255\040\001\255\255\255\255\ -\255\255\044\001\255\255\255\255\255\255\255\255\049\001\050\001\ +\255\255\255\255\106\001\255\255\255\255\109\001\255\255\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\060\001\061\001\062\001\063\001\064\001\255\255\255\255\ -\255\255\255\255\069\001\255\255\255\255\255\255\073\001\255\255\ -\255\255\255\255\077\001\078\001\255\255\255\255\255\255\255\255\ -\255\255\084\001\085\001\255\255\255\255\255\255\255\255\090\001\ +\255\255\255\255\255\255\255\255\124\001\255\255\126\001\255\255\ +\128\001\255\255\255\255\255\255\132\001\255\255\255\255\255\255\ +\255\255\137\001\255\255\139\001\140\001\255\255\255\255\143\001\ +\255\255\255\255\255\255\255\255\255\255\255\255\255\255\151\001\ +\255\255\153\001\255\255\255\255\255\255\157\001\158\001\159\001\ +\160\001\161\001\012\001\255\255\164\001\255\255\016\001\255\255\ +\168\001\019\001\255\255\021\001\255\255\255\255\255\255\255\255\ +\255\255\255\255\255\255\255\255\255\255\031\001\032\001\033\001\ +\034\001\255\255\255\255\255\255\255\255\255\255\040\001\255\255\ +\255\255\255\255\044\001\255\255\255\255\255\255\255\255\049\001\ +\050\001\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\060\001\061\001\062\001\063\001\064\001\255\255\ +\255\255\255\255\255\255\069\001\255\255\255\255\255\255\073\001\ +\255\255\255\255\255\255\077\001\078\001\255\255\255\255\255\255\ +\255\255\255\255\084\001\085\001\255\255\255\255\255\255\255\255\ +\090\001\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\255\255\255\255\106\001\ -\255\255\255\255\109\001\255\255\255\255\255\255\255\255\255\255\ +\106\001\255\255\255\255\109\001\255\255\255\255\255\255\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\124\001\255\255\126\001\255\255\128\001\255\255\255\255\ -\255\255\132\001\255\255\255\255\255\255\255\255\137\001\255\255\ -\139\001\140\001\255\255\255\255\143\001\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\151\001\255\255\153\001\255\255\ -\255\255\255\255\157\001\158\001\159\001\160\001\161\001\012\001\ -\255\255\164\001\255\255\016\001\255\255\168\001\019\001\255\255\ -\021\001\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\031\001\032\001\033\001\034\001\255\255\255\255\ -\255\255\255\255\255\255\040\001\255\255\255\255\255\255\044\001\ -\255\255\255\255\255\255\255\255\049\001\050\001\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\255\255\255\255\060\001\ -\061\001\062\001\063\001\064\001\255\255\255\255\255\255\255\255\ -\069\001\255\255\255\255\255\255\073\001\255\255\255\255\255\255\ -\077\001\078\001\255\255\255\255\255\255\255\255\255\255\084\001\ -\085\001\255\255\255\255\255\255\255\255\090\001\255\255\255\255\ +\255\255\255\255\124\001\255\255\126\001\255\255\128\001\255\255\ +\255\255\255\255\132\001\255\255\255\255\255\255\255\255\137\001\ +\255\255\139\001\140\001\255\255\255\255\143\001\255\255\255\255\ +\255\255\255\255\255\255\255\255\255\255\151\001\255\255\153\001\ +\255\255\255\255\255\255\157\001\158\001\159\001\160\001\161\001\ +\012\001\255\255\164\001\255\255\016\001\255\255\168\001\019\001\ +\255\255\021\001\255\255\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\255\255\031\001\032\001\033\001\034\001\255\255\ +\255\255\255\255\255\255\255\255\040\001\255\255\255\255\255\255\ +\044\001\255\255\255\255\255\255\255\255\049\001\050\001\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\106\001\255\255\255\255\ -\109\001\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\255\255\255\255\124\001\ -\255\255\126\001\255\255\128\001\255\255\255\255\255\255\132\001\ -\255\255\255\255\255\255\255\255\137\001\255\255\139\001\140\001\ -\255\255\255\255\143\001\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\151\001\255\255\153\001\255\255\255\255\255\255\ -\157\001\158\001\159\001\160\001\161\001\012\001\255\255\164\001\ -\255\255\016\001\255\255\168\001\019\001\255\255\255\255\255\255\ +\060\001\061\001\062\001\063\001\064\001\255\255\255\255\255\255\ +\255\255\069\001\255\255\255\255\255\255\073\001\255\255\255\255\ +\255\255\077\001\078\001\255\255\255\255\255\255\255\255\255\255\ +\084\001\085\001\255\255\255\255\255\255\255\255\090\001\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\031\001\032\001\033\001\034\001\255\255\255\255\255\255\255\255\ -\255\255\040\001\255\255\255\255\255\255\044\001\255\255\255\255\ -\255\255\255\255\049\001\050\001\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\060\001\061\001\062\001\ -\063\001\064\001\255\255\255\255\255\255\255\255\069\001\255\255\ -\255\255\255\255\073\001\255\255\255\255\255\255\077\001\078\001\ -\255\255\255\255\255\255\255\255\255\255\084\001\085\001\255\255\ -\255\255\255\255\255\255\090\001\255\255\255\255\255\255\255\255\ +\255\255\255\255\255\255\255\255\255\255\255\255\106\001\255\255\ +\255\255\109\001\255\255\255\255\255\255\255\255\255\255\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\106\001\255\255\255\255\109\001\255\255\ +\124\001\255\255\126\001\255\255\128\001\255\255\255\255\255\255\ +\132\001\255\255\255\255\255\255\255\255\137\001\255\255\139\001\ +\140\001\255\255\255\255\143\001\255\255\255\255\255\255\255\255\ +\255\255\255\255\255\255\151\001\255\255\153\001\255\255\255\255\ +\255\255\157\001\158\001\159\001\160\001\161\001\012\001\255\255\ +\164\001\255\255\016\001\255\255\168\001\019\001\255\255\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\124\001\255\255\126\001\ -\255\255\128\001\255\255\255\255\255\255\132\001\255\255\255\255\ -\255\255\255\255\137\001\255\255\139\001\140\001\255\255\255\255\ -\143\001\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\151\001\255\255\153\001\255\255\255\255\255\255\157\001\158\001\ -\159\001\160\001\161\001\012\001\255\255\164\001\255\255\016\001\ -\255\255\168\001\019\001\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\255\255\031\001\032\001\ -\033\001\034\001\255\255\255\255\255\255\255\255\255\255\040\001\ -\255\255\255\255\255\255\044\001\255\255\255\255\255\255\255\255\ -\049\001\050\001\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\060\001\061\001\062\001\063\001\064\001\ -\255\255\255\255\255\255\255\255\069\001\255\255\255\255\255\255\ -\073\001\255\255\255\255\255\255\077\001\078\001\255\255\255\255\ -\255\255\255\255\255\255\084\001\085\001\255\255\255\255\255\255\ -\255\255\090\001\255\255\255\255\255\255\255\255\255\255\255\255\ +\255\255\031\001\032\001\033\001\034\001\255\255\255\255\255\255\ +\255\255\255\255\040\001\255\255\255\255\255\255\044\001\255\255\ +\255\255\255\255\255\255\049\001\050\001\255\255\255\255\255\255\ +\255\255\255\255\255\255\255\255\255\255\255\255\060\001\061\001\ +\062\001\063\001\064\001\255\255\255\255\255\255\255\255\069\001\ +\255\255\255\255\255\255\073\001\255\255\255\255\255\255\077\001\ +\078\001\255\255\255\255\255\255\255\255\255\255\084\001\085\001\ +\255\255\255\255\255\255\255\255\090\001\255\255\255\255\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\106\001\255\255\255\255\109\001\255\255\255\255\255\255\ +\255\255\255\255\255\255\255\255\106\001\255\255\255\255\109\001\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\124\001\255\255\126\001\255\255\128\001\ -\255\255\255\255\255\255\132\001\255\255\255\255\255\255\255\255\ -\137\001\255\255\139\001\140\001\255\255\255\255\143\001\255\255\ -\255\255\255\255\255\255\255\255\255\255\255\255\151\001\255\255\ -\153\001\255\255\255\255\255\255\157\001\158\001\159\001\160\001\ -\161\001\012\001\255\255\164\001\255\255\016\001\255\255\168\001\ -\019\001\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\031\001\032\001\033\001\034\001\ -\255\255\255\255\255\255\255\255\255\255\040\001\255\255\255\255\ -\255\255\044\001\255\255\255\255\255\255\255\255\049\001\050\001\ +\255\255\255\255\255\255\255\255\255\255\255\255\124\001\255\255\ +\126\001\255\255\128\001\255\255\255\255\255\255\132\001\255\255\ +\255\255\255\255\255\255\137\001\255\255\139\001\140\001\255\255\ +\255\255\143\001\255\255\255\255\255\255\255\255\255\255\255\255\ +\255\255\151\001\255\255\153\001\255\255\255\255\255\255\157\001\ +\158\001\159\001\160\001\161\001\012\001\255\255\164\001\255\255\ +\016\001\255\255\168\001\019\001\255\255\255\255\255\255\255\255\ +\255\255\255\255\255\255\255\255\255\255\255\255\255\255\031\001\ +\032\001\033\001\034\001\255\255\255\255\255\255\255\255\255\255\ +\040\001\255\255\255\255\255\255\044\001\255\255\255\255\255\255\ +\255\255\049\001\050\001\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\255\255\255\255\060\001\061\001\062\001\063\001\ +\064\001\255\255\255\255\255\255\255\255\069\001\255\255\255\255\ +\255\255\073\001\255\255\255\255\255\255\077\001\078\001\255\255\ +\255\255\255\255\255\255\255\255\084\001\085\001\255\255\255\255\ +\255\255\255\255\090\001\255\255\255\255\255\255\255\255\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\060\001\061\001\062\001\063\001\064\001\255\255\255\255\ -\255\255\255\255\069\001\255\255\255\255\255\255\073\001\255\255\ -\255\255\255\255\077\001\078\001\255\255\255\255\255\255\255\255\ -\255\255\084\001\085\001\255\255\255\255\255\255\255\255\090\001\ +\255\255\255\255\106\001\255\255\255\255\109\001\255\255\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\255\255\255\255\106\001\ -\255\255\255\255\109\001\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\255\255\255\255\124\001\255\255\126\001\255\255\ +\128\001\255\255\255\255\255\255\132\001\255\255\255\255\255\255\ +\255\255\137\001\255\255\139\001\140\001\255\255\255\255\143\001\ +\255\255\255\255\255\255\255\255\255\255\255\255\255\255\151\001\ +\255\255\153\001\255\255\255\255\255\255\157\001\158\001\159\001\ +\160\001\161\001\012\001\255\255\164\001\255\255\016\001\255\255\ +\168\001\019\001\255\255\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\255\255\255\255\255\255\031\001\032\001\033\001\ +\034\001\255\255\255\255\255\255\255\255\255\255\040\001\255\255\ +\255\255\255\255\044\001\255\255\255\255\255\255\255\255\049\001\ +\050\001\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\060\001\061\001\062\001\063\001\064\001\255\255\ +\255\255\255\255\255\255\069\001\255\255\255\255\255\255\073\001\ +\255\255\255\255\255\255\077\001\078\001\255\255\255\255\255\255\ +\255\255\255\255\084\001\085\001\255\255\255\255\255\255\255\255\ +\090\001\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\124\001\255\255\126\001\255\255\128\001\255\255\255\255\ -\255\255\132\001\255\255\255\255\255\255\255\255\137\001\255\255\ -\139\001\140\001\255\255\255\255\143\001\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\151\001\255\255\153\001\255\255\ -\255\255\255\255\157\001\158\001\159\001\160\001\161\001\012\001\ -\255\255\164\001\255\255\016\001\255\255\168\001\019\001\255\255\ +\106\001\255\255\255\255\109\001\255\255\255\255\255\255\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\031\001\032\001\033\001\034\001\255\255\255\255\ -\255\255\255\255\255\255\040\001\255\255\255\255\255\255\044\001\ -\255\255\255\255\255\255\255\255\049\001\050\001\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\255\255\255\255\060\001\ -\061\001\062\001\063\001\064\001\255\255\255\255\255\255\255\255\ -\069\001\255\255\255\255\255\255\073\001\255\255\255\255\255\255\ -\077\001\078\001\255\255\255\255\255\255\255\255\255\255\084\001\ -\085\001\255\255\255\255\255\255\255\255\090\001\255\255\255\255\ +\255\255\255\255\124\001\255\255\126\001\255\255\128\001\255\255\ +\255\255\255\255\132\001\255\255\255\255\255\255\255\255\137\001\ +\255\255\139\001\140\001\255\255\255\255\143\001\255\255\255\255\ +\255\255\255\255\255\255\255\255\255\255\151\001\255\255\153\001\ +\255\255\255\255\255\255\157\001\158\001\159\001\160\001\161\001\ +\012\001\255\255\164\001\255\255\016\001\255\255\168\001\019\001\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\106\001\255\255\255\255\ -\109\001\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\255\255\255\255\124\001\ -\255\255\126\001\255\255\128\001\255\255\255\255\255\255\132\001\ -\255\255\255\255\255\255\255\255\137\001\255\255\139\001\140\001\ -\255\255\255\255\143\001\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\151\001\255\255\153\001\255\255\255\255\255\255\ -\157\001\158\001\159\001\160\001\161\001\012\001\255\255\164\001\ -\255\255\016\001\255\255\168\001\019\001\255\255\255\255\255\255\ +\255\255\255\255\255\255\031\001\032\001\033\001\034\001\255\255\ +\255\255\255\255\255\255\255\255\040\001\255\255\255\255\255\255\ +\044\001\255\255\255\255\255\255\255\255\049\001\050\001\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\031\001\032\001\033\001\034\001\255\255\255\255\255\255\255\255\ -\255\255\040\001\255\255\255\255\255\255\044\001\255\255\255\255\ -\255\255\255\255\049\001\050\001\255\255\255\255\255\255\255\255\ +\060\001\061\001\062\001\063\001\064\001\255\255\255\255\255\255\ +\255\255\069\001\255\255\255\255\255\255\073\001\255\255\255\255\ +\255\255\077\001\078\001\255\255\255\255\255\255\255\255\255\255\ +\084\001\085\001\255\255\255\255\255\255\255\255\090\001\255\255\ +\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\255\255\255\255\255\255\255\255\106\001\255\255\ +\255\255\109\001\255\255\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ +\124\001\255\255\126\001\255\255\128\001\255\255\255\255\255\255\ +\132\001\255\255\255\255\255\255\255\255\137\001\255\255\139\001\ +\140\001\255\255\255\255\143\001\255\255\255\255\255\255\255\255\ +\255\255\255\255\255\255\151\001\255\255\153\001\255\255\255\255\ +\255\255\157\001\158\001\159\001\160\001\161\001\012\001\255\255\ +\164\001\255\255\016\001\255\255\168\001\019\001\255\255\255\255\ +\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ +\255\255\031\001\032\001\033\001\034\001\255\255\255\255\255\255\ +\255\255\255\255\040\001\255\255\255\255\255\255\044\001\255\255\ +\255\255\255\255\255\255\049\001\050\001\255\255\255\255\255\255\ +\255\255\255\255\255\255\255\255\255\255\255\255\060\001\061\001\ +\062\001\063\001\064\001\255\255\255\255\255\255\255\255\069\001\ +\255\255\255\255\255\255\073\001\255\255\255\255\255\255\077\001\ +\078\001\255\255\255\255\255\255\255\255\255\255\084\001\085\001\ +\255\255\255\255\255\255\255\255\090\001\255\255\255\255\255\255\ +\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\255\255\255\255\106\001\255\255\255\255\109\001\ +\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\255\255\255\255\255\255\255\255\124\001\255\255\ +\126\001\255\255\128\001\255\255\008\001\255\255\132\001\255\255\ +\012\001\255\255\255\255\137\001\016\001\139\001\140\001\019\001\ +\255\255\143\001\022\001\023\001\255\255\025\001\255\255\255\255\ +\028\001\151\001\255\255\153\001\255\255\255\255\255\255\157\001\ +\158\001\159\001\160\001\161\001\040\001\255\255\164\001\255\255\ +\044\001\255\255\168\001\255\255\255\255\255\255\050\001\255\255\ +\052\001\255\255\054\001\255\255\255\255\255\255\255\255\255\255\ +\060\001\061\001\062\001\063\001\064\001\255\255\255\255\255\255\ +\255\255\255\255\255\255\255\255\255\255\073\001\255\255\255\255\ +\255\255\077\001\078\001\255\255\255\255\255\255\255\255\255\255\ +\084\001\085\001\255\255\255\255\255\255\255\255\090\001\255\255\ +\255\255\255\255\255\255\255\255\255\255\255\255\098\001\099\001\ +\100\001\101\001\102\001\103\001\104\001\105\001\106\001\107\001\ +\108\001\109\001\110\001\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ +\124\001\255\255\255\255\255\255\128\001\255\255\255\255\255\255\ +\132\001\008\001\009\001\010\001\011\001\012\001\255\255\139\001\ +\140\001\016\001\255\255\143\001\019\001\255\255\255\255\255\255\ +\255\255\255\255\255\255\151\001\255\255\153\001\029\001\255\255\ +\255\255\157\001\158\001\159\001\160\001\161\001\255\255\255\255\ +\255\255\255\255\255\255\255\255\043\001\044\001\045\001\255\255\ +\047\001\255\255\049\001\050\001\255\255\255\255\255\255\255\255\ \255\255\255\255\255\255\255\255\255\255\060\001\061\001\062\001\ -\063\001\064\001\255\255\255\255\255\255\255\255\069\001\255\255\ -\255\255\255\255\073\001\255\255\255\255\255\255\077\001\078\001\ +\063\001\064\001\255\255\255\255\255\255\255\255\069\001\070\001\ +\255\255\072\001\073\001\255\255\255\255\076\001\077\001\078\001\ \255\255\255\255\255\255\255\255\255\255\084\001\085\001\255\255\ -\255\255\255\255\255\255\090\001\255\255\255\255\255\255\255\255\ +\255\255\088\001\255\255\090\001\255\255\255\255\255\255\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ \255\255\255\255\255\255\106\001\255\255\255\255\109\001\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\124\001\255\255\126\001\ -\255\255\128\001\255\255\008\001\255\255\132\001\255\255\012\001\ -\255\255\255\255\137\001\016\001\139\001\140\001\019\001\255\255\ -\143\001\022\001\023\001\255\255\025\001\255\255\255\255\028\001\ -\151\001\255\255\153\001\255\255\255\255\255\255\157\001\158\001\ -\159\001\160\001\161\001\040\001\255\255\164\001\255\255\044\001\ -\255\255\168\001\255\255\255\255\255\255\050\001\255\255\052\001\ -\255\255\054\001\255\255\255\255\255\255\255\255\255\255\060\001\ +\255\255\255\255\255\255\122\001\255\255\124\001\255\255\255\255\ +\255\255\128\001\255\255\255\255\255\255\132\001\008\001\009\001\ +\010\001\011\001\012\001\013\001\139\001\140\001\016\001\255\255\ +\143\001\019\001\255\255\255\255\255\255\255\255\149\001\255\255\ +\151\001\255\255\153\001\029\001\255\255\255\255\157\001\158\001\ +\159\001\160\001\161\001\255\255\255\255\255\255\255\255\255\255\ +\255\255\043\001\044\001\045\001\255\255\047\001\255\255\049\001\ +\050\001\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\060\001\061\001\062\001\063\001\064\001\255\255\ +\255\255\255\255\255\255\069\001\070\001\255\255\255\255\073\001\ +\255\255\255\255\076\001\077\001\078\001\255\255\255\255\255\255\ +\255\255\255\255\084\001\085\001\255\255\255\255\088\001\255\255\ +\090\001\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ +\106\001\255\255\255\255\109\001\255\255\255\255\255\255\255\255\ +\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ +\122\001\255\255\124\001\255\255\255\255\255\255\128\001\255\255\ +\255\255\255\255\132\001\008\001\009\001\010\001\011\001\012\001\ +\255\255\139\001\140\001\016\001\255\255\143\001\019\001\255\255\ +\255\255\255\255\255\255\149\001\255\255\151\001\255\255\153\001\ +\029\001\255\255\255\255\157\001\158\001\159\001\160\001\161\001\ +\255\255\255\255\255\255\255\255\255\255\255\255\043\001\044\001\ +\045\001\255\255\047\001\255\255\049\001\050\001\255\255\255\255\ +\255\255\255\255\255\255\255\255\255\255\255\255\255\255\060\001\ \061\001\062\001\063\001\064\001\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\073\001\255\255\255\255\255\255\ +\069\001\070\001\255\255\255\255\073\001\255\255\255\255\076\001\ \077\001\078\001\255\255\255\255\255\255\255\255\255\255\084\001\ -\085\001\255\255\255\255\255\255\255\255\090\001\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\098\001\099\001\100\001\ -\101\001\102\001\103\001\104\001\105\001\106\001\107\001\108\001\ -\109\001\110\001\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\255\255\255\255\124\001\ +\085\001\255\255\255\255\088\001\255\255\090\001\255\255\255\255\ +\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\255\255\255\255\255\255\106\001\255\255\255\255\ +\109\001\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\255\255\255\255\255\255\122\001\255\255\124\001\ \255\255\255\255\255\255\128\001\255\255\255\255\255\255\132\001\ \008\001\009\001\010\001\011\001\012\001\255\255\139\001\140\001\ \016\001\255\255\143\001\019\001\255\255\255\255\255\255\255\255\ -\255\255\255\255\151\001\255\255\153\001\029\001\255\255\255\255\ +\149\001\255\255\151\001\255\255\153\001\029\001\255\255\255\255\ \157\001\158\001\159\001\160\001\161\001\255\255\255\255\255\255\ -\255\255\255\255\255\255\043\001\044\001\045\001\255\255\047\001\ +\255\255\255\255\255\255\255\255\044\001\255\255\255\255\047\001\ \255\255\049\001\050\001\255\255\255\255\255\255\255\255\255\255\ \255\255\255\255\255\255\255\255\060\001\061\001\062\001\063\001\ \064\001\255\255\255\255\255\255\255\255\069\001\070\001\255\255\ -\072\001\073\001\255\255\255\255\076\001\077\001\078\001\255\255\ +\255\255\073\001\255\255\255\255\076\001\077\001\078\001\255\255\ \255\255\255\255\255\255\255\255\084\001\085\001\255\255\255\255\ \088\001\255\255\090\001\255\255\255\255\255\255\255\255\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ @@ -5940,11 +6008,11 @@ let yycheck = "\002\000\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ \255\255\255\255\122\001\255\255\124\001\255\255\255\255\255\255\ \128\001\255\255\255\255\255\255\132\001\008\001\009\001\010\001\ -\011\001\012\001\013\001\139\001\140\001\016\001\255\255\143\001\ +\011\001\012\001\255\255\139\001\140\001\016\001\255\255\143\001\ \019\001\255\255\255\255\255\255\255\255\149\001\255\255\151\001\ \255\255\153\001\029\001\255\255\255\255\157\001\158\001\159\001\ \160\001\161\001\255\255\255\255\255\255\255\255\255\255\255\255\ -\043\001\044\001\045\001\255\255\047\001\255\255\049\001\050\001\ +\255\255\044\001\255\255\255\255\255\255\255\255\049\001\050\001\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ \255\255\060\001\061\001\062\001\063\001\064\001\255\255\255\255\ \255\255\255\255\069\001\070\001\255\255\255\255\073\001\255\255\ @@ -5959,8 +6027,8 @@ let yycheck = "\002\000\ \139\001\140\001\016\001\255\255\143\001\019\001\255\255\255\255\ \255\255\255\255\149\001\255\255\151\001\255\255\153\001\029\001\ \255\255\255\255\157\001\158\001\159\001\160\001\161\001\255\255\ -\255\255\255\255\255\255\255\255\255\255\043\001\044\001\045\001\ -\255\255\047\001\255\255\049\001\050\001\255\255\255\255\255\255\ +\255\255\255\255\255\255\255\255\255\255\255\255\044\001\255\255\ +\255\255\255\255\255\255\049\001\050\001\255\255\255\255\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\060\001\061\001\ \062\001\063\001\064\001\255\255\255\255\255\255\255\255\069\001\ \070\001\255\255\255\255\073\001\255\255\255\255\076\001\077\001\ @@ -5975,7 +6043,7 @@ let yycheck = "\002\000\ \255\255\143\001\019\001\255\255\255\255\255\255\255\255\149\001\ \255\255\151\001\255\255\153\001\029\001\255\255\255\255\157\001\ \158\001\159\001\160\001\161\001\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\044\001\255\255\255\255\047\001\255\255\ +\255\255\255\255\255\255\044\001\255\255\255\255\255\255\255\255\ \049\001\050\001\255\255\255\255\255\255\255\255\255\255\255\255\ \255\255\255\255\255\255\060\001\061\001\062\001\063\001\064\001\ \255\255\255\255\255\255\255\255\069\001\070\001\255\255\255\255\ @@ -5983,509 +6051,461 @@ let yycheck = "\002\000\ \255\255\255\255\255\255\084\001\085\001\255\255\255\255\088\001\ \255\255\090\001\255\255\255\255\255\255\255\255\255\255\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\106\001\255\255\255\255\109\001\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\122\001\255\255\124\001\255\255\255\255\255\255\128\001\ -\255\255\255\255\255\255\132\001\008\001\009\001\010\001\011\001\ -\012\001\255\255\139\001\140\001\016\001\255\255\143\001\019\001\ -\255\255\255\255\255\255\255\255\149\001\255\255\151\001\255\255\ -\153\001\029\001\255\255\255\255\157\001\158\001\159\001\160\001\ -\161\001\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\044\001\255\255\255\255\255\255\255\255\049\001\050\001\255\255\ -\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\060\001\061\001\062\001\063\001\064\001\255\255\255\255\255\255\ -\255\255\069\001\070\001\255\255\255\255\073\001\255\255\255\255\ -\076\001\077\001\078\001\255\255\255\255\255\255\255\255\255\255\ -\084\001\085\001\255\255\255\255\088\001\255\255\090\001\255\255\ +\255\255\106\001\255\255\255\255\109\001\008\001\009\001\010\001\ +\011\001\012\001\255\255\255\255\255\255\016\001\255\255\255\255\ +\019\001\122\001\255\255\124\001\255\255\255\255\255\255\128\001\ +\255\255\255\255\255\255\132\001\255\255\255\255\255\255\255\255\ +\255\255\255\255\139\001\140\001\255\255\255\255\143\001\255\255\ +\043\001\044\001\045\001\255\255\149\001\255\255\151\001\050\001\ +\153\001\255\255\255\255\255\255\157\001\158\001\159\001\160\001\ +\161\001\060\001\061\001\062\001\063\001\064\001\255\255\255\255\ +\255\255\255\255\069\001\255\255\255\255\255\255\073\001\255\255\ +\255\255\255\255\077\001\078\001\255\255\255\255\255\255\255\255\ +\255\255\084\001\085\001\255\255\255\255\088\001\255\255\090\001\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\255\255\106\001\255\255\ -\255\255\109\001\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\255\255\122\001\255\255\ -\124\001\255\255\255\255\255\255\128\001\255\255\255\255\255\255\ -\132\001\008\001\009\001\010\001\011\001\012\001\255\255\139\001\ -\140\001\016\001\255\255\143\001\019\001\255\255\255\255\255\255\ -\255\255\149\001\255\255\151\001\255\255\153\001\029\001\255\255\ -\255\255\157\001\158\001\159\001\160\001\161\001\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\044\001\255\255\255\255\ -\255\255\255\255\049\001\050\001\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\060\001\061\001\062\001\ -\063\001\064\001\255\255\255\255\255\255\255\255\069\001\070\001\ +\255\255\255\255\255\255\255\255\255\255\255\255\255\255\106\001\ +\255\255\255\255\109\001\008\001\009\001\010\001\011\001\012\001\ +\255\255\255\255\255\255\016\001\255\255\255\255\019\001\122\001\ +\255\255\124\001\255\255\255\255\255\255\128\001\255\255\255\255\ +\255\255\132\001\255\255\255\255\255\255\255\255\255\255\255\255\ +\139\001\140\001\255\255\255\255\143\001\255\255\043\001\044\001\ +\045\001\255\255\149\001\255\255\151\001\050\001\153\001\255\255\ +\255\255\255\255\157\001\158\001\159\001\160\001\161\001\060\001\ +\061\001\062\001\063\001\064\001\255\255\255\255\255\255\255\255\ +\069\001\255\255\255\255\255\255\073\001\255\255\255\255\255\255\ +\077\001\078\001\255\255\255\255\255\255\255\255\255\255\084\001\ +\085\001\255\255\255\255\088\001\255\255\090\001\255\255\255\255\ +\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\255\255\255\255\255\255\106\001\255\255\255\255\ +\109\001\008\001\009\001\010\001\011\001\012\001\255\255\255\255\ +\255\255\016\001\255\255\255\255\019\001\122\001\255\255\124\001\ +\255\255\255\255\255\255\128\001\255\255\255\255\255\255\132\001\ +\255\255\255\255\255\255\255\255\255\255\255\255\139\001\140\001\ +\255\255\255\255\143\001\255\255\255\255\044\001\255\255\255\255\ +\149\001\255\255\151\001\050\001\153\001\255\255\255\255\255\255\ +\157\001\158\001\159\001\160\001\161\001\060\001\061\001\062\001\ +\063\001\064\001\255\255\255\255\255\255\255\255\069\001\255\255\ \255\255\255\255\073\001\255\255\255\255\076\001\077\001\078\001\ \255\255\255\255\255\255\255\255\255\255\084\001\085\001\255\255\ \255\255\088\001\255\255\090\001\255\255\255\255\255\255\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\106\001\255\255\255\255\109\001\255\255\ -\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\122\001\255\255\124\001\255\255\255\255\ -\255\255\128\001\255\255\255\255\255\255\132\001\008\001\009\001\ -\010\001\011\001\012\001\255\255\139\001\140\001\016\001\255\255\ -\143\001\019\001\255\255\255\255\255\255\255\255\149\001\255\255\ -\151\001\255\255\153\001\029\001\255\255\255\255\157\001\158\001\ -\159\001\160\001\161\001\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\044\001\255\255\255\255\255\255\255\255\049\001\ -\050\001\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\060\001\061\001\062\001\063\001\064\001\255\255\ -\255\255\255\255\255\255\069\001\070\001\255\255\255\255\073\001\ -\255\255\255\255\076\001\077\001\078\001\255\255\255\255\255\255\ -\255\255\255\255\084\001\085\001\255\255\255\255\088\001\255\255\ -\090\001\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\255\255\106\001\255\255\255\255\109\001\008\001\ +\009\001\010\001\011\001\012\001\255\255\255\255\255\255\016\001\ +\255\255\255\255\019\001\122\001\255\255\124\001\255\255\255\255\ +\255\255\128\001\255\255\255\255\255\255\132\001\255\255\255\255\ +\255\255\255\255\255\255\255\255\139\001\140\001\255\255\255\255\ +\143\001\255\255\255\255\044\001\255\255\255\255\149\001\255\255\ +\151\001\050\001\153\001\255\255\255\255\255\255\157\001\158\001\ +\159\001\160\001\161\001\060\001\061\001\062\001\063\001\064\001\ +\255\255\255\255\255\255\255\255\069\001\255\255\255\255\255\255\ +\073\001\255\255\255\255\076\001\077\001\078\001\255\255\255\255\ +\255\255\255\255\255\255\084\001\085\001\255\255\255\255\088\001\ +\255\255\090\001\255\255\255\255\255\255\255\255\255\255\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\106\001\255\255\255\255\109\001\008\001\009\001\010\001\011\001\ -\012\001\255\255\255\255\255\255\016\001\255\255\255\255\019\001\ -\122\001\255\255\124\001\255\255\255\255\255\255\128\001\255\255\ -\255\255\255\255\132\001\255\255\255\255\255\255\255\255\255\255\ -\255\255\139\001\140\001\255\255\255\255\143\001\255\255\043\001\ -\044\001\045\001\255\255\149\001\255\255\151\001\050\001\153\001\ -\255\255\255\255\255\255\157\001\158\001\159\001\160\001\161\001\ -\060\001\061\001\062\001\063\001\064\001\255\255\255\255\255\255\ -\255\255\069\001\255\255\255\255\255\255\073\001\255\255\255\255\ -\255\255\077\001\078\001\255\255\255\255\255\255\255\255\255\255\ -\084\001\085\001\255\255\255\255\088\001\255\255\090\001\255\255\ +\255\255\106\001\255\255\255\255\109\001\008\001\009\001\010\001\ +\011\001\012\001\255\255\255\255\255\255\016\001\255\255\255\255\ +\019\001\122\001\255\255\124\001\255\255\255\255\255\255\128\001\ +\255\255\255\255\255\255\132\001\255\255\255\255\255\255\255\255\ +\255\255\255\255\139\001\140\001\255\255\255\255\143\001\255\255\ +\255\255\044\001\255\255\255\255\149\001\255\255\151\001\050\001\ +\153\001\255\255\255\255\255\255\157\001\158\001\159\001\160\001\ +\161\001\060\001\061\001\062\001\063\001\064\001\255\255\255\255\ +\255\255\255\255\069\001\255\255\255\255\255\255\073\001\255\255\ +\255\255\255\255\077\001\078\001\255\255\255\255\255\255\255\255\ +\255\255\084\001\085\001\255\255\255\255\088\001\255\255\090\001\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\255\255\106\001\255\255\ -\255\255\109\001\008\001\009\001\010\001\011\001\012\001\255\255\ -\255\255\255\255\016\001\255\255\255\255\019\001\122\001\255\255\ -\124\001\255\255\255\255\255\255\128\001\255\255\255\255\255\255\ -\132\001\255\255\255\255\255\255\255\255\255\255\255\255\139\001\ -\140\001\255\255\255\255\143\001\255\255\043\001\044\001\045\001\ -\255\255\149\001\255\255\151\001\050\001\153\001\255\255\255\255\ -\255\255\157\001\158\001\159\001\160\001\161\001\060\001\061\001\ -\062\001\063\001\064\001\255\255\255\255\255\255\255\255\069\001\ -\255\255\255\255\255\255\073\001\255\255\255\255\255\255\077\001\ -\078\001\255\255\255\255\255\255\255\255\255\255\084\001\085\001\ -\255\255\255\255\088\001\255\255\090\001\255\255\255\255\255\255\ +\255\255\255\255\255\255\255\255\255\255\255\255\255\255\106\001\ +\255\255\255\255\109\001\008\001\009\001\010\001\011\001\012\001\ +\255\255\255\255\255\255\016\001\255\255\255\255\019\001\122\001\ +\255\255\124\001\255\255\255\255\255\255\128\001\255\255\255\255\ +\255\255\132\001\255\255\255\255\255\255\255\255\255\255\255\255\ +\139\001\140\001\255\255\255\255\143\001\255\255\255\255\044\001\ +\255\255\255\255\149\001\255\255\151\001\050\001\153\001\255\255\ +\255\255\255\255\157\001\158\001\159\001\160\001\161\001\060\001\ +\061\001\062\001\063\001\064\001\255\255\255\255\255\255\255\255\ +\069\001\255\255\255\255\255\255\073\001\255\255\255\255\255\255\ +\077\001\078\001\255\255\255\255\255\255\255\255\255\255\084\001\ +\085\001\255\255\255\255\088\001\255\255\090\001\255\255\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\106\001\255\255\255\255\109\001\ -\008\001\009\001\010\001\011\001\012\001\255\255\255\255\255\255\ -\016\001\255\255\255\255\019\001\122\001\255\255\124\001\255\255\ -\255\255\255\255\128\001\255\255\255\255\255\255\132\001\255\255\ -\255\255\255\255\255\255\255\255\255\255\139\001\140\001\255\255\ -\255\255\143\001\255\255\255\255\044\001\255\255\255\255\149\001\ -\255\255\151\001\050\001\153\001\255\255\255\255\255\255\157\001\ -\158\001\159\001\160\001\161\001\060\001\061\001\062\001\063\001\ -\064\001\255\255\255\255\255\255\255\255\069\001\255\255\255\255\ -\255\255\073\001\255\255\255\255\076\001\077\001\078\001\255\255\ -\255\255\255\255\255\255\255\255\084\001\085\001\255\255\255\255\ -\088\001\255\255\090\001\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\255\255\255\255\255\255\106\001\255\255\255\255\ +\109\001\008\001\009\001\010\001\011\001\012\001\255\255\255\255\ +\255\255\016\001\255\255\255\255\019\001\122\001\255\255\124\001\ +\255\255\255\255\255\255\128\001\255\255\255\255\255\255\132\001\ +\255\255\255\255\255\255\255\255\255\255\255\255\139\001\140\001\ +\255\255\255\255\143\001\255\255\255\255\044\001\255\255\255\255\ +\149\001\255\255\151\001\050\001\153\001\255\255\255\255\255\255\ +\157\001\158\001\159\001\160\001\161\001\060\001\061\001\062\001\ +\063\001\064\001\255\255\255\255\255\255\255\255\069\001\255\255\ +\255\255\255\255\073\001\255\255\255\255\255\255\077\001\078\001\ +\255\255\255\255\255\255\255\255\255\255\084\001\085\001\255\255\ +\255\255\088\001\255\255\090\001\255\255\255\255\255\255\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\106\001\255\255\255\255\109\001\008\001\009\001\ -\010\001\011\001\012\001\255\255\255\255\255\255\016\001\255\255\ -\255\255\019\001\122\001\255\255\124\001\255\255\255\255\255\255\ -\128\001\255\255\255\255\255\255\132\001\255\255\255\255\255\255\ -\255\255\255\255\255\255\139\001\140\001\255\255\255\255\143\001\ -\255\255\255\255\044\001\255\255\255\255\149\001\255\255\151\001\ -\050\001\153\001\255\255\255\255\255\255\157\001\158\001\159\001\ -\160\001\161\001\060\001\061\001\062\001\063\001\064\001\255\255\ -\255\255\255\255\255\255\069\001\255\255\255\255\255\255\073\001\ -\255\255\255\255\076\001\077\001\078\001\255\255\255\255\255\255\ -\255\255\255\255\084\001\085\001\255\255\255\255\088\001\255\255\ -\090\001\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\255\255\106\001\255\255\255\255\109\001\008\001\ +\009\001\010\001\011\001\012\001\255\255\255\255\255\255\016\001\ +\255\255\255\255\019\001\122\001\255\255\124\001\255\255\255\255\ +\255\255\128\001\255\255\255\255\255\255\132\001\255\255\255\255\ +\255\255\255\255\255\255\255\255\139\001\140\001\255\255\255\255\ +\143\001\255\255\255\255\044\001\255\255\255\255\149\001\255\255\ +\151\001\050\001\153\001\255\255\255\255\255\255\157\001\158\001\ +\159\001\160\001\161\001\060\001\061\001\062\001\063\001\064\001\ +\255\255\255\255\255\255\255\255\069\001\255\255\255\255\255\255\ +\255\255\255\255\255\255\076\001\077\001\078\001\255\255\255\255\ +\255\255\255\255\255\255\084\001\085\001\255\255\255\255\088\001\ +\255\255\090\001\255\255\255\255\255\255\255\255\255\255\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\106\001\255\255\255\255\109\001\008\001\009\001\010\001\011\001\ -\012\001\255\255\255\255\255\255\016\001\255\255\255\255\019\001\ -\122\001\255\255\124\001\255\255\255\255\255\255\128\001\255\255\ -\255\255\255\255\132\001\255\255\255\255\255\255\255\255\255\255\ -\255\255\139\001\140\001\255\255\255\255\143\001\255\255\255\255\ -\044\001\255\255\255\255\149\001\255\255\151\001\050\001\153\001\ -\255\255\255\255\255\255\157\001\158\001\159\001\160\001\161\001\ -\060\001\061\001\062\001\063\001\064\001\255\255\255\255\255\255\ -\255\255\069\001\255\255\255\255\255\255\073\001\255\255\255\255\ -\255\255\077\001\078\001\255\255\255\255\255\255\255\255\255\255\ -\084\001\085\001\255\255\255\255\088\001\255\255\090\001\255\255\ +\255\255\106\001\255\255\255\255\109\001\255\255\255\255\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\255\255\106\001\255\255\ -\255\255\109\001\008\001\009\001\010\001\011\001\012\001\255\255\ -\255\255\255\255\016\001\255\255\255\255\019\001\122\001\255\255\ -\124\001\255\255\255\255\255\255\128\001\255\255\255\255\255\255\ -\132\001\255\255\255\255\255\255\255\255\255\255\255\255\139\001\ -\140\001\255\255\255\255\143\001\255\255\255\255\044\001\255\255\ -\255\255\149\001\255\255\151\001\050\001\153\001\255\255\255\255\ -\255\255\157\001\158\001\159\001\160\001\161\001\060\001\061\001\ -\062\001\063\001\064\001\255\255\255\255\255\255\255\255\069\001\ -\255\255\255\255\255\255\073\001\255\255\255\255\255\255\077\001\ -\078\001\255\255\255\255\255\255\255\255\255\255\084\001\085\001\ -\255\255\255\255\088\001\255\255\090\001\255\255\255\255\255\255\ +\255\255\122\001\255\255\124\001\255\255\255\255\255\255\128\001\ +\255\255\012\001\255\255\132\001\255\255\016\001\255\255\255\255\ +\019\001\255\255\139\001\140\001\255\255\255\255\143\001\255\255\ +\255\255\255\255\255\255\255\255\149\001\255\255\151\001\255\255\ +\153\001\255\255\255\255\255\255\157\001\158\001\159\001\160\001\ +\161\001\044\001\255\255\255\255\255\255\255\255\049\001\050\001\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\106\001\255\255\255\255\109\001\ -\008\001\009\001\010\001\011\001\012\001\255\255\255\255\255\255\ -\016\001\255\255\255\255\019\001\122\001\255\255\124\001\255\255\ -\255\255\255\255\128\001\255\255\255\255\255\255\132\001\255\255\ -\255\255\255\255\255\255\255\255\255\255\139\001\140\001\255\255\ -\255\255\143\001\255\255\255\255\044\001\255\255\255\255\149\001\ -\255\255\151\001\050\001\153\001\255\255\255\255\255\255\157\001\ -\158\001\159\001\160\001\161\001\060\001\061\001\062\001\063\001\ -\064\001\255\255\255\255\255\255\255\255\069\001\255\255\255\255\ -\255\255\073\001\255\255\255\255\255\255\077\001\078\001\255\255\ -\255\255\255\255\255\255\255\255\084\001\085\001\255\255\255\255\ -\088\001\255\255\090\001\255\255\255\255\255\255\255\255\255\255\ +\255\255\060\001\061\001\062\001\063\001\064\001\255\255\255\255\ +\255\255\255\255\069\001\255\255\255\255\255\255\073\001\255\255\ +\255\255\255\255\077\001\078\001\255\255\255\255\255\255\255\255\ +\255\255\084\001\085\001\255\255\255\255\255\255\255\255\090\001\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\106\001\255\255\255\255\109\001\008\001\009\001\ -\010\001\011\001\012\001\255\255\255\255\255\255\016\001\255\255\ -\255\255\019\001\122\001\255\255\124\001\255\255\255\255\255\255\ -\128\001\255\255\255\255\255\255\132\001\255\255\255\255\255\255\ -\255\255\255\255\255\255\139\001\140\001\255\255\255\255\143\001\ -\255\255\255\255\044\001\255\255\255\255\149\001\255\255\151\001\ -\050\001\153\001\255\255\255\255\255\255\157\001\158\001\159\001\ -\160\001\161\001\060\001\061\001\062\001\063\001\064\001\255\255\ -\255\255\255\255\255\255\069\001\255\255\255\255\255\255\255\255\ -\255\255\255\255\076\001\077\001\078\001\255\255\255\255\255\255\ -\255\255\255\255\084\001\085\001\255\255\255\255\088\001\255\255\ -\090\001\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\255\255\255\255\255\255\255\255\255\255\106\001\ +\255\255\255\255\109\001\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\255\255\012\001\255\255\255\255\255\255\016\001\ +\255\255\124\001\019\001\255\255\255\255\128\001\255\255\255\255\ +\255\255\132\001\255\255\255\255\255\255\255\255\255\255\255\255\ +\139\001\140\001\255\255\255\255\143\001\255\255\255\255\255\255\ +\255\255\255\255\255\255\044\001\151\001\255\255\153\001\255\255\ +\049\001\050\001\157\001\158\001\159\001\160\001\161\001\255\255\ +\255\255\164\001\255\255\060\001\061\001\062\001\063\001\064\001\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\106\001\255\255\255\255\109\001\255\255\255\255\255\255\255\255\ +\073\001\255\255\255\255\255\255\077\001\078\001\255\255\255\255\ +\255\255\255\255\255\255\084\001\085\001\255\255\255\255\255\255\ +\255\255\090\001\255\255\255\255\255\255\255\255\255\255\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\122\001\255\255\124\001\255\255\255\255\255\255\128\001\255\255\ -\012\001\255\255\132\001\255\255\016\001\255\255\255\255\019\001\ -\255\255\139\001\140\001\255\255\255\255\143\001\255\255\255\255\ -\255\255\255\255\255\255\149\001\255\255\151\001\255\255\153\001\ -\255\255\255\255\255\255\157\001\158\001\159\001\160\001\161\001\ -\044\001\255\255\255\255\255\255\255\255\049\001\050\001\255\255\ +\255\255\106\001\255\255\255\255\109\001\255\255\255\255\255\255\ +\255\255\255\255\255\255\255\255\011\001\012\001\255\255\255\255\ +\255\255\016\001\255\255\124\001\019\001\255\255\255\255\128\001\ +\255\255\255\255\255\255\132\001\255\255\255\255\029\001\255\255\ +\255\255\255\255\139\001\140\001\255\255\255\255\143\001\255\255\ +\255\255\255\255\255\255\255\255\255\255\044\001\151\001\255\255\ +\153\001\255\255\049\001\050\001\157\001\158\001\159\001\160\001\ +\161\001\255\255\255\255\164\001\255\255\060\001\061\001\062\001\ +\063\001\064\001\255\255\255\255\255\255\255\255\069\001\070\001\ +\255\255\255\255\073\001\255\255\255\255\076\001\077\001\078\001\ +\255\255\255\255\255\255\255\255\255\255\084\001\085\001\255\255\ +\255\255\255\255\255\255\090\001\255\255\255\255\255\255\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\060\001\061\001\062\001\063\001\064\001\255\255\255\255\255\255\ -\255\255\069\001\255\255\255\255\255\255\073\001\255\255\255\255\ -\255\255\077\001\078\001\255\255\255\255\255\255\255\255\255\255\ -\084\001\085\001\255\255\255\255\255\255\255\255\090\001\255\255\ +\255\255\255\255\255\255\106\001\255\255\255\255\109\001\255\255\ +\255\255\255\255\011\001\012\001\255\255\255\255\255\255\016\001\ +\255\255\255\255\019\001\255\255\255\255\124\001\255\255\255\255\ +\255\255\128\001\255\255\255\255\255\255\132\001\255\255\255\255\ +\255\255\255\255\255\255\255\255\139\001\140\001\255\255\255\255\ +\143\001\255\255\255\255\044\001\255\255\255\255\149\001\255\255\ +\151\001\050\001\153\001\255\255\255\255\255\255\157\001\158\001\ +\159\001\160\001\161\001\060\001\061\001\062\001\063\001\064\001\ +\255\255\255\255\255\255\255\255\069\001\255\255\255\255\255\255\ +\073\001\255\255\255\255\076\001\077\001\078\001\255\255\255\255\ +\255\255\255\255\255\255\084\001\085\001\255\255\255\255\255\255\ +\255\255\090\001\255\255\255\255\255\255\255\255\255\255\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\255\255\106\001\255\255\ -\255\255\109\001\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\012\001\255\255\255\255\255\255\016\001\255\255\ -\124\001\019\001\255\255\255\255\128\001\255\255\255\255\255\255\ -\132\001\255\255\255\255\255\255\255\255\255\255\255\255\139\001\ -\140\001\255\255\255\255\143\001\255\255\255\255\255\255\255\255\ -\255\255\255\255\044\001\151\001\255\255\153\001\255\255\049\001\ -\050\001\157\001\158\001\159\001\160\001\161\001\255\255\255\255\ -\164\001\255\255\060\001\061\001\062\001\063\001\064\001\255\255\ -\255\255\255\255\255\255\255\255\255\255\255\255\255\255\073\001\ -\255\255\255\255\255\255\077\001\078\001\255\255\255\255\255\255\ -\255\255\255\255\084\001\085\001\255\255\255\255\255\255\255\255\ -\090\001\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ +\255\255\106\001\255\255\255\255\109\001\255\255\255\255\255\255\ +\011\001\012\001\255\255\255\255\255\255\016\001\255\255\255\255\ +\019\001\255\255\255\255\124\001\255\255\255\255\255\255\128\001\ +\255\255\255\255\255\255\132\001\255\255\255\255\255\255\255\255\ +\255\255\255\255\139\001\140\001\255\255\255\255\143\001\255\255\ +\255\255\044\001\255\255\255\255\149\001\255\255\151\001\050\001\ +\153\001\255\255\255\255\255\255\157\001\158\001\159\001\160\001\ +\161\001\060\001\061\001\062\001\063\001\064\001\255\255\255\255\ +\255\255\255\255\069\001\255\255\255\255\255\255\073\001\255\255\ +\255\255\255\255\077\001\078\001\255\255\255\255\255\255\255\255\ +\255\255\084\001\085\001\255\255\255\255\255\255\255\255\090\001\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\106\001\255\255\255\255\109\001\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\011\001\012\001\255\255\255\255\255\255\ -\016\001\255\255\124\001\019\001\255\255\255\255\128\001\255\255\ -\255\255\255\255\132\001\255\255\255\255\029\001\255\255\255\255\ -\255\255\139\001\140\001\255\255\255\255\143\001\255\255\255\255\ -\255\255\255\255\255\255\255\255\044\001\151\001\255\255\153\001\ -\255\255\049\001\050\001\157\001\158\001\159\001\160\001\161\001\ -\255\255\255\255\164\001\255\255\060\001\061\001\062\001\063\001\ -\064\001\255\255\255\255\255\255\255\255\069\001\070\001\255\255\ -\255\255\073\001\255\255\255\255\076\001\077\001\078\001\255\255\ -\255\255\255\255\255\255\255\255\084\001\085\001\255\255\255\255\ -\255\255\255\255\090\001\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\255\255\255\255\255\255\255\255\255\255\106\001\ +\255\255\255\255\109\001\255\255\255\255\255\255\011\001\012\001\ +\255\255\255\255\255\255\016\001\255\255\255\255\019\001\255\255\ +\255\255\124\001\255\255\255\255\255\255\128\001\255\255\255\255\ +\255\255\132\001\255\255\255\255\255\255\255\255\255\255\255\255\ +\139\001\140\001\255\255\255\255\143\001\255\255\255\255\044\001\ +\255\255\255\255\149\001\255\255\151\001\050\001\153\001\255\255\ +\255\255\255\255\157\001\158\001\159\001\160\001\161\001\060\001\ +\061\001\062\001\063\001\064\001\255\255\255\255\255\255\255\255\ +\069\001\255\255\255\255\255\255\073\001\255\255\255\255\255\255\ +\077\001\078\001\255\255\255\255\255\255\255\255\255\255\084\001\ +\085\001\255\255\255\255\255\255\255\255\090\001\255\255\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\106\001\255\255\255\255\109\001\255\255\255\255\ -\255\255\011\001\012\001\255\255\255\255\255\255\016\001\255\255\ -\255\255\019\001\255\255\255\255\124\001\255\255\255\255\255\255\ -\128\001\255\255\255\255\255\255\132\001\255\255\255\255\255\255\ -\255\255\255\255\255\255\139\001\140\001\255\255\255\255\143\001\ -\255\255\255\255\044\001\255\255\255\255\149\001\255\255\151\001\ -\050\001\153\001\255\255\255\255\255\255\157\001\158\001\159\001\ -\160\001\161\001\060\001\061\001\062\001\063\001\064\001\255\255\ -\255\255\255\255\255\255\069\001\255\255\255\255\255\255\073\001\ -\255\255\255\255\076\001\077\001\078\001\255\255\255\255\255\255\ -\255\255\255\255\084\001\085\001\255\255\255\255\255\255\255\255\ -\090\001\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\255\255\255\255\255\255\106\001\255\255\255\255\ +\109\001\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ +\011\001\255\255\013\001\255\255\255\255\255\255\255\255\124\001\ +\019\001\255\255\255\255\128\001\255\255\255\255\255\255\132\001\ +\255\255\255\255\029\001\255\255\255\255\255\255\139\001\140\001\ +\255\255\255\255\143\001\255\255\255\255\255\255\255\255\255\255\ +\149\001\044\001\151\001\255\255\153\001\255\255\049\001\050\001\ +\157\001\158\001\159\001\160\001\161\001\255\255\255\255\255\255\ +\255\255\060\001\061\001\062\001\063\001\064\001\255\255\255\255\ +\255\255\255\255\069\001\255\255\255\255\255\255\073\001\255\255\ +\255\255\076\001\255\255\078\001\255\255\255\255\255\255\255\255\ +\255\255\084\001\085\001\255\255\255\255\255\255\255\255\090\001\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\106\001\255\255\255\255\109\001\255\255\255\255\255\255\011\001\ -\012\001\255\255\255\255\255\255\016\001\255\255\255\255\019\001\ -\255\255\255\255\124\001\255\255\255\255\255\255\128\001\255\255\ -\255\255\255\255\132\001\255\255\255\255\255\255\255\255\255\255\ -\255\255\139\001\140\001\255\255\255\255\143\001\255\255\255\255\ -\044\001\255\255\255\255\149\001\255\255\151\001\050\001\153\001\ -\255\255\255\255\255\255\157\001\158\001\159\001\160\001\161\001\ -\060\001\061\001\062\001\063\001\064\001\255\255\255\255\255\255\ -\255\255\069\001\255\255\255\255\255\255\073\001\255\255\255\255\ -\255\255\077\001\078\001\255\255\255\255\255\255\255\255\255\255\ -\084\001\085\001\255\255\255\255\255\255\255\255\090\001\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\255\255\106\001\255\255\ -\255\255\109\001\255\255\255\255\255\255\011\001\012\001\255\255\ -\255\255\255\255\016\001\255\255\255\255\019\001\255\255\255\255\ -\124\001\255\255\255\255\255\255\128\001\255\255\255\255\255\255\ -\132\001\255\255\255\255\255\255\255\255\255\255\255\255\139\001\ -\140\001\255\255\255\255\143\001\255\255\255\255\044\001\255\255\ -\255\255\149\001\255\255\151\001\050\001\153\001\255\255\255\255\ -\255\255\157\001\158\001\159\001\160\001\161\001\060\001\061\001\ -\062\001\063\001\064\001\255\255\255\255\255\255\255\255\069\001\ -\255\255\255\255\255\255\073\001\255\255\255\255\255\255\077\001\ -\078\001\255\255\255\255\255\255\255\255\255\255\084\001\085\001\ -\255\255\255\255\255\255\255\255\090\001\255\255\255\255\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\106\001\255\255\255\255\109\001\ -\255\255\255\255\255\255\255\255\255\255\255\255\255\255\011\001\ -\255\255\255\255\255\255\255\255\255\255\255\255\124\001\019\001\ -\255\255\255\255\128\001\255\255\255\255\255\255\132\001\255\255\ -\255\255\029\001\255\255\255\255\255\255\139\001\140\001\255\255\ -\255\255\143\001\255\255\255\255\255\255\255\255\255\255\149\001\ -\044\001\151\001\255\255\153\001\255\255\049\001\050\001\157\001\ -\158\001\159\001\160\001\161\001\255\255\255\255\255\255\255\255\ -\060\001\061\001\062\001\063\001\064\001\255\255\255\255\255\255\ -\255\255\069\001\070\001\255\255\255\255\073\001\255\255\255\255\ -\076\001\255\255\078\001\255\255\255\255\255\255\255\255\255\255\ -\084\001\085\001\255\255\255\255\255\255\255\255\090\001\255\255\ +\255\255\255\255\011\001\255\255\255\255\255\255\255\255\255\255\ +\255\255\124\001\019\001\255\255\255\255\128\001\255\255\255\255\ +\255\255\132\001\255\255\255\255\029\001\255\255\255\255\255\255\ +\139\001\140\001\255\255\255\255\143\001\255\255\255\255\255\255\ +\255\255\255\255\255\255\044\001\151\001\255\255\153\001\255\255\ +\049\001\050\001\157\001\158\001\159\001\160\001\161\001\255\255\ +\255\255\255\255\255\255\060\001\061\001\062\001\063\001\064\001\ +\255\255\255\255\255\255\255\255\069\001\255\255\255\255\255\255\ +\073\001\255\255\255\255\076\001\255\255\078\001\255\255\255\255\ +\255\255\255\255\255\255\084\001\085\001\255\255\255\255\255\255\ +\255\255\090\001\255\255\255\255\255\255\255\255\255\255\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\255\255\255\255\011\001\255\255\255\255\255\255\ +\255\255\255\255\255\255\124\001\019\001\255\255\255\255\128\001\ +\255\255\255\255\255\255\132\001\255\255\255\255\029\001\136\001\ +\255\255\255\255\139\001\140\001\255\255\255\255\143\001\255\255\ +\255\255\255\255\255\255\255\255\255\255\044\001\151\001\255\255\ +\153\001\255\255\049\001\050\001\157\001\158\001\159\001\160\001\ +\161\001\255\255\255\255\255\255\255\255\060\001\061\001\062\001\ +\063\001\064\001\255\255\255\255\255\255\255\255\069\001\255\255\ +\255\255\255\255\073\001\255\255\255\255\076\001\255\255\078\001\ +\255\255\255\255\255\255\255\255\255\255\084\001\085\001\255\255\ +\255\255\255\255\255\255\090\001\255\255\255\255\255\255\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\011\001\255\255\013\001\255\255\255\255\255\255\255\255\ -\124\001\019\001\255\255\255\255\128\001\255\255\255\255\255\255\ -\132\001\255\255\255\255\029\001\255\255\255\255\255\255\139\001\ -\140\001\255\255\255\255\143\001\255\255\255\255\255\255\255\255\ -\255\255\255\255\044\001\151\001\255\255\153\001\255\255\049\001\ -\050\001\157\001\158\001\159\001\160\001\161\001\255\255\255\255\ -\255\255\255\255\060\001\061\001\062\001\063\001\064\001\255\255\ -\255\255\255\255\255\255\069\001\255\255\255\255\255\255\073\001\ -\255\255\255\255\076\001\255\255\078\001\255\255\255\255\255\255\ -\255\255\255\255\084\001\085\001\255\255\255\255\255\255\255\255\ -\090\001\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\255\255\255\255\255\255\255\255\011\001\255\255\ +\255\255\255\255\255\255\255\255\255\255\124\001\019\001\255\255\ +\255\255\128\001\255\255\255\255\255\255\132\001\255\255\255\255\ +\029\001\136\001\255\255\255\255\139\001\140\001\255\255\255\255\ +\143\001\255\255\255\255\255\255\255\255\255\255\255\255\044\001\ +\151\001\255\255\153\001\255\255\049\001\050\001\157\001\158\001\ +\159\001\160\001\161\001\255\255\255\255\255\255\255\255\060\001\ +\061\001\062\001\063\001\064\001\255\255\255\255\255\255\255\255\ +\069\001\255\255\255\255\255\255\073\001\255\255\255\255\076\001\ +\255\255\078\001\255\255\255\255\255\255\255\255\255\255\084\001\ +\085\001\255\255\255\255\255\255\255\255\090\001\255\255\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\011\001\255\255\255\255\255\255\255\255\ -\255\255\255\255\124\001\019\001\255\255\255\255\128\001\255\255\ -\255\255\255\255\132\001\255\255\255\255\029\001\255\255\255\255\ -\255\255\139\001\140\001\255\255\255\255\143\001\255\255\255\255\ -\255\255\255\255\255\255\255\255\044\001\151\001\255\255\153\001\ -\255\255\049\001\050\001\157\001\158\001\159\001\160\001\161\001\ -\255\255\255\255\255\255\255\255\060\001\061\001\062\001\063\001\ -\064\001\255\255\255\255\255\255\255\255\069\001\255\255\255\255\ -\255\255\073\001\255\255\255\255\076\001\255\255\078\001\255\255\ -\255\255\255\255\255\255\255\255\084\001\085\001\255\255\255\255\ -\255\255\255\255\090\001\255\255\255\255\255\255\255\255\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\011\001\255\255\255\255\ -\255\255\255\255\255\255\255\255\124\001\019\001\255\255\255\255\ -\128\001\255\255\255\255\255\255\132\001\255\255\255\255\029\001\ -\136\001\255\255\255\255\139\001\140\001\255\255\255\255\143\001\ -\255\255\255\255\255\255\255\255\255\255\255\255\044\001\151\001\ -\255\255\153\001\255\255\049\001\050\001\157\001\158\001\159\001\ -\160\001\161\001\255\255\255\255\255\255\255\255\060\001\061\001\ -\062\001\063\001\064\001\255\255\255\255\255\255\255\255\069\001\ -\255\255\255\255\255\255\073\001\255\255\255\255\076\001\255\255\ -\078\001\255\255\255\255\255\255\255\255\255\255\084\001\085\001\ -\255\255\255\255\255\255\255\255\090\001\255\255\255\255\255\255\ +\011\001\255\255\255\255\255\255\255\255\255\255\255\255\124\001\ +\019\001\255\255\255\255\128\001\255\255\255\255\255\255\132\001\ +\255\255\255\255\029\001\136\001\255\255\255\255\139\001\140\001\ +\255\255\255\255\143\001\255\255\255\255\255\255\255\255\255\255\ +\255\255\044\001\151\001\255\255\153\001\255\255\049\001\050\001\ +\157\001\158\001\159\001\160\001\161\001\255\255\255\255\255\255\ +\255\255\060\001\061\001\062\001\063\001\064\001\255\255\255\255\ +\255\255\255\255\069\001\255\255\255\255\255\255\073\001\255\255\ +\255\255\076\001\255\255\078\001\255\255\255\255\255\255\255\255\ +\255\255\084\001\085\001\255\255\255\255\255\255\255\255\090\001\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\255\255\255\255\011\001\ -\255\255\255\255\255\255\255\255\255\255\255\255\124\001\019\001\ -\255\255\255\255\128\001\255\255\255\255\255\255\132\001\255\255\ -\255\255\029\001\136\001\255\255\255\255\139\001\140\001\255\255\ -\255\255\143\001\255\255\255\255\255\255\255\255\255\255\255\255\ -\044\001\151\001\255\255\153\001\255\255\049\001\050\001\157\001\ -\158\001\159\001\160\001\161\001\255\255\255\255\255\255\255\255\ -\060\001\061\001\062\001\063\001\064\001\255\255\255\255\255\255\ -\255\255\069\001\255\255\255\255\255\255\073\001\255\255\255\255\ -\076\001\255\255\078\001\255\255\255\255\255\255\255\255\255\255\ -\084\001\085\001\255\255\255\255\255\255\255\255\090\001\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\011\001\255\255\255\255\255\255\255\255\255\255\ +\255\255\124\001\019\001\255\255\255\255\128\001\255\255\255\255\ +\255\255\132\001\255\255\255\255\029\001\255\255\255\255\255\255\ +\139\001\140\001\255\255\142\001\143\001\255\255\255\255\255\255\ +\255\255\255\255\255\255\044\001\151\001\255\255\153\001\255\255\ +\049\001\050\001\157\001\158\001\159\001\160\001\161\001\255\255\ +\255\255\255\255\255\255\060\001\061\001\062\001\063\001\064\001\ +\255\255\255\255\255\255\255\255\069\001\255\255\255\255\255\255\ +\073\001\255\255\255\255\076\001\255\255\078\001\255\255\255\255\ +\255\255\255\255\255\255\084\001\085\001\255\255\255\255\255\255\ +\255\255\090\001\255\255\255\255\255\255\255\255\255\255\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\011\001\255\255\255\255\255\255\255\255\255\255\255\255\ -\124\001\019\001\255\255\255\255\128\001\255\255\255\255\255\255\ -\132\001\255\255\255\255\029\001\136\001\255\255\255\255\139\001\ -\140\001\255\255\255\255\143\001\255\255\255\255\255\255\255\255\ -\255\255\255\255\044\001\151\001\255\255\153\001\255\255\049\001\ -\050\001\157\001\158\001\159\001\160\001\161\001\255\255\255\255\ -\255\255\255\255\060\001\061\001\062\001\063\001\064\001\255\255\ -\255\255\255\255\255\255\069\001\255\255\255\255\255\255\073\001\ -\255\255\255\255\076\001\255\255\078\001\255\255\255\255\255\255\ -\255\255\255\255\084\001\085\001\255\255\255\255\255\255\255\255\ -\090\001\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\255\255\255\255\011\001\255\255\255\255\255\255\ +\255\255\255\255\255\255\124\001\019\001\255\255\255\255\128\001\ +\255\255\255\255\255\255\132\001\255\255\255\255\029\001\255\255\ +\255\255\255\255\139\001\140\001\255\255\255\255\143\001\255\255\ +\255\255\255\255\255\255\255\255\255\255\044\001\151\001\255\255\ +\153\001\255\255\049\001\050\001\157\001\158\001\159\001\160\001\ +\161\001\255\255\255\255\255\255\255\255\060\001\061\001\062\001\ +\063\001\064\001\255\255\255\255\255\255\255\255\069\001\255\255\ +\255\255\255\255\073\001\255\255\255\255\076\001\255\255\078\001\ +\255\255\255\255\255\255\255\255\255\255\084\001\085\001\255\255\ +\255\255\255\255\255\255\090\001\255\255\255\255\255\255\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\011\001\255\255\255\255\255\255\255\255\ -\255\255\255\255\124\001\019\001\255\255\255\255\128\001\255\255\ -\255\255\255\255\132\001\255\255\255\255\029\001\255\255\255\255\ -\255\255\139\001\140\001\255\255\142\001\143\001\255\255\255\255\ -\255\255\255\255\255\255\255\255\044\001\151\001\255\255\153\001\ -\255\255\049\001\050\001\157\001\158\001\159\001\160\001\161\001\ -\255\255\255\255\255\255\255\255\060\001\061\001\062\001\063\001\ -\064\001\255\255\255\255\255\255\255\255\069\001\255\255\255\255\ -\255\255\073\001\255\255\255\255\076\001\255\255\078\001\255\255\ -\255\255\255\255\255\255\255\255\084\001\085\001\255\255\255\255\ -\255\255\255\255\090\001\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\255\255\255\255\255\255\255\255\011\001\255\255\ +\255\255\255\255\255\255\255\255\255\255\124\001\019\001\255\255\ +\255\255\128\001\255\255\255\255\255\255\132\001\255\255\255\255\ +\029\001\255\255\255\255\255\255\139\001\140\001\255\255\255\255\ +\143\001\255\255\255\255\255\255\255\255\255\255\255\255\044\001\ +\151\001\255\255\153\001\255\255\049\001\050\001\157\001\158\001\ +\159\001\160\001\161\001\255\255\255\255\255\255\255\255\060\001\ +\061\001\062\001\063\001\064\001\255\255\255\255\255\255\255\255\ +\069\001\255\255\255\255\255\255\073\001\255\255\255\255\076\001\ +\255\255\078\001\255\255\255\255\255\255\255\255\255\255\084\001\ +\085\001\255\255\255\255\255\255\255\255\090\001\255\255\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\011\001\255\255\255\255\ -\255\255\255\255\255\255\255\255\124\001\019\001\255\255\255\255\ -\128\001\255\255\255\255\255\255\132\001\255\255\255\255\029\001\ -\255\255\255\255\255\255\139\001\140\001\255\255\255\255\143\001\ -\255\255\255\255\255\255\255\255\255\255\255\255\044\001\151\001\ -\255\255\153\001\255\255\049\001\050\001\157\001\158\001\159\001\ -\160\001\161\001\255\255\255\255\255\255\255\255\060\001\061\001\ -\062\001\063\001\064\001\255\255\255\255\255\255\255\255\069\001\ -\255\255\255\255\255\255\073\001\255\255\255\255\076\001\255\255\ -\078\001\255\255\255\255\255\255\255\255\255\255\084\001\085\001\ -\255\255\255\255\255\255\255\255\090\001\255\255\255\255\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ +\255\255\012\001\255\255\255\255\255\255\016\001\255\255\124\001\ +\019\001\255\255\255\255\128\001\255\255\255\255\255\255\132\001\ +\255\255\255\255\029\001\255\255\255\255\255\255\139\001\140\001\ +\255\255\255\255\143\001\255\255\255\255\255\255\255\255\255\255\ +\255\255\044\001\151\001\255\255\153\001\255\255\049\001\050\001\ +\157\001\158\001\159\001\160\001\161\001\255\255\255\255\255\255\ +\255\255\060\001\061\001\062\001\063\001\064\001\255\255\255\255\ +\255\255\255\255\255\255\255\255\255\255\255\255\073\001\255\255\ +\255\255\076\001\077\001\078\001\255\255\255\255\255\255\255\255\ +\255\255\084\001\085\001\255\255\255\255\255\255\255\255\090\001\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\255\255\255\255\011\001\ -\255\255\255\255\255\255\255\255\255\255\255\255\124\001\019\001\ -\255\255\255\255\128\001\255\255\255\255\255\255\132\001\255\255\ -\255\255\029\001\255\255\255\255\255\255\139\001\140\001\255\255\ -\255\255\143\001\255\255\255\255\255\255\255\255\255\255\255\255\ -\044\001\151\001\255\255\153\001\255\255\049\001\050\001\157\001\ -\158\001\159\001\160\001\161\001\255\255\255\255\255\255\255\255\ -\060\001\061\001\062\001\063\001\064\001\255\255\255\255\255\255\ -\255\255\069\001\255\255\255\255\255\255\073\001\255\255\255\255\ -\076\001\255\255\078\001\255\255\255\255\255\255\255\255\255\255\ -\084\001\085\001\255\255\255\255\255\255\255\255\090\001\255\255\ +\255\255\255\255\255\255\255\255\255\255\255\255\255\255\106\001\ +\255\255\255\255\109\001\255\255\255\255\255\255\255\255\012\001\ +\255\255\255\255\255\255\016\001\255\255\255\255\019\001\255\255\ +\255\255\124\001\255\255\255\255\255\255\128\001\255\255\255\255\ +\255\255\132\001\255\255\255\255\255\255\255\255\255\255\255\255\ +\139\001\140\001\255\255\255\255\143\001\255\255\255\255\044\001\ +\255\255\255\255\255\255\255\255\151\001\050\001\153\001\255\255\ +\255\255\255\255\157\001\158\001\159\001\160\001\161\001\060\001\ +\061\001\062\001\063\001\064\001\255\255\255\255\255\255\255\255\ +\255\255\255\255\255\255\255\255\073\001\255\255\255\255\255\255\ +\077\001\078\001\255\255\255\255\255\255\255\255\255\255\084\001\ +\085\001\255\255\255\255\255\255\255\255\090\001\255\255\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\255\255\255\255\255\255\106\001\255\255\255\255\ +\109\001\255\255\255\255\255\255\255\255\012\001\255\255\255\255\ +\255\255\016\001\255\255\255\255\019\001\255\255\255\255\124\001\ +\255\255\255\255\255\255\128\001\255\255\255\255\255\255\132\001\ +\255\255\255\255\255\255\136\001\255\255\255\255\139\001\140\001\ +\255\255\255\255\143\001\255\255\255\255\044\001\255\255\255\255\ +\255\255\255\255\151\001\050\001\153\001\255\255\255\255\255\255\ +\157\001\158\001\159\001\160\001\161\001\060\001\061\001\062\001\ +\063\001\064\001\255\255\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\073\001\255\255\255\255\255\255\077\001\078\001\ +\255\255\255\255\255\255\255\255\255\255\084\001\085\001\255\255\ +\255\255\255\255\255\255\090\001\255\255\255\255\255\255\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\255\255\106\001\255\255\255\255\109\001\255\255\ +\255\255\255\255\255\255\012\001\255\255\255\255\255\255\016\001\ +\255\255\255\255\019\001\255\255\255\255\124\001\255\255\255\255\ +\255\255\128\001\255\255\255\255\255\255\132\001\255\255\255\255\ +\255\255\136\001\255\255\255\255\139\001\140\001\255\255\255\255\ +\143\001\255\255\255\255\044\001\255\255\255\255\255\255\255\255\ +\151\001\050\001\153\001\255\255\255\255\255\255\157\001\158\001\ +\159\001\160\001\161\001\060\001\061\001\062\001\063\001\064\001\ +\255\255\255\255\255\255\255\255\069\001\255\255\255\255\255\255\ +\073\001\255\255\255\255\255\255\077\001\078\001\255\255\255\255\ +\255\255\255\255\255\255\084\001\085\001\255\255\255\255\255\255\ +\255\255\090\001\255\255\255\255\255\255\255\255\255\255\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\012\001\255\255\255\255\255\255\016\001\255\255\ -\124\001\019\001\255\255\255\255\128\001\255\255\255\255\255\255\ -\132\001\255\255\255\255\029\001\255\255\255\255\255\255\139\001\ -\140\001\255\255\255\255\143\001\255\255\255\255\255\255\255\255\ -\255\255\255\255\044\001\151\001\255\255\153\001\255\255\049\001\ -\050\001\157\001\158\001\159\001\160\001\161\001\255\255\255\255\ -\255\255\255\255\060\001\061\001\062\001\063\001\064\001\255\255\ -\255\255\255\255\255\255\255\255\255\255\255\255\255\255\073\001\ -\255\255\255\255\076\001\077\001\078\001\255\255\255\255\255\255\ -\255\255\255\255\084\001\085\001\255\255\255\255\255\255\255\255\ -\090\001\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ +\255\255\106\001\255\255\255\255\109\001\255\255\255\255\255\255\ +\255\255\012\001\255\255\255\255\255\255\016\001\255\255\255\255\ +\019\001\255\255\255\255\124\001\255\255\255\255\255\255\128\001\ +\255\255\255\255\255\255\132\001\255\255\255\255\255\255\255\255\ +\255\255\255\255\139\001\140\001\255\255\255\255\143\001\255\255\ +\255\255\044\001\255\255\255\255\255\255\255\255\151\001\050\001\ +\153\001\255\255\255\255\255\255\157\001\158\001\159\001\160\001\ +\161\001\060\001\061\001\062\001\063\001\064\001\255\255\255\255\ +\255\255\255\255\255\255\255\255\255\255\255\255\073\001\255\255\ +\255\255\255\255\077\001\078\001\255\255\255\255\255\255\255\255\ +\255\255\084\001\085\001\255\255\255\255\255\255\255\255\090\001\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\106\001\255\255\255\255\109\001\255\255\255\255\255\255\255\255\ -\012\001\255\255\255\255\255\255\016\001\255\255\255\255\019\001\ -\255\255\255\255\124\001\255\255\255\255\255\255\128\001\255\255\ -\255\255\255\255\132\001\255\255\255\255\255\255\255\255\255\255\ -\255\255\139\001\140\001\255\255\255\255\143\001\255\255\255\255\ -\044\001\255\255\255\255\255\255\255\255\151\001\050\001\153\001\ -\255\255\255\255\255\255\157\001\158\001\159\001\160\001\161\001\ -\060\001\061\001\062\001\063\001\064\001\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\073\001\255\255\255\255\ -\255\255\077\001\078\001\255\255\255\255\255\255\255\255\255\255\ -\084\001\085\001\255\255\255\255\255\255\255\255\090\001\255\255\ +\255\255\255\255\255\255\255\255\255\255\255\255\255\255\106\001\ +\255\255\255\255\109\001\255\255\255\255\255\255\255\255\012\001\ +\255\255\255\255\255\255\016\001\255\255\255\255\019\001\255\255\ +\255\255\124\001\255\255\255\255\255\255\128\001\255\255\255\255\ +\255\255\132\001\255\255\255\255\255\255\136\001\255\255\255\255\ +\139\001\140\001\255\255\255\255\143\001\255\255\255\255\044\001\ +\255\255\255\255\255\255\255\255\151\001\050\001\153\001\255\255\ +\255\255\255\255\157\001\158\001\159\001\160\001\161\001\060\001\ +\061\001\062\001\063\001\064\001\255\255\255\255\255\255\255\255\ +\255\255\255\255\255\255\255\255\073\001\255\255\255\255\255\255\ +\077\001\078\001\255\255\255\255\255\255\255\255\255\255\084\001\ +\085\001\255\255\255\255\255\255\255\255\090\001\255\255\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\255\255\106\001\255\255\ -\255\255\109\001\255\255\255\255\255\255\255\255\012\001\255\255\ -\255\255\255\255\016\001\255\255\255\255\019\001\255\255\255\255\ -\124\001\255\255\255\255\255\255\128\001\255\255\255\255\255\255\ -\132\001\255\255\255\255\255\255\136\001\255\255\255\255\139\001\ -\140\001\255\255\255\255\143\001\255\255\255\255\044\001\255\255\ -\255\255\255\255\255\255\151\001\050\001\153\001\255\255\255\255\ -\255\255\157\001\158\001\159\001\160\001\161\001\060\001\061\001\ -\062\001\063\001\064\001\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\073\001\255\255\255\255\255\255\077\001\ -\078\001\255\255\255\255\255\255\255\255\255\255\084\001\085\001\ -\255\255\255\255\255\255\255\255\090\001\255\255\255\255\255\255\ +\255\255\255\255\255\255\255\255\255\255\106\001\255\255\255\255\ +\109\001\255\255\255\255\255\255\255\255\012\001\255\255\255\255\ +\255\255\016\001\255\255\255\255\019\001\255\255\255\255\124\001\ +\255\255\255\255\255\255\128\001\255\255\255\255\255\255\132\001\ +\255\255\255\255\255\255\136\001\255\255\255\255\139\001\140\001\ +\255\255\255\255\143\001\255\255\255\255\044\001\255\255\255\255\ +\255\255\255\255\151\001\050\001\153\001\255\255\255\255\255\255\ +\157\001\158\001\159\001\160\001\161\001\060\001\061\001\062\001\ +\063\001\064\001\255\255\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\073\001\255\255\255\255\255\255\077\001\078\001\ +\255\255\255\255\255\255\255\255\255\255\084\001\085\001\255\255\ +\255\255\255\255\255\255\090\001\255\255\255\255\255\255\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\106\001\255\255\255\255\109\001\ -\255\255\255\255\255\255\255\255\012\001\255\255\255\255\255\255\ -\016\001\255\255\255\255\019\001\255\255\255\255\124\001\255\255\ -\255\255\255\255\128\001\255\255\255\255\255\255\132\001\255\255\ -\255\255\255\255\136\001\255\255\255\255\139\001\140\001\255\255\ -\255\255\143\001\255\255\255\255\044\001\255\255\255\255\255\255\ -\255\255\151\001\050\001\153\001\255\255\255\255\255\255\157\001\ -\158\001\159\001\160\001\161\001\060\001\061\001\062\001\063\001\ -\064\001\255\255\255\255\255\255\255\255\069\001\255\255\255\255\ -\255\255\073\001\255\255\255\255\255\255\077\001\078\001\255\255\ -\255\255\255\255\255\255\255\255\084\001\085\001\255\255\255\255\ -\255\255\255\255\090\001\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\255\255\106\001\255\255\255\255\109\001\255\255\ +\255\255\255\255\255\255\012\001\255\255\255\255\255\255\016\001\ +\255\255\255\255\019\001\255\255\255\255\124\001\255\255\255\255\ +\255\255\128\001\255\255\255\255\255\255\132\001\255\255\255\255\ +\255\255\136\001\255\255\255\255\139\001\140\001\255\255\255\255\ +\143\001\255\255\255\255\044\001\255\255\255\255\255\255\255\255\ +\151\001\050\001\153\001\255\255\255\255\255\255\157\001\158\001\ +\159\001\160\001\161\001\060\001\061\001\062\001\063\001\064\001\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\106\001\255\255\255\255\109\001\255\255\255\255\ -\255\255\255\255\012\001\255\255\255\255\255\255\016\001\255\255\ -\255\255\019\001\255\255\255\255\124\001\255\255\255\255\255\255\ -\128\001\255\255\255\255\255\255\132\001\255\255\255\255\255\255\ -\255\255\255\255\255\255\139\001\140\001\255\255\255\255\143\001\ -\255\255\255\255\044\001\255\255\255\255\255\255\255\255\151\001\ -\050\001\153\001\255\255\255\255\255\255\157\001\158\001\159\001\ -\160\001\161\001\060\001\061\001\062\001\063\001\064\001\255\255\ -\255\255\255\255\255\255\255\255\255\255\255\255\255\255\073\001\ -\255\255\255\255\255\255\077\001\078\001\255\255\255\255\255\255\ -\255\255\255\255\084\001\085\001\255\255\255\255\255\255\255\255\ -\090\001\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ +\073\001\255\255\255\255\255\255\077\001\078\001\255\255\255\255\ +\255\255\255\255\255\255\084\001\085\001\255\255\255\255\255\255\ +\255\255\090\001\255\255\255\255\255\255\255\255\255\255\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\106\001\255\255\255\255\109\001\255\255\255\255\255\255\255\255\ -\012\001\255\255\255\255\255\255\016\001\255\255\255\255\019\001\ -\255\255\255\255\124\001\255\255\255\255\255\255\128\001\255\255\ -\255\255\255\255\132\001\255\255\255\255\255\255\136\001\255\255\ -\255\255\139\001\140\001\255\255\255\255\143\001\255\255\255\255\ -\044\001\255\255\255\255\255\255\255\255\151\001\050\001\153\001\ -\255\255\255\255\255\255\157\001\158\001\159\001\160\001\161\001\ -\060\001\061\001\062\001\063\001\064\001\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\073\001\255\255\255\255\ -\255\255\077\001\078\001\255\255\255\255\255\255\255\255\255\255\ -\084\001\085\001\255\255\255\255\255\255\255\255\090\001\255\255\ +\255\255\106\001\255\255\255\255\109\001\255\255\255\255\255\255\ +\255\255\012\001\255\255\255\255\255\255\016\001\255\255\255\255\ +\019\001\255\255\255\255\124\001\255\255\255\255\255\255\128\001\ +\255\255\255\255\255\255\132\001\255\255\255\255\255\255\136\001\ +\255\255\255\255\139\001\140\001\255\255\255\255\143\001\255\255\ +\255\255\044\001\255\255\255\255\255\255\255\255\151\001\050\001\ +\153\001\255\255\255\255\255\255\157\001\158\001\159\001\160\001\ +\161\001\060\001\061\001\062\001\063\001\064\001\255\255\255\255\ +\255\255\255\255\255\255\255\255\255\255\255\255\073\001\255\255\ +\255\255\255\255\077\001\078\001\255\255\255\255\255\255\255\255\ +\255\255\084\001\085\001\255\255\255\255\255\255\255\255\090\001\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\255\255\106\001\255\255\ -\255\255\109\001\255\255\255\255\255\255\255\255\012\001\255\255\ -\255\255\255\255\016\001\255\255\255\255\019\001\255\255\255\255\ -\124\001\255\255\255\255\255\255\128\001\255\255\255\255\255\255\ -\132\001\255\255\255\255\255\255\136\001\255\255\255\255\139\001\ -\140\001\255\255\255\255\143\001\255\255\255\255\044\001\255\255\ -\255\255\255\255\255\255\151\001\050\001\153\001\255\255\255\255\ -\255\255\157\001\158\001\159\001\160\001\161\001\060\001\061\001\ -\062\001\063\001\064\001\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\073\001\255\255\255\255\255\255\077\001\ -\078\001\255\255\255\255\255\255\255\255\255\255\084\001\085\001\ -\255\255\255\255\255\255\255\255\090\001\255\255\255\255\255\255\ +\255\255\255\255\255\255\255\255\255\255\255\255\255\255\106\001\ +\255\255\255\255\109\001\255\255\255\255\255\255\255\255\012\001\ +\255\255\255\255\255\255\016\001\255\255\255\255\019\001\255\255\ +\255\255\124\001\255\255\255\255\255\255\128\001\255\255\255\255\ +\255\255\132\001\255\255\255\255\255\255\255\255\255\255\255\255\ +\139\001\140\001\255\255\255\255\143\001\255\255\255\255\044\001\ +\255\255\255\255\255\255\255\255\151\001\050\001\153\001\255\255\ +\255\255\255\255\157\001\158\001\159\001\160\001\161\001\060\001\ +\061\001\062\001\063\001\064\001\255\255\255\255\255\255\255\255\ +\255\255\255\255\255\255\255\255\073\001\255\255\255\255\255\255\ +\077\001\078\001\255\255\255\255\255\255\255\255\255\255\084\001\ +\085\001\255\255\255\255\255\255\255\255\090\001\255\255\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\106\001\255\255\255\255\109\001\ -\255\255\255\255\255\255\255\255\012\001\255\255\255\255\255\255\ -\016\001\255\255\255\255\019\001\255\255\255\255\124\001\255\255\ -\255\255\255\255\128\001\255\255\255\255\255\255\132\001\255\255\ -\255\255\255\255\136\001\255\255\255\255\139\001\140\001\255\255\ -\255\255\143\001\255\255\255\255\044\001\255\255\255\255\255\255\ -\255\255\151\001\050\001\153\001\255\255\255\255\255\255\157\001\ -\158\001\159\001\160\001\161\001\060\001\061\001\062\001\063\001\ -\064\001\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\073\001\255\255\255\255\255\255\077\001\078\001\255\255\ -\255\255\255\255\255\255\255\255\084\001\085\001\255\255\255\255\ -\255\255\255\255\090\001\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\255\255\255\255\255\255\106\001\255\255\255\255\ +\109\001\255\255\255\255\255\255\255\255\012\001\255\255\255\255\ +\255\255\016\001\255\255\255\255\019\001\255\255\255\255\124\001\ +\255\255\255\255\255\255\128\001\255\255\255\255\255\255\132\001\ +\255\255\255\255\255\255\255\255\255\255\255\255\139\001\140\001\ +\255\255\255\255\143\001\255\255\255\255\044\001\255\255\255\255\ +\255\255\255\255\151\001\050\001\153\001\255\255\255\255\255\255\ +\157\001\158\001\159\001\160\001\161\001\060\001\061\001\062\001\ +\063\001\064\001\255\255\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\073\001\255\255\255\255\255\255\077\001\078\001\ +\255\255\255\255\255\255\255\255\255\255\084\001\085\001\255\255\ +\255\255\255\255\255\255\090\001\255\255\255\255\255\255\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\106\001\255\255\255\255\109\001\255\255\255\255\ -\255\255\255\255\012\001\255\255\255\255\255\255\016\001\255\255\ -\255\255\019\001\255\255\255\255\124\001\255\255\255\255\255\255\ -\128\001\255\255\255\255\255\255\132\001\255\255\255\255\255\255\ -\136\001\255\255\255\255\139\001\140\001\255\255\255\255\143\001\ -\255\255\255\255\044\001\255\255\255\255\255\255\255\255\151\001\ -\050\001\153\001\255\255\255\255\255\255\157\001\158\001\159\001\ -\160\001\161\001\060\001\061\001\062\001\063\001\064\001\255\255\ -\255\255\255\255\255\255\255\255\255\255\255\255\255\255\073\001\ -\255\255\255\255\255\255\077\001\078\001\255\255\255\255\255\255\ -\255\255\255\255\084\001\085\001\255\255\255\255\255\255\255\255\ -\090\001\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\255\255\106\001\255\255\020\001\109\001\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\106\001\255\255\255\255\109\001\255\255\255\255\255\255\255\255\ -\012\001\255\255\255\255\255\255\016\001\255\255\255\255\019\001\ -\255\255\255\255\124\001\255\255\255\255\255\255\128\001\255\255\ -\255\255\255\255\132\001\255\255\255\255\255\255\255\255\255\255\ -\255\255\139\001\140\001\255\255\255\255\143\001\255\255\255\255\ -\044\001\255\255\255\255\255\255\255\255\151\001\050\001\153\001\ -\255\255\255\255\255\255\157\001\158\001\159\001\160\001\161\001\ -\060\001\061\001\062\001\063\001\064\001\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\073\001\255\255\255\255\ -\255\255\077\001\078\001\255\255\255\255\255\255\255\255\255\255\ -\084\001\085\001\255\255\255\255\255\255\255\255\090\001\255\255\ +\255\255\255\255\255\255\255\255\035\001\124\001\255\255\255\255\ +\255\255\128\001\255\255\042\001\255\255\132\001\255\255\046\001\ +\255\255\255\255\255\255\255\255\139\001\140\001\255\255\255\255\ +\143\001\056\001\255\255\255\255\059\001\255\255\255\255\255\255\ +\151\001\255\255\153\001\255\255\255\255\068\001\157\001\158\001\ +\159\001\160\001\161\001\255\255\255\255\255\255\255\255\255\255\ +\255\255\080\001\255\255\003\001\255\255\255\255\255\255\255\255\ +\255\255\255\255\089\001\255\255\255\255\092\001\255\255\255\255\ +\255\255\255\255\097\001\255\255\255\255\255\255\022\001\023\001\ +\255\255\025\001\255\255\255\255\028\001\255\255\255\255\003\001\ +\111\001\112\001\113\001\114\001\115\001\116\001\117\001\118\001\ +\040\001\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\022\001\023\001\052\001\025\001\054\001\255\255\ +\028\001\255\255\255\255\255\255\255\255\255\255\141\001\255\255\ +\255\255\255\255\255\255\146\001\040\001\255\255\255\255\255\255\ +\255\255\255\255\255\255\154\001\255\255\255\255\255\255\255\255\ +\052\001\081\001\054\001\255\255\255\255\255\255\255\255\166\001\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\255\255\106\001\255\255\ -\255\255\109\001\255\255\255\255\255\255\255\255\012\001\255\255\ -\255\255\255\255\016\001\255\255\255\255\019\001\255\255\255\255\ -\124\001\255\255\255\255\255\255\128\001\255\255\255\255\255\255\ -\132\001\255\255\255\255\255\255\255\255\255\255\255\255\139\001\ -\140\001\255\255\255\255\143\001\255\255\255\255\044\001\255\255\ -\255\255\255\255\255\255\151\001\050\001\153\001\255\255\255\255\ -\255\255\157\001\158\001\159\001\160\001\161\001\060\001\061\001\ -\062\001\063\001\064\001\255\255\255\255\003\001\255\255\255\255\ -\255\255\255\255\255\255\073\001\255\255\255\255\255\255\077\001\ -\078\001\255\255\255\255\255\255\255\255\255\255\084\001\085\001\ -\022\001\023\001\255\255\025\001\090\001\255\255\028\001\255\255\ -\255\255\255\255\255\255\255\255\255\255\255\255\255\255\003\001\ -\255\255\255\255\040\001\255\255\106\001\255\255\255\255\109\001\ -\255\255\255\255\255\255\255\255\255\255\255\255\052\001\255\255\ -\054\001\255\255\022\001\023\001\255\255\025\001\124\001\255\255\ -\028\001\255\255\128\001\255\255\255\255\255\255\132\001\255\255\ -\255\255\255\255\255\255\255\255\040\001\139\001\140\001\255\255\ -\255\255\143\001\255\255\081\001\255\255\255\255\255\255\255\255\ -\052\001\151\001\054\001\153\001\255\255\255\255\255\255\157\001\ -\158\001\159\001\160\001\161\001\098\001\099\001\100\001\101\001\ -\102\001\103\001\104\001\105\001\106\001\107\001\108\001\255\255\ -\110\001\255\255\255\255\255\255\255\255\081\001\255\255\255\255\ +\255\255\255\255\098\001\099\001\100\001\101\001\102\001\103\001\ +\104\001\105\001\106\001\107\001\108\001\081\001\110\001\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ \003\001\255\255\255\255\255\255\255\255\255\255\098\001\099\001\ \100\001\101\001\102\001\103\001\104\001\105\001\106\001\107\001\ -\108\001\255\255\110\001\022\001\023\001\255\255\025\001\149\001\ -\255\255\028\001\255\255\255\255\003\001\255\255\255\255\255\255\ +\108\001\255\255\110\001\022\001\023\001\255\255\025\001\255\255\ +\255\255\028\001\255\255\255\255\003\001\149\001\255\255\255\255\ \255\255\255\255\255\255\255\255\255\255\040\001\255\255\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\022\001\ \023\001\052\001\025\001\054\001\255\255\028\001\255\255\255\255\ @@ -6500,25 +6520,21 @@ let yycheck = "\002\000\ \255\255\255\255\255\255\098\001\099\001\100\001\101\001\102\001\ \103\001\104\001\105\001\106\001\107\001\108\001\255\255\110\001\ \022\001\023\001\255\255\025\001\255\255\255\255\028\001\255\255\ -\255\255\003\001\149\001\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\149\001\255\255\255\255\255\255\255\255\255\255\ \255\255\255\255\040\001\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\022\001\023\001\052\001\025\001\ -\054\001\255\255\028\001\255\255\255\255\255\255\149\001\255\255\ -\255\255\255\255\255\255\255\255\255\255\255\255\040\001\255\255\ +\255\255\255\255\255\255\255\255\255\255\255\255\052\001\255\255\ +\054\001\255\255\255\255\255\255\255\255\255\255\149\001\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\052\001\081\001\054\001\255\255\255\255\255\255\ +\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ +\255\255\255\255\255\255\081\001\255\255\255\255\255\255\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ \255\255\255\255\255\255\255\255\098\001\099\001\100\001\101\001\ -\102\001\103\001\104\001\105\001\106\001\107\001\108\001\081\001\ +\102\001\103\001\104\001\105\001\106\001\107\001\108\001\255\255\ \110\001\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\098\001\099\001\100\001\101\001\102\001\103\001\104\001\105\001\ -\106\001\107\001\108\001\255\255\110\001\255\255\255\255\255\255\ -\255\255\255\255\255\255\255\255\255\255\255\255\255\255\149\001\ -\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ \255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\255\ -\255\255\255\255\255\255\149\001" +\255\255\255\255\255\255\255\255\255\255\255\255\255\255\149\001" let yynames_const = "\ AMP\000\ @@ -6699,1008 +6715,1023 @@ let yyact = [| (fun _ -> failwith "parser") ; (fun __caml_parser_env -> Obj.repr( -# 238 "parse.mly" +# 240 "parse.mly" ( ( None )) -# 6691 "parse.ml" +# 6707 "parse.ml" : 'option___anonymous_0_)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 0 : 'typ) in Obj.repr( -# 240 "parse.mly" +# 242 "parse.mly" (let (_1, t) = ((), _2) in let x = (t) in ( Some x )) -# 6700 "parse.ml" +# 6716 "parse.ml" : 'option___anonymous_0_)) ; (fun __caml_parser_env -> Obj.repr( -# 246 "parse.mly" +# 248 "parse.mly" ( ( None )) -# 6706 "parse.ml" +# 6722 "parse.ml" : 'option___anonymous_1_)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 0 : 'typ) in Obj.repr( -# 248 "parse.mly" +# 250 "parse.mly" (let (_1, t) = ((), _2) in let x = (t) in ( Some x )) -# 6715 "parse.ml" +# 6731 "parse.ml" : 'option___anonymous_1_)) ; (fun __caml_parser_env -> Obj.repr( -# 254 "parse.mly" +# 256 "parse.mly" ( ( None )) -# 6721 "parse.ml" +# 6737 "parse.ml" : 'option___anonymous_12_)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 0 : 'thunk2_typ_) in Obj.repr( -# 256 "parse.mly" +# 258 "parse.mly" (let (_1, tactic) = ((), _2) in let x = (tactic) in ( Some x )) -# 6730 "parse.ml" +# 6746 "parse.ml" : 'option___anonymous_12_)) ; (fun __caml_parser_env -> Obj.repr( -# 262 "parse.mly" +# 264 "parse.mly" ( ( None )) -# 6736 "parse.ml" +# 6752 "parse.ml" : 'option___anonymous_13_)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 1 : 'noSeqTerm) in Obj.repr( -# 264 "parse.mly" +# 266 "parse.mly" (let (_1, e, _3) = ((), _2, ()) in let x = let phi = ( {e with level=Formula} ) in (phi) in ( Some x )) -# 6748 "parse.ml" +# 6764 "parse.ml" : 'option___anonymous_13_)) ; (fun __caml_parser_env -> Obj.repr( -# 273 "parse.mly" +# 275 "parse.mly" ( ( None )) -# 6754 "parse.ml" +# 6770 "parse.ml" : 'option___anonymous_2_)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 0 : FStar_Parser_AST.term) in Obj.repr( -# 275 "parse.mly" +# 277 "parse.mly" (let (_1, tm) = ((), _2) in let x = (tm) in ( Some x )) -# 6763 "parse.ml" +# 6779 "parse.ml" : 'option___anonymous_2_)) ; (fun __caml_parser_env -> Obj.repr( -# 281 "parse.mly" +# 283 "parse.mly" ( ( None )) -# 6769 "parse.ml" +# 6785 "parse.ml" : 'option___anonymous_5_)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 0 : 'thunk_atomicTerm_) in Obj.repr( -# 283 "parse.mly" +# 285 "parse.mly" (let (_1, tactic) = ((), _2) in let x = (tactic) in ( Some x )) -# 6778 "parse.ml" +# 6794 "parse.ml" : 'option___anonymous_5_)) ; (fun __caml_parser_env -> Obj.repr( -# 289 "parse.mly" +# 291 "parse.mly" ( ( None )) -# 6784 "parse.ml" +# 6800 "parse.ml" : 'option___anonymous_6_)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 0 : FStar_Ident.ident) in Obj.repr( -# 291 "parse.mly" +# 293 "parse.mly" (let (_1, i) = ((), _2) in let x = (i) in ( Some x )) -# 6793 "parse.ml" +# 6809 "parse.ml" : 'option___anonymous_6_)) ; (fun __caml_parser_env -> Obj.repr( -# 297 "parse.mly" +# 299 "parse.mly" ( ( None )) -# 6799 "parse.ml" +# 6815 "parse.ml" : 'option___anonymous_7_)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 0 : FStar_Ident.ident) in Obj.repr( -# 299 "parse.mly" +# 301 "parse.mly" (let (_1, i) = ((), _2) in let x = (i) in ( Some x )) -# 6808 "parse.ml" +# 6824 "parse.ml" : 'option___anonymous_7_)) ; (fun __caml_parser_env -> Obj.repr( -# 305 "parse.mly" +# 307 "parse.mly" ( ( None )) -# 6814 "parse.ml" +# 6830 "parse.ml" : 'option___anonymous_8_)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 0 : 'thunk_typ_) in Obj.repr( -# 307 "parse.mly" +# 309 "parse.mly" (let (_1, tactic) = ((), _2) in let x = (tactic) in ( Some x )) -# 6823 "parse.ml" +# 6839 "parse.ml" : 'option___anonymous_8_)) ; (fun __caml_parser_env -> Obj.repr( -# 313 "parse.mly" +# 315 "parse.mly" ( ( None )) -# 6829 "parse.ml" +# 6845 "parse.ml" : 'option___anonymous_9_)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 0 : 'thunk_typ_) in Obj.repr( -# 315 "parse.mly" +# 317 "parse.mly" (let (_1, tactic) = ((), _2) in let x = (tactic) in ( Some x )) -# 6838 "parse.ml" +# 6854 "parse.ml" : 'option___anonymous_9_)) ; (fun __caml_parser_env -> Obj.repr( -# 321 "parse.mly" +# 323 "parse.mly" ( ( None )) -# 6844 "parse.ml" +# 6860 "parse.ml" : 'option_ascribeKind_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'ascribeKind) in Obj.repr( -# 323 "parse.mly" +# 325 "parse.mly" (let x = _1 in ( Some x )) -# 6852 "parse.ml" +# 6868 "parse.ml" : 'option_ascribeKind_)) ; (fun __caml_parser_env -> Obj.repr( -# 328 "parse.mly" +# 330 "parse.mly" ( ( None )) -# 6858 "parse.ml" +# 6874 "parse.ml" : 'option_ascribeTyp_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'ascribeTyp) in Obj.repr( -# 330 "parse.mly" +# 332 "parse.mly" (let x = _1 in ( Some x )) -# 6866 "parse.ml" +# 6882 "parse.ml" : 'option_ascribeTyp_)) ; (fun __caml_parser_env -> Obj.repr( -# 335 "parse.mly" +# 337 "parse.mly" ( ( None )) -# 6872 "parse.ml" +# 6888 "parse.ml" : 'option_constructorPayload_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'constructorPayload) in Obj.repr( -# 337 "parse.mly" +# 339 "parse.mly" (let x = _1 in ( Some x )) -# 6880 "parse.ml" +# 6896 "parse.ml" : 'option_constructorPayload_)) ; (fun __caml_parser_env -> Obj.repr( -# 342 "parse.mly" +# 344 "parse.mly" ( ( None )) -# 6886 "parse.ml" +# 6902 "parse.ml" : 'option_fsTypeArgs_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'fsTypeArgs) in Obj.repr( -# 344 "parse.mly" +# 346 "parse.mly" (let x = _1 in ( Some x )) -# 6894 "parse.ml" +# 6910 "parse.ml" : 'option_fsTypeArgs_)) ; (fun __caml_parser_env -> Obj.repr( -# 349 "parse.mly" +# 351 "parse.mly" ( ( None )) -# 6900 "parse.ml" +# 6916 "parse.ml" : 'option_match_returning_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'match_returning) in Obj.repr( -# 351 "parse.mly" +# 353 "parse.mly" (let x = _1 in ( Some x )) -# 6908 "parse.ml" +# 6924 "parse.ml" : 'option_match_returning_)) ; (fun __caml_parser_env -> Obj.repr( -# 356 "parse.mly" +# 358 "parse.mly" ( ( None )) -# 6914 "parse.ml" +# 6930 "parse.ml" : 'option_pair_hasSort_simpleTerm__)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 1 : 'hasSort) in let _2 = (Parsing.peek_val __caml_parser_env 0 : 'simpleTerm) in Obj.repr( -# 358 "parse.mly" +# 360 "parse.mly" (let (x, y) = (_1, _2) in let x = ( (x, y) ) in ( Some x )) -# 6924 "parse.ml" +# 6940 "parse.ml" : 'option_pair_hasSort_simpleTerm__)) ; (fun __caml_parser_env -> Obj.repr( -# 364 "parse.mly" +# 366 "parse.mly" ( ( None )) -# 6930 "parse.ml" +# 6946 "parse.ml" : 'option_string_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'string) in Obj.repr( -# 366 "parse.mly" +# 368 "parse.mly" (let x = _1 in ( Some x )) -# 6938 "parse.ml" +# 6954 "parse.ml" : 'option_string_)) ; (fun __caml_parser_env -> Obj.repr( -# 371 "parse.mly" +# 373 "parse.mly" ( ( None )) -# 6944 "parse.ml" +# 6960 "parse.ml" : 'option_term_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : FStar_Parser_AST.term) in Obj.repr( -# 373 "parse.mly" +# 375 "parse.mly" (let x = _1 in ( Some x )) -# 6952 "parse.ml" +# 6968 "parse.ml" : 'option_term_)) ; (fun __caml_parser_env -> Obj.repr( -# 378 "parse.mly" +# 380 "parse.mly" ( ( false )) -# 6958 "parse.ml" +# 6974 "parse.ml" : 'boption_SQUIGGLY_RARROW_)) ; (fun __caml_parser_env -> Obj.repr( -# 380 "parse.mly" +# 382 "parse.mly" (let _1 = () in ( true )) -# 6965 "parse.ml" +# 6981 "parse.ml" : 'boption_SQUIGGLY_RARROW_)) ; (fun __caml_parser_env -> Obj.repr( -# 385 "parse.mly" +# 387 "parse.mly" ( ( [] )) -# 6971 "parse.ml" +# 6987 "parse.ml" : 'loption_separated_nonempty_list_COMMA_appTerm__)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'separated_nonempty_list_COMMA_appTerm_) in Obj.repr( -# 387 "parse.mly" +# 389 "parse.mly" (let x = _1 in ( x )) -# 6979 "parse.ml" +# 6995 "parse.ml" : 'loption_separated_nonempty_list_COMMA_appTerm__)) ; (fun __caml_parser_env -> Obj.repr( -# 392 "parse.mly" +# 394 "parse.mly" ( ( [] )) -# 6985 "parse.ml" +# 7001 "parse.ml" : 'loption_separated_nonempty_list_SEMICOLON_ident__)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'separated_nonempty_list_SEMICOLON_ident_) in Obj.repr( -# 394 "parse.mly" +# 396 "parse.mly" (let x = _1 in ( x )) -# 6993 "parse.ml" +# 7009 "parse.ml" : 'loption_separated_nonempty_list_SEMICOLON_ident__)) ; (fun __caml_parser_env -> Obj.repr( -# 399 "parse.mly" +# 401 "parse.mly" ( ( [] )) -# 6999 "parse.ml" +# 7015 "parse.ml" : 'loption_separated_nonempty_list_SEMICOLON_tuplePattern__)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'separated_nonempty_list_SEMICOLON_tuplePattern_) in Obj.repr( -# 401 "parse.mly" +# 403 "parse.mly" (let x = _1 in ( x )) -# 7007 "parse.ml" +# 7023 "parse.ml" : 'loption_separated_nonempty_list_SEMICOLON_tuplePattern__)) ; (fun __caml_parser_env -> Obj.repr( -# 406 "parse.mly" +# 408 "parse.mly" ( ( [] )) -# 7013 "parse.ml" +# 7029 "parse.ml" : 'list___anonymous_11_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 2 : string) in let _2 = (Parsing.peek_val __caml_parser_env 1 : 'letoperatorbinding) in let _3 = (Parsing.peek_val __caml_parser_env 0 : 'list___anonymous_11_) in Obj.repr( -# 408 "parse.mly" +# 410 "parse.mly" (let (op, b, xs) = (_1, _2, _3) in let x = let op = ( mk_ident ("and" ^ op, rhs parseState 1) ) in ((op, b)) in ( x :: xs )) -# 7027 "parse.ml" +# 7043 "parse.ml" : 'list___anonymous_11_)) ; (fun __caml_parser_env -> Obj.repr( -# 417 "parse.mly" +# 419 "parse.mly" ( ( [] )) -# 7033 "parse.ml" +# 7049 "parse.ml" : 'list___anonymous_14_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 1 : 'argTerm) in let _2 = (Parsing.peek_val __caml_parser_env 0 : 'list___anonymous_14_) in Obj.repr( -# 419 "parse.mly" +# 421 "parse.mly" (let (t, xs) = (_1, _2) in let x = (t) in ( x :: xs )) -# 7043 "parse.ml" +# 7059 "parse.ml" : 'list___anonymous_14_)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 2 : 'recordExp) in let _4 = (Parsing.peek_val __caml_parser_env 0 : 'list___anonymous_14_) in Obj.repr( -# 423 "parse.mly" +# 425 "parse.mly" (let (_2, t, _4, xs) = ((), _2, (), _4) in let x = let h = ( Nothing ) in (h, t) in ( x :: xs )) -# 7056 "parse.ml" +# 7072 "parse.ml" : 'list___anonymous_14_)) ; (fun __caml_parser_env -> let _3 = (Parsing.peek_val __caml_parser_env 2 : 'recordExp) in let _5 = (Parsing.peek_val __caml_parser_env 0 : 'list___anonymous_14_) in Obj.repr( -# 430 "parse.mly" +# 432 "parse.mly" (let (_1, _2, t, _4, xs) = ((), (), _3, (), _5) in let x = let h = ( Hash ) in (h, t) in ( x :: xs )) -# 7069 "parse.ml" +# 7085 "parse.ml" : 'list___anonymous_14_)) ; (fun __caml_parser_env -> Obj.repr( -# 439 "parse.mly" +# 441 "parse.mly" ( ( [] )) -# 7075 "parse.ml" +# 7091 "parse.ml" : 'list___anonymous_15_)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 1 : 'qlident) in let _3 = (Parsing.peek_val __caml_parser_env 0 : 'list___anonymous_15_) in Obj.repr( -# 441 "parse.mly" +# 443 "parse.mly" (let (_1, id, xs) = ((), _2, _3) in let x = (id) in ( x :: xs )) -# 7085 "parse.ml" +# 7101 "parse.ml" : 'list___anonymous_15_)) ; (fun __caml_parser_env -> Obj.repr( -# 447 "parse.mly" +# 449 "parse.mly" ( ( [] )) -# 7091 "parse.ml" +# 7107 "parse.ml" : 'list___anonymous_4_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 1 : 'binder) in let _2 = (Parsing.peek_val __caml_parser_env 0 : 'list___anonymous_4_) in Obj.repr( -# 449 "parse.mly" +# 451 "parse.mly" (let (b, xs) = (_1, _2) in let x = ([b]) in ( x :: xs )) -# 7101 "parse.ml" +# 7117 "parse.ml" : 'list___anonymous_4_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 1 : 'multiBinder) in let _2 = (Parsing.peek_val __caml_parser_env 0 : 'list___anonymous_4_) in Obj.repr( -# 453 "parse.mly" +# 455 "parse.mly" (let (bs, xs) = (_1, _2) in let x = (bs) in ( x :: xs )) -# 7111 "parse.ml" +# 7127 "parse.ml" : 'list___anonymous_4_)) ; (fun __caml_parser_env -> Obj.repr( -# 459 "parse.mly" +# 461 "parse.mly" ( ( [] )) -# 7117 "parse.ml" +# 7133 "parse.ml" : 'list_argTerm_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 1 : 'argTerm) in let _2 = (Parsing.peek_val __caml_parser_env 0 : 'list_argTerm_) in Obj.repr( -# 461 "parse.mly" +# 463 "parse.mly" (let (x, xs) = (_1, _2) in ( x :: xs )) -# 7126 "parse.ml" +# 7142 "parse.ml" : 'list_argTerm_)) ; (fun __caml_parser_env -> Obj.repr( -# 466 "parse.mly" +# 468 "parse.mly" ( ( [] )) -# 7132 "parse.ml" +# 7148 "parse.ml" : 'list_atomicTerm_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 1 : 'atomicTerm) in let _2 = (Parsing.peek_val __caml_parser_env 0 : 'list_atomicTerm_) in Obj.repr( -# 468 "parse.mly" +# 470 "parse.mly" (let (x, xs) = (_1, _2) in ( x :: xs )) -# 7141 "parse.ml" +# 7157 "parse.ml" : 'list_atomicTerm_)) ; (fun __caml_parser_env -> Obj.repr( -# 473 "parse.mly" +# 475 "parse.mly" ( ( [] )) -# 7147 "parse.ml" +# 7163 "parse.ml" : 'list_attr_letbinding_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 1 : 'attr_letbinding) in let _2 = (Parsing.peek_val __caml_parser_env 0 : 'list_attr_letbinding_) in Obj.repr( -# 475 "parse.mly" +# 477 "parse.mly" (let (x, xs) = (_1, _2) in ( x :: xs )) -# 7156 "parse.ml" +# 7172 "parse.ml" : 'list_attr_letbinding_)) ; (fun __caml_parser_env -> Obj.repr( -# 480 "parse.mly" +# 482 "parse.mly" ( ( [] )) -# 7162 "parse.ml" +# 7178 "parse.ml" : 'list_calcStep_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 1 : 'calcStep) in let _2 = (Parsing.peek_val __caml_parser_env 0 : 'list_calcStep_) in Obj.repr( -# 482 "parse.mly" +# 484 "parse.mly" (let (x, xs) = (_1, _2) in ( x :: xs )) -# 7171 "parse.ml" +# 7187 "parse.ml" : 'list_calcStep_)) ; (fun __caml_parser_env -> Obj.repr( -# 487 "parse.mly" +# 489 "parse.mly" ( ( [] )) -# 7177 "parse.ml" +# 7193 "parse.ml" : 'list_constructorDecl_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 1 : 'constructorDecl) in let _2 = (Parsing.peek_val __caml_parser_env 0 : 'list_constructorDecl_) in Obj.repr( -# 489 "parse.mly" +# 491 "parse.mly" (let (x, xs) = (_1, _2) in ( x :: xs )) -# 7186 "parse.ml" +# 7202 "parse.ml" : 'list_constructorDecl_)) ; (fun __caml_parser_env -> Obj.repr( -# 494 "parse.mly" +# 496 "parse.mly" ( ( [] )) -# 7192 "parse.ml" +# 7208 "parse.ml" : 'list_decl_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 1 : 'decl) in let _2 = (Parsing.peek_val __caml_parser_env 0 : 'list_decl_) in Obj.repr( -# 496 "parse.mly" +# 498 "parse.mly" (let (x, xs) = (_1, _2) in ( x :: xs )) -# 7201 "parse.ml" +# 7217 "parse.ml" : 'list_decl_)) ; (fun __caml_parser_env -> Obj.repr( -# 501 "parse.mly" +# 503 "parse.mly" ( ( [] )) -# 7207 "parse.ml" +# 7223 "parse.ml" : 'list_decoration_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 1 : 'decoration) in let _2 = (Parsing.peek_val __caml_parser_env 0 : 'list_decoration_) in Obj.repr( -# 503 "parse.mly" +# 505 "parse.mly" (let (x, xs) = (_1, _2) in ( x :: xs )) -# 7216 "parse.ml" +# 7232 "parse.ml" : 'list_decoration_)) ; (fun __caml_parser_env -> Obj.repr( -# 508 "parse.mly" +# 510 "parse.mly" ( ( [] )) -# 7222 "parse.ml" +# 7238 "parse.ml" : 'list_multiBinder_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 1 : 'multiBinder) in let _2 = (Parsing.peek_val __caml_parser_env 0 : 'list_multiBinder_) in Obj.repr( -# 510 "parse.mly" +# 512 "parse.mly" (let (x, xs) = (_1, _2) in ( x :: xs )) -# 7231 "parse.ml" +# 7247 "parse.ml" : 'list_multiBinder_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'aqualifiedWithAttrs_lident_) in Obj.repr( -# 515 "parse.mly" +# 517 "parse.mly" (let x = _1 in ( [ x ] )) -# 7239 "parse.ml" +# 7255 "parse.ml" : 'nonempty_list_aqualifiedWithAttrs_lident__)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 1 : 'aqualifiedWithAttrs_lident_) in let _2 = (Parsing.peek_val __caml_parser_env 0 : 'nonempty_list_aqualifiedWithAttrs_lident__) in Obj.repr( -# 518 "parse.mly" +# 520 "parse.mly" (let (x, xs) = (_1, _2) in ( x :: xs )) -# 7248 "parse.ml" +# 7264 "parse.ml" : 'nonempty_list_aqualifiedWithAttrs_lident__)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'aqualifiedWithAttrs_lidentOrUnderscore_) in Obj.repr( -# 523 "parse.mly" +# 525 "parse.mly" (let x = _1 in ( [ x ] )) -# 7256 "parse.ml" +# 7272 "parse.ml" : 'nonempty_list_aqualifiedWithAttrs_lidentOrUnderscore__)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 1 : 'aqualifiedWithAttrs_lidentOrUnderscore_) in let _2 = (Parsing.peek_val __caml_parser_env 0 : 'nonempty_list_aqualifiedWithAttrs_lidentOrUnderscore__) in Obj.repr( -# 526 "parse.mly" +# 528 "parse.mly" (let (x, xs) = (_1, _2) in ( x :: xs )) -# 7265 "parse.ml" +# 7281 "parse.ml" : 'nonempty_list_aqualifiedWithAttrs_lidentOrUnderscore__)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'atomicPattern) in Obj.repr( -# 531 "parse.mly" +# 533 "parse.mly" (let x = _1 in ( [ x ] )) -# 7273 "parse.ml" +# 7289 "parse.ml" : 'nonempty_list_atomicPattern_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 1 : 'atomicPattern) in let _2 = (Parsing.peek_val __caml_parser_env 0 : 'nonempty_list_atomicPattern_) in Obj.repr( -# 534 "parse.mly" +# 536 "parse.mly" (let (x, xs) = (_1, _2) in ( x :: xs )) -# 7282 "parse.ml" +# 7298 "parse.ml" : 'nonempty_list_atomicPattern_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'atomicTerm) in Obj.repr( -# 539 "parse.mly" +# 541 "parse.mly" (let x = _1 in ( [ x ] )) -# 7290 "parse.ml" +# 7306 "parse.ml" : 'nonempty_list_atomicTerm_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 1 : 'atomicTerm) in let _2 = (Parsing.peek_val __caml_parser_env 0 : 'nonempty_list_atomicTerm_) in Obj.repr( -# 542 "parse.mly" +# 544 "parse.mly" (let (x, xs) = (_1, _2) in ( x :: xs )) -# 7299 "parse.ml" +# 7315 "parse.ml" : 'nonempty_list_atomicTerm_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'atomicUniverse) in Obj.repr( -# 547 "parse.mly" +# 549 "parse.mly" (let x = _1 in ( [ x ] )) -# 7307 "parse.ml" +# 7323 "parse.ml" : 'nonempty_list_atomicUniverse_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 1 : 'atomicUniverse) in let _2 = (Parsing.peek_val __caml_parser_env 0 : 'nonempty_list_atomicUniverse_) in Obj.repr( -# 550 "parse.mly" +# 552 "parse.mly" (let (x, xs) = (_1, _2) in ( x :: xs )) -# 7316 "parse.ml" +# 7332 "parse.ml" : 'nonempty_list_atomicUniverse_)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 1 : FStar_Parser_AST.term) in Obj.repr( -# 555 "parse.mly" +# 557 "parse.mly" (let (_1, e, _3) = ((), _2, ()) in let x = ( mk_ident (".()", rhs parseState 1), e, rhs2 parseState 1 3 ) in ( [ x ] )) -# 7325 "parse.ml" +# 7341 "parse.ml" : 'nonempty_list_dotOperator_)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 1 : FStar_Parser_AST.term) in Obj.repr( -# 559 "parse.mly" +# 561 "parse.mly" (let (_1, e, _3) = ((), _2, ()) in let x = ( mk_ident (".[]", rhs parseState 1), e, rhs2 parseState 1 3 ) in ( [ x ] )) -# 7334 "parse.ml" +# 7350 "parse.ml" : 'nonempty_list_dotOperator_)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 1 : FStar_Parser_AST.term) in Obj.repr( -# 563 "parse.mly" +# 565 "parse.mly" (let (_1, e, _3) = ((), _2, ()) in let x = ( mk_ident (".[||]", rhs parseState 1), e, rhs2 parseState 1 3 ) in ( [ x ] )) -# 7343 "parse.ml" +# 7359 "parse.ml" : 'nonempty_list_dotOperator_)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 1 : FStar_Parser_AST.term) in Obj.repr( -# 567 "parse.mly" +# 569 "parse.mly" (let (_1, e, _3) = ((), _2, ()) in let x = ( mk_ident (".(||)", rhs parseState 1), e, rhs2 parseState 1 3 ) in ( [ x ] )) -# 7352 "parse.ml" +# 7368 "parse.ml" : 'nonempty_list_dotOperator_)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 2 : FStar_Parser_AST.term) in let _4 = (Parsing.peek_val __caml_parser_env 0 : 'nonempty_list_dotOperator_) in Obj.repr( -# 571 "parse.mly" +# 573 "parse.mly" (let (_1, e, _3, xs) = ((), _2, (), _4) in let x = ( mk_ident (".()", rhs parseState 1), e, rhs2 parseState 1 3 ) in ( x :: xs )) -# 7362 "parse.ml" +# 7378 "parse.ml" : 'nonempty_list_dotOperator_)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 2 : FStar_Parser_AST.term) in let _4 = (Parsing.peek_val __caml_parser_env 0 : 'nonempty_list_dotOperator_) in Obj.repr( -# 575 "parse.mly" +# 577 "parse.mly" (let (_1, e, _3, xs) = ((), _2, (), _4) in let x = ( mk_ident (".[]", rhs parseState 1), e, rhs2 parseState 1 3 ) in ( x :: xs )) -# 7372 "parse.ml" +# 7388 "parse.ml" : 'nonempty_list_dotOperator_)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 2 : FStar_Parser_AST.term) in let _4 = (Parsing.peek_val __caml_parser_env 0 : 'nonempty_list_dotOperator_) in Obj.repr( -# 579 "parse.mly" +# 581 "parse.mly" (let (_1, e, _3, xs) = ((), _2, (), _4) in let x = ( mk_ident (".[||]", rhs parseState 1), e, rhs2 parseState 1 3 ) in ( x :: xs )) -# 7382 "parse.ml" +# 7398 "parse.ml" : 'nonempty_list_dotOperator_)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 2 : FStar_Parser_AST.term) in let _4 = (Parsing.peek_val __caml_parser_env 0 : 'nonempty_list_dotOperator_) in Obj.repr( -# 583 "parse.mly" +# 585 "parse.mly" (let (_1, e, _3, xs) = ((), _2, (), _4) in let x = ( mk_ident (".(||)", rhs parseState 1), e, rhs2 parseState 1 3 ) in ( x :: xs )) -# 7392 "parse.ml" +# 7408 "parse.ml" : 'nonempty_list_dotOperator_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'patternOrMultibinder) in Obj.repr( -# 589 "parse.mly" +# 591 "parse.mly" (let x = _1 in ( [ x ] )) -# 7400 "parse.ml" +# 7416 "parse.ml" : 'nonempty_list_patternOrMultibinder_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 1 : 'patternOrMultibinder) in let _2 = (Parsing.peek_val __caml_parser_env 0 : 'nonempty_list_patternOrMultibinder_) in Obj.repr( -# 592 "parse.mly" +# 594 "parse.mly" (let (x, xs) = (_1, _2) in ( x :: xs )) -# 7409 "parse.ml" +# 7425 "parse.ml" : 'nonempty_list_patternOrMultibinder_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'letbinding) in Obj.repr( -# 597 "parse.mly" +# 599 "parse.mly" (let x = _1 in ( [ x ] )) -# 7417 "parse.ml" +# 7433 "parse.ml" : 'separated_nonempty_list_AND_letbinding_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 2 : 'letbinding) in let _3 = (Parsing.peek_val __caml_parser_env 0 : 'separated_nonempty_list_AND_letbinding_) in Obj.repr( -# 600 "parse.mly" +# 602 "parse.mly" (let (x, _2, xs) = (_1, (), _3) in ( x :: xs )) -# 7426 "parse.ml" +# 7442 "parse.ml" : 'separated_nonempty_list_AND_letbinding_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'typeDecl) in Obj.repr( -# 605 "parse.mly" +# 607 "parse.mly" (let x = _1 in ( [ x ] )) -# 7434 "parse.ml" +# 7450 "parse.ml" : 'separated_nonempty_list_AND_typeDecl_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 2 : 'typeDecl) in let _3 = (Parsing.peek_val __caml_parser_env 0 : 'separated_nonempty_list_AND_typeDecl_) in Obj.repr( -# 608 "parse.mly" +# 610 "parse.mly" (let (x, _2, xs) = (_1, (), _3) in ( x :: xs )) -# 7443 "parse.ml" +# 7459 "parse.ml" : 'separated_nonempty_list_AND_typeDecl_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'tuplePattern) in Obj.repr( -# 613 "parse.mly" +# 615 "parse.mly" (let x = _1 in ( [ x ] )) -# 7451 "parse.ml" +# 7467 "parse.ml" : 'separated_nonempty_list_BAR_tuplePattern_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 2 : 'tuplePattern) in let _3 = (Parsing.peek_val __caml_parser_env 0 : 'separated_nonempty_list_BAR_tuplePattern_) in Obj.repr( -# 616 "parse.mly" +# 618 "parse.mly" (let (x, _2, xs) = (_1, (), _3) in ( x :: xs )) -# 7460 "parse.ml" +# 7476 "parse.ml" : 'separated_nonempty_list_BAR_tuplePattern_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'appTerm) in Obj.repr( -# 621 "parse.mly" +# 623 "parse.mly" (let x = _1 in ( [ x ] )) -# 7468 "parse.ml" +# 7484 "parse.ml" : 'separated_nonempty_list_COMMA_appTerm_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 2 : 'appTerm) in let _3 = (Parsing.peek_val __caml_parser_env 0 : 'separated_nonempty_list_COMMA_appTerm_) in Obj.repr( -# 624 "parse.mly" +# 626 "parse.mly" (let (x, _2, xs) = (_1, (), _3) in ( x :: xs )) -# 7477 "parse.ml" +# 7493 "parse.ml" : 'separated_nonempty_list_COMMA_appTerm_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'atomicTerm) in Obj.repr( -# 629 "parse.mly" +# 631 "parse.mly" (let x = _1 in ( [ x ] )) -# 7485 "parse.ml" +# 7501 "parse.ml" : 'separated_nonempty_list_COMMA_atomicTerm_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 2 : 'atomicTerm) in let _3 = (Parsing.peek_val __caml_parser_env 0 : 'separated_nonempty_list_COMMA_atomicTerm_) in Obj.repr( -# 632 "parse.mly" +# 634 "parse.mly" (let (x, _2, xs) = (_1, (), _3) in ( x :: xs )) -# 7494 "parse.ml" +# 7510 "parse.ml" : 'separated_nonempty_list_COMMA_atomicTerm_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'constructorPattern) in Obj.repr( -# 637 "parse.mly" +# 639 "parse.mly" (let x = _1 in ( [ x ] )) -# 7502 "parse.ml" +# 7518 "parse.ml" : 'separated_nonempty_list_COMMA_constructorPattern_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 2 : 'constructorPattern) in let _3 = (Parsing.peek_val __caml_parser_env 0 : 'separated_nonempty_list_COMMA_constructorPattern_) in Obj.repr( -# 640 "parse.mly" +# 642 "parse.mly" (let (x, _2, xs) = (_1, (), _3) in ( x :: xs )) -# 7511 "parse.ml" +# 7527 "parse.ml" : 'separated_nonempty_list_COMMA_constructorPattern_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'tmEq) in Obj.repr( -# 645 "parse.mly" +# 647 "parse.mly" (let x = _1 in ( [ x ] )) -# 7519 "parse.ml" +# 7535 "parse.ml" : 'separated_nonempty_list_COMMA_tmEq_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 2 : 'tmEq) in let _3 = (Parsing.peek_val __caml_parser_env 0 : 'separated_nonempty_list_COMMA_tmEq_) in Obj.repr( -# 648 "parse.mly" +# 650 "parse.mly" (let (x, _2, xs) = (_1, (), _3) in ( x :: xs )) -# 7528 "parse.ml" +# 7544 "parse.ml" : 'separated_nonempty_list_COMMA_tmEq_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'tvar) in Obj.repr( -# 653 "parse.mly" +# 655 "parse.mly" (let x = _1 in ( [ x ] )) -# 7536 "parse.ml" +# 7552 "parse.ml" : 'separated_nonempty_list_COMMA_tvar_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 2 : 'tvar) in let _3 = (Parsing.peek_val __caml_parser_env 0 : 'separated_nonempty_list_COMMA_tvar_) in Obj.repr( -# 656 "parse.mly" +# 658 "parse.mly" (let (x, _2, xs) = (_1, (), _3) in ( x :: xs )) -# 7545 "parse.ml" +# 7561 "parse.ml" : 'separated_nonempty_list_COMMA_tvar_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'conjunctivePat) in Obj.repr( -# 661 "parse.mly" +# 663 "parse.mly" (let x = _1 in ( [ x ] )) -# 7553 "parse.ml" +# 7569 "parse.ml" : 'separated_nonempty_list_DISJUNCTION_conjunctivePat_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 2 : 'conjunctivePat) in let _3 = (Parsing.peek_val __caml_parser_env 0 : 'separated_nonempty_list_DISJUNCTION_conjunctivePat_) in Obj.repr( -# 664 "parse.mly" +# 666 "parse.mly" (let (x, _2, xs) = (_1, (), _3) in ( x :: xs )) -# 7562 "parse.ml" +# 7578 "parse.ml" : 'separated_nonempty_list_DISJUNCTION_conjunctivePat_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'appTerm) in Obj.repr( -# 669 "parse.mly" +# 671 "parse.mly" (let x = _1 in ( [ x ] )) -# 7570 "parse.ml" +# 7586 "parse.ml" : 'separated_nonempty_list_SEMICOLON_appTerm_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 2 : 'appTerm) in let _3 = (Parsing.peek_val __caml_parser_env 0 : 'separated_nonempty_list_SEMICOLON_appTerm_) in Obj.repr( -# 672 "parse.mly" +# 674 "parse.mly" (let (x, _2, xs) = (_1, (), _3) in ( x :: xs )) -# 7579 "parse.ml" +# 7595 "parse.ml" : 'separated_nonempty_list_SEMICOLON_appTerm_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'effectDecl) in Obj.repr( -# 677 "parse.mly" +# 679 "parse.mly" (let x = _1 in ( [ x ] )) -# 7587 "parse.ml" +# 7603 "parse.ml" : 'separated_nonempty_list_SEMICOLON_effectDecl_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 2 : 'effectDecl) in let _3 = (Parsing.peek_val __caml_parser_env 0 : 'separated_nonempty_list_SEMICOLON_effectDecl_) in Obj.repr( -# 680 "parse.mly" +# 682 "parse.mly" (let (x, _2, xs) = (_1, (), _3) in ( x :: xs )) -# 7596 "parse.ml" +# 7612 "parse.ml" : 'separated_nonempty_list_SEMICOLON_effectDecl_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'ident) in Obj.repr( -# 685 "parse.mly" +# 687 "parse.mly" (let x = _1 in ( [ x ] )) -# 7604 "parse.ml" +# 7620 "parse.ml" : 'separated_nonempty_list_SEMICOLON_ident_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 2 : 'ident) in let _3 = (Parsing.peek_val __caml_parser_env 0 : 'separated_nonempty_list_SEMICOLON_ident_) in Obj.repr( -# 688 "parse.mly" +# 690 "parse.mly" (let (x, _2, xs) = (_1, (), _3) in ( x :: xs )) -# 7613 "parse.ml" +# 7629 "parse.ml" : 'separated_nonempty_list_SEMICOLON_ident_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'tuplePattern) in Obj.repr( -# 693 "parse.mly" +# 695 "parse.mly" (let x = _1 in ( [ x ] )) -# 7621 "parse.ml" +# 7637 "parse.ml" : 'separated_nonempty_list_SEMICOLON_tuplePattern_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 2 : 'tuplePattern) in let _3 = (Parsing.peek_val __caml_parser_env 0 : 'separated_nonempty_list_SEMICOLON_tuplePattern_) in Obj.repr( -# 696 "parse.mly" +# 698 "parse.mly" (let (x, _2, xs) = (_1, (), _3) in ( x :: xs )) -# 7630 "parse.ml" +# 7646 "parse.ml" : 'separated_nonempty_list_SEMICOLON_tuplePattern_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 1 : 'list_decl_) in Obj.repr( -# 701 "parse.mly" +# 703 "parse.mly" (let (decls, _2) = (_1, ()) in ( as_frag decls )) -# 7640 "parse.ml" +# 7656 "parse.ml" : FStar_Parser_AST.inputFragment)) +; (fun __caml_parser_env -> + Obj.repr( +# 710 "parse.mly" + (let _1 = () in + ( None )) +# 7663 "parse.ml" + : FStar_Parser_AST.decl option)) +; (fun __caml_parser_env -> + let _1 = (Parsing.peek_val __caml_parser_env 0 : 'decl) in + Obj.repr( +# 713 "parse.mly" + (let d = _1 in + ( Some d )) +# 7671 "parse.ml" + : FStar_Parser_AST.decl option)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 0 : 'string) in Obj.repr( -# 708 "parse.mly" +# 718 "parse.mly" (let (_1, s) = ((), _2) in ( SetOptions s )) -# 7648 "parse.ml" +# 7679 "parse.ml" : 'pragma)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 0 : 'option_string_) in Obj.repr( -# 711 "parse.mly" +# 721 "parse.mly" (let (_1, s_opt) = ((), _2) in ( ResetOptions s_opt )) -# 7656 "parse.ml" +# 7687 "parse.ml" : 'pragma)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 0 : 'option_string_) in Obj.repr( -# 714 "parse.mly" +# 724 "parse.mly" (let (_1, s_opt) = ((), _2) in ( PushOptions s_opt )) -# 7664 "parse.ml" +# 7695 "parse.ml" : 'pragma)) ; (fun __caml_parser_env -> Obj.repr( -# 717 "parse.mly" +# 727 "parse.mly" (let _1 = () in ( PopOptions )) -# 7671 "parse.ml" +# 7702 "parse.ml" : 'pragma)) ; (fun __caml_parser_env -> Obj.repr( -# 720 "parse.mly" +# 730 "parse.mly" (let _1 = () in ( RestartSolver )) -# 7678 "parse.ml" +# 7709 "parse.ml" : 'pragma)) ; (fun __caml_parser_env -> Obj.repr( -# 723 "parse.mly" +# 733 "parse.mly" (let _1 = () in ( PrintEffectsGraph )) -# 7685 "parse.ml" +# 7716 "parse.ml" : 'pragma)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 1 : 'list_atomicTerm_) in Obj.repr( -# 728 "parse.mly" +# 738 "parse.mly" (let (_1, x, _3) = ((), _2, ()) in ( let _ = @@ -7711,68 +7742,68 @@ let x = ( mk_ident (".(||)", rhs | _ -> () in x )) -# 7701 "parse.ml" +# 7732 "parse.ml" : 'attribute)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 1 : 'right_flexible_list_SEMICOLON_noSeqTerm_) in Obj.repr( -# 739 "parse.mly" +# 749 "parse.mly" (let (_1, l, _3) = ((), _2, ()) in let x = ( l ) in ( x )) -# 7710 "parse.ml" +# 7741 "parse.ml" : 'attribute)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'attribute) in Obj.repr( -# 745 "parse.mly" +# 755 "parse.mly" (let x = _1 in ( DeclAttributes x )) -# 7718 "parse.ml" +# 7749 "parse.ml" : 'decoration)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'qualifier) in Obj.repr( -# 748 "parse.mly" +# 758 "parse.mly" (let x = _1 in ( Qualifier x )) -# 7726 "parse.ml" +# 7757 "parse.ml" : 'decoration)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 2 : 'uident) in let _4 = (Parsing.peek_val __caml_parser_env 0 : 'noSeqTerm) in Obj.repr( -# 753 "parse.mly" +# 763 "parse.mly" (let (_1, lid, _3, e) = ((), _2, (), _4) in let phi = ( {e with level=Formula} ) in ( mk_decl (Assume(lid, phi)) (rhs2 parseState 1 4) [ Qualifier Assumption ] )) -# 7736 "parse.ml" +# 7767 "parse.ml" : 'decl)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 1 : 'list_decoration_) in let _2 = (Parsing.peek_val __caml_parser_env 0 : 'rawDecl) in Obj.repr( -# 757 "parse.mly" +# 767 "parse.mly" (let (ds, decl) = (_1, _2) in ( mk_decl decl (rhs parseState 2) ds )) -# 7745 "parse.ml" +# 7776 "parse.ml" : 'decl)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 1 : 'list_decoration_) in let _2 = (Parsing.peek_val __caml_parser_env 0 : 'typeclassDecl) in Obj.repr( -# 760 "parse.mly" +# 770 "parse.mly" (let (ds, decl) = (_1, _2) in ( let (decl, extra_attrs) = decl in let d = mk_decl decl (rhs parseState 2) ds in { d with attrs = extra_attrs @ d.attrs } )) -# 7757 "parse.ml" +# 7788 "parse.ml" : 'decl)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 0 : 'typeDecl) in Obj.repr( -# 768 "parse.mly" +# 778 "parse.mly" (let (_1, tcdef) = ((), _2) in ( (* Only a single type decl allowed, but construct it the same as for multiple ones. @@ -7783,13 +7814,13 @@ let phi = ( {e with level=Formula} ) in (* No attrs yet, but perhaps we want a `class` attribute *) (d, []) )) -# 7773 "parse.ml" +# 7804 "parse.ml" : 'typeclassDecl)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 1 : 'letqualifier) in let _3 = (Parsing.peek_val __caml_parser_env 0 : 'letbinding) in Obj.repr( -# 779 "parse.mly" +# 789 "parse.mly" (let (_1, q, lb) = ((), _2, _3) in ( (* Making a single letbinding *) @@ -7802,89 +7833,89 @@ let phi = ( {e with level=Formula} ) in (d, [at]) )) -# 7792 "parse.ml" +# 7823 "parse.ml" : 'typeclassDecl)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'pragma) in Obj.repr( -# 794 "parse.mly" +# 804 "parse.mly" (let p = _1 in ( Pragma p )) -# 7800 "parse.ml" +# 7831 "parse.ml" : 'rawDecl)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 0 : 'quident) in Obj.repr( -# 797 "parse.mly" +# 807 "parse.mly" (let (_1, uid) = ((), _2) in ( Open uid )) -# 7808 "parse.ml" +# 7839 "parse.ml" : 'rawDecl)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 0 : 'quident) in Obj.repr( -# 800 "parse.mly" +# 810 "parse.mly" (let (_1, uid) = ((), _2) in ( Friend uid )) -# 7816 "parse.ml" +# 7847 "parse.ml" : 'rawDecl)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 0 : 'quident) in Obj.repr( -# 803 "parse.mly" +# 813 "parse.mly" (let (_1, uid) = ((), _2) in ( Include uid )) -# 7824 "parse.ml" +# 7855 "parse.ml" : 'rawDecl)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 2 : 'uident) in let _4 = (Parsing.peek_val __caml_parser_env 0 : 'quident) in Obj.repr( -# 806 "parse.mly" +# 816 "parse.mly" (let (_1, uid1, _3, uid2) = ((), _2, (), _4) in ( ModuleAbbrev(uid1, uid2) )) -# 7833 "parse.ml" +# 7864 "parse.ml" : 'rawDecl)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 0 : 'qlident) in Obj.repr( -# 809 "parse.mly" +# 819 "parse.mly" (let (_1, _2) = ((), _2) in ( raise_error (Fatal_SyntaxError, "Syntax error: expected a module name") (rhs parseState 2) )) -# 7841 "parse.ml" +# 7872 "parse.ml" : 'rawDecl)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 0 : 'quident) in Obj.repr( -# 812 "parse.mly" +# 822 "parse.mly" (let (_1, uid) = ((), _2) in ( TopLevelModule uid )) -# 7849 "parse.ml" +# 7880 "parse.ml" : 'rawDecl)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 0 : 'separated_nonempty_list_AND_typeDecl_) in Obj.repr( -# 815 "parse.mly" +# 825 "parse.mly" (let (_1, tcdefs) = ((), _2) in ( Tycon (false, false, tcdefs) )) -# 7857 "parse.ml" +# 7888 "parse.ml" : 'rawDecl)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 3 : 'uident) in let _3 = (Parsing.peek_val __caml_parser_env 2 : 'typars) in let _5 = (Parsing.peek_val __caml_parser_env 0 : 'typ) in Obj.repr( -# 818 "parse.mly" +# 828 "parse.mly" (let (_1, uid, tparams, _4, t) = ((), _2, _3, (), _5) in ( Tycon(true, false, [(TyconAbbrev(uid, tparams, None, t))]) )) -# 7867 "parse.ml" +# 7898 "parse.ml" : 'rawDecl)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 2 : bool) in let _2 = (Parsing.peek_val __caml_parser_env 1 : 'letqualifier) in let _3 = (Parsing.peek_val __caml_parser_env 0 : 'separated_nonempty_list_AND_letbinding_) in Obj.repr( -# 821 "parse.mly" +# 831 "parse.mly" (let (_1, q, lbs) = (_1, _2, _3) in ( let r = rhs2 parseState 1 3 in @@ -7893,25 +7924,25 @@ let phi = ( {e with level=Formula} ) in then raise_error (Fatal_MultipleLetBinding, "Unexpected multiple let-binding (Did you forget some rec qualifier ?)") r; TopLevelLet(q, lbs) )) -# 7883 "parse.ml" +# 7914 "parse.ml" : 'rawDecl)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 0 : 'constant) in Obj.repr( -# 830 "parse.mly" +# 840 "parse.mly" (let (_1, c) = ((), _2) in ( (* This is just to provide a better error than "syntax error" *) raise_error (Fatal_SyntaxError, "Syntax error: constants are not allowed in val declarations") (rhs2 parseState 1 2) )) -# 7894 "parse.ml" +# 7925 "parse.ml" : 'rawDecl)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 3 : FStar_Ident.ident) in let _3 = (Parsing.peek_val __caml_parser_env 2 : 'list_multiBinder_) in let _5 = (Parsing.peek_val __caml_parser_env 0 : 'typ) in Obj.repr( -# 836 "parse.mly" +# 846 "parse.mly" (let (_1, id, bss, _4, t) = ((), _2, _3, (), _5) in let lid = ( id ) in ( @@ -7920,14 +7951,14 @@ let lid = ( id ) in | bs -> mk_term (Product(bs, t)) (rhs2 parseState 3 5) Type_level in Val(lid, t) )) -# 7910 "parse.ml" +# 7941 "parse.ml" : 'rawDecl)) ; (fun __caml_parser_env -> let _3 = (Parsing.peek_val __caml_parser_env 4 : string) in let _5 = (Parsing.peek_val __caml_parser_env 2 : 'list_multiBinder_) in let _7 = (Parsing.peek_val __caml_parser_env 0 : 'typ) in Obj.repr( -# 845 "parse.mly" +# 855 "parse.mly" (let (_1, _1_inlined1, op, _3, bss, _4, t) = ((), (), _3, (), _5, (), _7) in let lid = let id = ( mk_ident (op, rhs parseState 1) ) in @@ -7939,14 +7970,14 @@ in | bs -> mk_term (Product(bs, t)) (rhs2 parseState 3 5) Type_level in Val(lid, t) )) -# 7929 "parse.ml" +# 7960 "parse.ml" : 'rawDecl)) ; (fun __caml_parser_env -> let _3 = (Parsing.peek_val __caml_parser_env 4 : 'binop_name) in let _5 = (Parsing.peek_val __caml_parser_env 2 : 'list_multiBinder_) in let _7 = (Parsing.peek_val __caml_parser_env 0 : 'typ) in Obj.repr( -# 857 "parse.mly" +# 867 "parse.mly" (let (_1, _1_inlined1, op, _3, bss, _4, t) = ((), (), _3, (), _5, (), _7) in let lid = let id = ( op ) in @@ -7958,14 +7989,14 @@ in | bs -> mk_term (Product(bs, t)) (rhs2 parseState 3 5) Type_level in Val(lid, t) )) -# 7948 "parse.ml" +# 7979 "parse.ml" : 'rawDecl)) ; (fun __caml_parser_env -> let _3 = (Parsing.peek_val __caml_parser_env 4 : string) in let _5 = (Parsing.peek_val __caml_parser_env 2 : 'list_multiBinder_) in let _7 = (Parsing.peek_val __caml_parser_env 0 : 'typ) in Obj.repr( -# 869 "parse.mly" +# 879 "parse.mly" (let (_1, _1_inlined1, op, _3, bss, _4, t) = ((), (), _3, (), _5, (), _7) in let lid = let id = ( mk_ident (op, rhs parseState 1) ) in @@ -7977,14 +8008,14 @@ in | bs -> mk_term (Product(bs, t)) (rhs2 parseState 3 5) Type_level in Val(lid, t) )) -# 7967 "parse.ml" +# 7998 "parse.ml" : 'rawDecl)) ; (fun __caml_parser_env -> let _3 = (Parsing.peek_val __caml_parser_env 4 : string) in let _5 = (Parsing.peek_val __caml_parser_env 2 : 'list_multiBinder_) in let _7 = (Parsing.peek_val __caml_parser_env 0 : 'typ) in Obj.repr( -# 881 "parse.mly" +# 891 "parse.mly" (let (_1, _1_inlined1, op, _3, bss, _4, t) = ((), (), _3, (), _5, (), _7) in let lid = let id = @@ -7999,14 +8030,14 @@ in | bs -> mk_term (Product(bs, t)) (rhs2 parseState 3 5) Type_level in Val(lid, t) )) -# 7989 "parse.ml" +# 8020 "parse.ml" : 'rawDecl)) ; (fun __caml_parser_env -> let _3 = (Parsing.peek_val __caml_parser_env 4 : string) in let _5 = (Parsing.peek_val __caml_parser_env 2 : 'list_multiBinder_) in let _7 = (Parsing.peek_val __caml_parser_env 0 : 'typ) in Obj.repr( -# 896 "parse.mly" +# 906 "parse.mly" (let (_1, _1_inlined1, op, _3, bss, _4, t) = ((), (), _3, (), _5, (), _7) in let lid = let id = @@ -8021,74 +8052,74 @@ in | bs -> mk_term (Product(bs, t)) (rhs2 parseState 3 5) Type_level in Val(lid, t) )) -# 8011 "parse.ml" +# 8042 "parse.ml" : 'rawDecl)) ; (fun __caml_parser_env -> let _3 = (Parsing.peek_val __caml_parser_env 2 : 'loption_separated_nonempty_list_SEMICOLON_ident__) in let _5 = (Parsing.peek_val __caml_parser_env 0 : 'thunk_atomicTerm_) in Obj.repr( -# 911 "parse.mly" +# 921 "parse.mly" (let (_1, _2, xs, _4, t) = ((), (), _3, (), _5) in let ids = ( xs ) in ( Splice (ids, t) )) -# 8021 "parse.ml" +# 8052 "parse.ml" : 'rawDecl)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 1 : 'uident) in let _3 = (Parsing.peek_val __caml_parser_env 0 : 'option___anonymous_0_) in Obj.repr( -# 915 "parse.mly" +# 925 "parse.mly" (let (_1, lid, t_opt) = ((), _2, _3) in ( Exception(lid, t_opt) )) -# 8030 "parse.ml" +# 8061 "parse.ml" : 'rawDecl)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 0 : 'newEffect) in Obj.repr( -# 918 "parse.mly" +# 928 "parse.mly" (let (_1, ne) = ((), _2) in ( NewEffect ne )) -# 8038 "parse.ml" +# 8069 "parse.ml" : 'rawDecl)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 0 : 'effectDefinition) in Obj.repr( -# 921 "parse.mly" +# 931 "parse.mly" (let (_1, ne) = ((), _2) in ( LayeredEffect ne )) -# 8046 "parse.ml" +# 8077 "parse.ml" : 'rawDecl)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 0 : 'layeredEffectDefinition) in Obj.repr( -# 924 "parse.mly" +# 934 "parse.mly" (let (_1, ne) = ((), _2) in ( LayeredEffect ne )) -# 8054 "parse.ml" +# 8085 "parse.ml" : 'rawDecl)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 0 : 'subEffect) in Obj.repr( -# 927 "parse.mly" +# 937 "parse.mly" (let (_1, se) = ((), _2) in ( SubEffect se )) -# 8062 "parse.ml" +# 8093 "parse.ml" : 'rawDecl)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 0 : 'polymonadic_bind) in Obj.repr( -# 930 "parse.mly" +# 940 "parse.mly" (let (_1, b) = ((), _2) in ( Polymonadic_bind b )) -# 8070 "parse.ml" +# 8101 "parse.ml" : 'rawDecl)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 0 : 'polymonadic_subcomp) in Obj.repr( -# 933 "parse.mly" +# 943 "parse.mly" (let (_1, c) = ((), _2) in ( Polymonadic_subcomp c )) -# 8078 "parse.ml" +# 8109 "parse.ml" : 'rawDecl)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 3 : 'ident) in @@ -8096,162 +8127,162 @@ let ids = ( xs ) in let _3 = (Parsing.peek_val __caml_parser_env 1 : 'option_ascribeKind_) in let _4 = (Parsing.peek_val __caml_parser_env 0 : 'typeDefinition) in Obj.repr( -# 938 "parse.mly" +# 948 "parse.mly" (let (lid, tparams, ascr_opt, tcdef) = (_1, _2, _3, _4) in ( tcdef lid tparams ascr_opt )) -# 8089 "parse.ml" +# 8120 "parse.ml" : 'typeDecl)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'tvarinsts) in Obj.repr( -# 943 "parse.mly" +# 953 "parse.mly" (let x = _1 in ( x )) -# 8097 "parse.ml" +# 8128 "parse.ml" : 'typars)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'binders) in Obj.repr( -# 946 "parse.mly" +# 956 "parse.mly" (let x = _1 in ( x )) -# 8105 "parse.ml" +# 8136 "parse.ml" : 'typars)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 1 : 'separated_nonempty_list_COMMA_tvar_) in Obj.repr( -# 951 "parse.mly" +# 961 "parse.mly" (let (_1, tvs, _3) = ((), _2, ()) in ( map (fun tv -> mk_binder (TVariable(tv)) (range_of_id tv) Kind None) tvs )) -# 8113 "parse.ml" +# 8144 "parse.ml" : 'tvarinsts)) ; (fun __caml_parser_env -> Obj.repr( -# 956 "parse.mly" +# 966 "parse.mly" ( ( (fun id binders kopt -> check_id id; TyconAbstract(id, binders, kopt)) )) -# 8119 "parse.ml" +# 8150 "parse.ml" : 'typeDefinition)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 0 : 'typ) in Obj.repr( -# 958 "parse.mly" +# 968 "parse.mly" (let (_1, t) = ((), _2) in ( (fun id binders kopt -> check_id id; TyconAbbrev(id, binders, kopt, t)) )) -# 8127 "parse.ml" +# 8158 "parse.ml" : 'typeDefinition)) ; (fun __caml_parser_env -> let _3 = (Parsing.peek_val __caml_parser_env 1 : 'right_flexible_nonempty_list_SEMICOLON_recordFieldDecl_) in Obj.repr( -# 961 "parse.mly" +# 971 "parse.mly" (let (_1, _1_inlined1, record_field_decls, _3) = ((), (), _3, ()) in let record_field_decls = ( record_field_decls ) in let attrs_opt = ( None ) in ( (fun id binders kopt -> check_id id; TyconRecord(id, binders, kopt, none_to_empty_list attrs_opt, record_field_decls)) )) -# 8137 "parse.ml" +# 8168 "parse.ml" : 'typeDefinition)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 3 : 'binderAttributes) in let _4 = (Parsing.peek_val __caml_parser_env 1 : 'right_flexible_nonempty_list_SEMICOLON_recordFieldDecl_) in Obj.repr( -# 966 "parse.mly" +# 976 "parse.mly" (let (_1, x, _1_inlined1, record_field_decls, _3) = ((), _2, (), _4, ()) in let record_field_decls = ( record_field_decls ) in let attrs_opt = ( Some x ) in ( (fun id binders kopt -> check_id id; TyconRecord(id, binders, kopt, none_to_empty_list attrs_opt, record_field_decls)) )) -# 8148 "parse.ml" +# 8179 "parse.ml" : 'typeDefinition)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 0 : 'list_constructorDecl_) in Obj.repr( -# 971 "parse.mly" +# 981 "parse.mly" (let (_1, ct_decls) = ((), _2) in ( (fun id binders kopt -> check_id id; TyconVariant(id, binders, kopt, ct_decls)) )) -# 8156 "parse.ml" +# 8187 "parse.ml" : 'typeDefinition)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 2 : 'aqualifiedWithAttrs_lidentOrOperator_) in let _3 = (Parsing.peek_val __caml_parser_env 0 : 'typ) in Obj.repr( -# 976 "parse.mly" +# 986 "parse.mly" (let (qualified_lid, _2, t) = (_1, (), _3) in ( let (qual, attrs), lid = qualified_lid in (lid, qual, attrs, t) )) -# 8168 "parse.ml" +# 8199 "parse.ml" : 'recordFieldDecl)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 0 : 'typ) in Obj.repr( -# 984 "parse.mly" +# 994 "parse.mly" (let (_1, t) = ((), _2) in (VpArbitrary t)) -# 8176 "parse.ml" +# 8207 "parse.ml" : 'constructorPayload)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 0 : 'typ) in Obj.repr( -# 987 "parse.mly" +# 997 "parse.mly" (let (_1, t) = ((), _2) in (VpOfNotation t)) -# 8184 "parse.ml" +# 8215 "parse.ml" : 'constructorPayload)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 2 : 'right_flexible_nonempty_list_SEMICOLON_recordFieldDecl_) in let _4 = (Parsing.peek_val __caml_parser_env 0 : 'option___anonymous_1_) in Obj.repr( -# 990 "parse.mly" +# 1000 "parse.mly" (let (_1, record_field_decls, _3, opt) = ((), _2, (), _4) in let fields = ( record_field_decls ) in (VpRecord(fields, opt))) -# 8194 "parse.ml" +# 8225 "parse.ml" : 'constructorPayload)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 1 : 'uident) in let _3 = (Parsing.peek_val __caml_parser_env 0 : 'option_constructorPayload_) in Obj.repr( -# 996 "parse.mly" +# 1006 "parse.mly" (let (_1, uid, payload) = ((), _2, _3) in let attrs_opt = ( None ) in ( uid, payload, none_to_empty_list attrs_opt )) -# 8204 "parse.ml" +# 8235 "parse.ml" : 'constructorDecl)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 2 : 'binderAttributes) in let _3 = (Parsing.peek_val __caml_parser_env 1 : 'uident) in let _4 = (Parsing.peek_val __caml_parser_env 0 : 'option_constructorPayload_) in Obj.repr( -# 1000 "parse.mly" +# 1010 "parse.mly" (let (_1, x, uid, payload) = ((), _2, _3, _4) in let attrs_opt = ( Some x ) in ( uid, payload, none_to_empty_list attrs_opt )) -# 8215 "parse.ml" +# 8246 "parse.ml" : 'constructorDecl)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 0 : 'letbinding) in Obj.repr( -# 1006 "parse.mly" +# 1016 "parse.mly" (let (_2, lb) = ((), _2) in let attr = ( None ) in ( attr, lb )) -# 8224 "parse.ml" +# 8255 "parse.ml" : 'attr_letbinding)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 2 : 'attribute) in let _3 = (Parsing.peek_val __caml_parser_env 0 : 'letbinding) in Obj.repr( -# 1010 "parse.mly" +# 1020 "parse.mly" (let (x, _2, lb) = (_1, (), _3) in let attr = ( Some x ) in ( attr, lb )) -# 8234 "parse.ml" +# 8265 "parse.ml" : 'attr_letbinding)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 2 : 'tuplePattern) in let _2 = (Parsing.peek_val __caml_parser_env 1 : 'option_ascribeTyp_) in let _3 = (Parsing.peek_val __caml_parser_env 0 : 'option___anonymous_2_) in Obj.repr( -# 1016 "parse.mly" +# 1026 "parse.mly" (let (pat, ascr_opt, tm) = (_1, _2, _3) in ( let h tm @@ -8267,7 +8298,7 @@ let attr = ( Some x ) in h (mk_term (Var v) (rhs parseState 1) Expr) | _ -> raise_error (Fatal_SyntaxError, "Syntax error: let-punning expects a name, not a pattern") (rhs parseState 2) )) -# 8257 "parse.ml" +# 8288 "parse.ml" : 'letoperatorbinding)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 5 : 'maybeFocus) in @@ -8276,7 +8307,7 @@ let attr = ( Some x ) in let _4 = (Parsing.peek_val __caml_parser_env 2 : 'option_ascribeTyp_) in let _6 = (Parsing.peek_val __caml_parser_env 0 : FStar_Parser_AST.term) in Obj.repr( -# 1034 "parse.mly" +# 1044 "parse.mly" (let (focus_opt, id, lbp, ascr_opt, _5, tm) = (_1, _2, _3, _4, (), _6) in let lid = ( id ) in ( @@ -8287,7 +8318,7 @@ let lid = ( id ) in | None -> (focus_opt, (pat, tm)) | Some t -> (focus_opt, (mk_pattern (PatAscribed(pat, t)) pos, tm)) )) -# 8277 "parse.ml" +# 8308 "parse.ml" : 'letbinding)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 7 : 'maybeFocus) in @@ -8296,7 +8327,7 @@ let lid = ( id ) in let _6 = (Parsing.peek_val __caml_parser_env 2 : 'option_ascribeTyp_) in let _8 = (Parsing.peek_val __caml_parser_env 0 : FStar_Parser_AST.term) in Obj.repr( -# 1045 "parse.mly" +# 1055 "parse.mly" (let (focus_opt, _1, op, _3, lbp, ascr_opt, _5, tm) = (_1, (), _3, (), _5, _6, (), _8) in let lid = let id = ( mk_ident (op, rhs parseState 1) ) in @@ -8310,7 +8341,7 @@ in | None -> (focus_opt, (pat, tm)) | Some t -> (focus_opt, (mk_pattern (PatAscribed(pat, t)) pos, tm)) )) -# 8300 "parse.ml" +# 8331 "parse.ml" : 'letbinding)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 7 : 'maybeFocus) in @@ -8319,7 +8350,7 @@ in let _6 = (Parsing.peek_val __caml_parser_env 2 : 'option_ascribeTyp_) in let _8 = (Parsing.peek_val __caml_parser_env 0 : FStar_Parser_AST.term) in Obj.repr( -# 1059 "parse.mly" +# 1069 "parse.mly" (let (focus_opt, _1, op, _3, lbp, ascr_opt, _5, tm) = (_1, (), _3, (), _5, _6, (), _8) in let lid = let id = ( op ) in @@ -8333,7 +8364,7 @@ in | None -> (focus_opt, (pat, tm)) | Some t -> (focus_opt, (mk_pattern (PatAscribed(pat, t)) pos, tm)) )) -# 8323 "parse.ml" +# 8354 "parse.ml" : 'letbinding)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 7 : 'maybeFocus) in @@ -8342,7 +8373,7 @@ in let _6 = (Parsing.peek_val __caml_parser_env 2 : 'option_ascribeTyp_) in let _8 = (Parsing.peek_val __caml_parser_env 0 : FStar_Parser_AST.term) in Obj.repr( -# 1073 "parse.mly" +# 1083 "parse.mly" (let (focus_opt, _1, op, _3, lbp, ascr_opt, _5, tm) = (_1, (), _3, (), _5, _6, (), _8) in let lid = let id = ( mk_ident (op, rhs parseState 1) ) in @@ -8356,7 +8387,7 @@ in | None -> (focus_opt, (pat, tm)) | Some t -> (focus_opt, (mk_pattern (PatAscribed(pat, t)) pos, tm)) )) -# 8346 "parse.ml" +# 8377 "parse.ml" : 'letbinding)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 7 : 'maybeFocus) in @@ -8365,7 +8396,7 @@ in let _6 = (Parsing.peek_val __caml_parser_env 2 : 'option_ascribeTyp_) in let _8 = (Parsing.peek_val __caml_parser_env 0 : FStar_Parser_AST.term) in Obj.repr( -# 1087 "parse.mly" +# 1097 "parse.mly" (let (focus_opt, _1, op, _3, lbp, ascr_opt, _5, tm) = (_1, (), _3, (), _5, _6, (), _8) in let lid = let id = @@ -8382,7 +8413,7 @@ in | None -> (focus_opt, (pat, tm)) | Some t -> (focus_opt, (mk_pattern (PatAscribed(pat, t)) pos, tm)) )) -# 8372 "parse.ml" +# 8403 "parse.ml" : 'letbinding)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 7 : 'maybeFocus) in @@ -8391,7 +8422,7 @@ in let _6 = (Parsing.peek_val __caml_parser_env 2 : 'option_ascribeTyp_) in let _8 = (Parsing.peek_val __caml_parser_env 0 : FStar_Parser_AST.term) in Obj.repr( -# 1104 "parse.mly" +# 1114 "parse.mly" (let (focus_opt, _1, op, _3, lbp, ascr_opt, _5, tm) = (_1, (), _3, (), _5, _6, (), _8) in let lid = let id = @@ -8408,7 +8439,7 @@ in | None -> (focus_opt, (pat, tm)) | Some t -> (focus_opt, (mk_pattern (PatAscribed(pat, t)) pos, tm)) )) -# 8398 "parse.ml" +# 8429 "parse.ml" : 'letbinding)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 4 : 'maybeFocus) in @@ -8416,45 +8447,45 @@ in let _3 = (Parsing.peek_val __caml_parser_env 2 : 'ascribeTyp) in let _5 = (Parsing.peek_val __caml_parser_env 0 : FStar_Parser_AST.term) in Obj.repr( -# 1121 "parse.mly" +# 1131 "parse.mly" (let (focus_opt, pat, ascr, _4, tm) = (_1, _2, _3, (), _5) in ( focus_opt, (mk_pattern (PatAscribed(pat, ascr)) (rhs2 parseState 1 4), tm) )) -# 8409 "parse.ml" +# 8440 "parse.ml" : 'letbinding)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 3 : 'maybeFocus) in let _2 = (Parsing.peek_val __caml_parser_env 2 : 'tuplePattern) in let _4 = (Parsing.peek_val __caml_parser_env 0 : FStar_Parser_AST.term) in Obj.repr( -# 1124 "parse.mly" +# 1134 "parse.mly" (let (focus_opt, pat, _3, tm) = (_1, _2, (), _4) in ( focus_opt, (pat, tm) )) -# 8419 "parse.ml" +# 8450 "parse.ml" : 'letbinding)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'effectRedefinition) in Obj.repr( -# 1129 "parse.mly" +# 1139 "parse.mly" (let ed = _1 in ( ed )) -# 8427 "parse.ml" +# 8458 "parse.ml" : 'newEffect)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'effectDefinition) in Obj.repr( -# 1132 "parse.mly" +# 1142 "parse.mly" (let ed = _1 in ( ed )) -# 8435 "parse.ml" +# 8466 "parse.ml" : 'newEffect)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 2 : 'uident) in let _3 = (Parsing.peek_val __caml_parser_env 0 : 'simpleTerm) in Obj.repr( -# 1137 "parse.mly" +# 1147 "parse.mly" (let (lid, _2, t) = (_1, (), _3) in ( RedefineEffect(lid, [], t) )) -# 8444 "parse.ml" +# 8475 "parse.ml" : 'effectRedefinition)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 6 : 'uident) in @@ -8462,17 +8493,17 @@ in let _5 = (Parsing.peek_val __caml_parser_env 3 : 'tmArrow_tmNoEq_) in let _7 = (Parsing.peek_val __caml_parser_env 1 : 'separated_nonempty_list_SEMICOLON_effectDecl_) in Obj.repr( -# 1142 "parse.mly" +# 1152 "parse.mly" (let (_1, lid, bs, _4, typ, _6, eds, _8) = ((), _2, _3, (), _5, (), _7, ()) in ( DefineEffect(lid, bs, typ, eds) )) -# 8455 "parse.ml" +# 8486 "parse.ml" : 'effectDefinition)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 4 : 'uident) in let _3 = (Parsing.peek_val __caml_parser_env 3 : 'binders) in let _5 = (Parsing.peek_val __caml_parser_env 1 : 'tmNoEq) in Obj.repr( -# 1147 "parse.mly" +# 1157 "parse.mly" (let (_1, lid, bs, _4, r, _6) = ((), _2, _3, (), _5, ()) in ( let typ = (* bs -> Effect *) @@ -8499,27 +8530,27 @@ in "Syntax error: layered effect combinators should be declared as a record") r.range in DefineEffect (lid, [], typ, decls r) )) -# 8489 "parse.ml" +# 8520 "parse.ml" : 'layeredEffectDefinition)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 3 : FStar_Ident.ident) in let _2 = (Parsing.peek_val __caml_parser_env 2 : 'binders) in let _4 = (Parsing.peek_val __caml_parser_env 0 : 'simpleTerm) in Obj.repr( -# 1176 "parse.mly" +# 1186 "parse.mly" (let (lid, action_params, _3, t) = (_1, _2, (), _4) in ( mk_decl (Tycon (false, false, [TyconAbbrev(lid, action_params, None, t)])) (rhs2 parseState 1 3) [] )) -# 8499 "parse.ml" +# 8530 "parse.ml" : 'effectDecl)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 4 : 'quident) in let _3 = (Parsing.peek_val __caml_parser_env 2 : 'quident) in let _5 = (Parsing.peek_val __caml_parser_env 0 : 'simpleTerm) in Obj.repr( -# 1181 "parse.mly" +# 1191 "parse.mly" (let (src_eff, _2, tgt_eff, _4, lift) = (_1, (), _3, (), _5) in - ( { msource = src_eff; mdest = tgt_eff; lift_op = NonReifiableLift lift } )) -# 8509 "parse.ml" + ( { msource = src_eff; mdest = tgt_eff; lift_op = NonReifiableLift lift; braced=false } )) +# 8540 "parse.ml" : 'subEffect)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 7 : 'quident) in @@ -8527,7 +8558,7 @@ in let _5 = (Parsing.peek_val __caml_parser_env 3 : string) in let _7 = (Parsing.peek_val __caml_parser_env 1 : 'simpleTerm) in Obj.repr( -# 1184 "parse.mly" +# 1194 "parse.mly" (let (src_eff, _2, tgt_eff, _4, x, _2_inlined1, y, _7) = (_1, (), _3, (), _5, (), _7, ()) in let lift2_opt = ( None ) in let lift1 = ( (x, y) ) in @@ -8536,9 +8567,9 @@ let lift1 = ( (x, y) ) in | None -> begin match lift1 with | ("lift", lift) -> - { msource = src_eff; mdest = tgt_eff; lift_op = LiftForFree lift } + { msource = src_eff; mdest = tgt_eff; lift_op = LiftForFree lift; braced=true } | ("lift_wp", lift_wp) -> - { msource = src_eff; mdest = tgt_eff; lift_op = NonReifiableLift lift_wp } + { msource = src_eff; mdest = tgt_eff; lift_op = NonReifiableLift lift_wp; braced=true } | _ -> raise_error (Fatal_UnexpectedIdentifier, "Unexpected identifier; expected {'lift', and possibly 'lift_wp'}") (lhs parseState) end @@ -8549,9 +8580,9 @@ let lift1 = ( (x, y) ) in | "lift", "lift_wp" -> tm2, tm1 | _ -> raise_error (Fatal_UnexpectedIdentifier, "Unexpected identifier; expected {'lift', 'lift_wp'}") (lhs parseState) in - { msource = src_eff; mdest = tgt_eff; lift_op = ReifiableLift (lift, lift_wp) } + { msource = src_eff; mdest = tgt_eff; lift_op = ReifiableLift (lift, lift_wp); braced=true } )) -# 8541 "parse.ml" +# 8572 "parse.ml" : 'subEffect)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 11 : 'quident) in @@ -8561,7 +8592,7 @@ let lift1 = ( (x, y) ) in let _9 = (Parsing.peek_val __caml_parser_env 3 : string) in let _11 = (Parsing.peek_val __caml_parser_env 1 : 'simpleTerm) in Obj.repr( -# 1208 "parse.mly" +# 1218 "parse.mly" (let (src_eff, _2, tgt_eff, _4, x, _2_inlined1, y, _1, id, _2_inlined2, y_inlined1, _7) = (_1, (), _3, (), _5, (), _7, (), _9, (), _11, ()) in let lift2_opt = let y = y_inlined1 in @@ -8577,9 +8608,9 @@ let lift1 = ( (x, y) ) in | None -> begin match lift1 with | ("lift", lift) -> - { msource = src_eff; mdest = tgt_eff; lift_op = LiftForFree lift } + { msource = src_eff; mdest = tgt_eff; lift_op = LiftForFree lift; braced=true } | ("lift_wp", lift_wp) -> - { msource = src_eff; mdest = tgt_eff; lift_op = NonReifiableLift lift_wp } + { msource = src_eff; mdest = tgt_eff; lift_op = NonReifiableLift lift_wp; braced=true } | _ -> raise_error (Fatal_UnexpectedIdentifier, "Unexpected identifier; expected {'lift', and possibly 'lift_wp'}") (lhs parseState) end @@ -8590,9 +8621,9 @@ let lift1 = ( (x, y) ) in | "lift", "lift_wp" -> tm2, tm1 | _ -> raise_error (Fatal_UnexpectedIdentifier, "Unexpected identifier; expected {'lift', 'lift_wp'}") (lhs parseState) in - { msource = src_eff; mdest = tgt_eff; lift_op = ReifiableLift (lift, lift_wp) } + { msource = src_eff; mdest = tgt_eff; lift_op = ReifiableLift (lift, lift_wp); braced=true } )) -# 8582 "parse.ml" +# 8613 "parse.ml" : 'subEffect)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 7 : 'quident) in @@ -8600,437 +8631,437 @@ let lift1 = ( (x, y) ) in let _7 = (Parsing.peek_val __caml_parser_env 2 : 'quident) in let _9 = (Parsing.peek_val __caml_parser_env 0 : 'simpleTerm) in Obj.repr( -# 1241 "parse.mly" +# 1251 "parse.mly" (let (_1, m_eff, _3, n_eff, _5, _6, p_eff, _8, bind) = ((), _2, (), _4, (), (), _7, (), _9) in ( (m_eff, n_eff, p_eff, bind) )) -# 8593 "parse.ml" +# 8624 "parse.ml" : 'polymonadic_bind)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 4 : 'quident) in let _3 = (Parsing.peek_val __caml_parser_env 2 : 'quident) in let _5 = (Parsing.peek_val __caml_parser_env 0 : 'simpleTerm) in Obj.repr( -# 1246 "parse.mly" +# 1256 "parse.mly" (let (m_eff, _2, n_eff, _4, subcomp) = (_1, (), _3, (), _5) in ( (m_eff, n_eff, subcomp) )) -# 8603 "parse.ml" +# 8634 "parse.ml" : 'polymonadic_subcomp)) ; (fun __caml_parser_env -> Obj.repr( -# 1251 "parse.mly" +# 1261 "parse.mly" (let _1 = () in ( Assumption )) -# 8610 "parse.ml" +# 8641 "parse.ml" : 'qualifier)) ; (fun __caml_parser_env -> Obj.repr( -# 1254 "parse.mly" +# 1264 "parse.mly" (let _1 = () in ( raise_error (Fatal_InlineRenamedAsUnfold, "The 'inline' qualifier has been renamed to 'unfold'") (lhs parseState) )) -# 8619 "parse.ml" +# 8650 "parse.ml" : 'qualifier)) ; (fun __caml_parser_env -> Obj.repr( -# 1259 "parse.mly" +# 1269 "parse.mly" (let _1 = () in ( raise_error (Fatal_UnfoldableDeprecated, "The 'unfoldable' qualifier is no longer denotable; it is the default qualifier so just omit it") (lhs parseState) )) -# 8628 "parse.ml" +# 8659 "parse.ml" : 'qualifier)) ; (fun __caml_parser_env -> Obj.repr( -# 1264 "parse.mly" +# 1274 "parse.mly" (let _1 = () in ( Inline_for_extraction )) -# 8637 "parse.ml" +# 8668 "parse.ml" : 'qualifier)) ; (fun __caml_parser_env -> Obj.repr( -# 1269 "parse.mly" +# 1279 "parse.mly" (let _1 = () in ( Unfold_for_unification_and_vcgen )) -# 8646 "parse.ml" +# 8677 "parse.ml" : 'qualifier)) ; (fun __caml_parser_env -> Obj.repr( -# 1274 "parse.mly" +# 1284 "parse.mly" (let _1 = () in ( Irreducible )) -# 8653 "parse.ml" +# 8684 "parse.ml" : 'qualifier)) ; (fun __caml_parser_env -> Obj.repr( -# 1277 "parse.mly" +# 1287 "parse.mly" (let _1 = () in ( NoExtract )) -# 8660 "parse.ml" +# 8691 "parse.ml" : 'qualifier)) ; (fun __caml_parser_env -> Obj.repr( -# 1280 "parse.mly" +# 1290 "parse.mly" (let _1 = () in ( DefaultEffect )) -# 8667 "parse.ml" +# 8698 "parse.ml" : 'qualifier)) ; (fun __caml_parser_env -> Obj.repr( -# 1283 "parse.mly" +# 1293 "parse.mly" (let _1 = () in ( TotalEffect )) -# 8674 "parse.ml" +# 8705 "parse.ml" : 'qualifier)) ; (fun __caml_parser_env -> Obj.repr( -# 1286 "parse.mly" +# 1296 "parse.mly" (let _1 = () in ( Private )) -# 8681 "parse.ml" +# 8712 "parse.ml" : 'qualifier)) ; (fun __caml_parser_env -> Obj.repr( -# 1289 "parse.mly" +# 1299 "parse.mly" (let _1 = () in ( Noeq )) -# 8688 "parse.ml" +# 8719 "parse.ml" : 'qualifier)) ; (fun __caml_parser_env -> Obj.repr( -# 1292 "parse.mly" +# 1302 "parse.mly" (let _1 = () in ( Unopteq )) -# 8695 "parse.ml" +# 8726 "parse.ml" : 'qualifier)) ; (fun __caml_parser_env -> Obj.repr( -# 1295 "parse.mly" +# 1305 "parse.mly" (let _1 = () in ( New )) -# 8702 "parse.ml" +# 8733 "parse.ml" : 'qualifier)) ; (fun __caml_parser_env -> Obj.repr( -# 1298 "parse.mly" +# 1308 "parse.mly" (let _1 = () in ( log_issue (lhs parseState) (Warning_logicqualifier, logic_qualifier_deprecation_warning); Logic )) -# 8711 "parse.ml" +# 8742 "parse.ml" : 'qualifier)) ; (fun __caml_parser_env -> Obj.repr( -# 1303 "parse.mly" +# 1313 "parse.mly" (let _1 = () in ( Opaque )) -# 8718 "parse.ml" +# 8749 "parse.ml" : 'qualifier)) ; (fun __caml_parser_env -> Obj.repr( -# 1306 "parse.mly" +# 1316 "parse.mly" (let _1 = () in ( Reifiable )) -# 8725 "parse.ml" +# 8756 "parse.ml" : 'qualifier)) ; (fun __caml_parser_env -> Obj.repr( -# 1309 "parse.mly" +# 1319 "parse.mly" (let _1 = () in ( Reflectable )) -# 8732 "parse.ml" +# 8763 "parse.ml" : 'qualifier)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'boption_SQUIGGLY_RARROW_) in Obj.repr( -# 1314 "parse.mly" +# 1324 "parse.mly" (let b = _1 in ( b )) -# 8740 "parse.ml" +# 8771 "parse.ml" : 'maybeFocus)) ; (fun __caml_parser_env -> Obj.repr( -# 1319 "parse.mly" +# 1329 "parse.mly" (let _1 = () in ( Rec )) -# 8747 "parse.ml" +# 8778 "parse.ml" : 'letqualifier)) ; (fun __caml_parser_env -> Obj.repr( -# 1322 "parse.mly" +# 1332 "parse.mly" ( ( NoLetQualifier )) -# 8753 "parse.ml" +# 8784 "parse.ml" : 'letqualifier)) ; (fun __caml_parser_env -> let _3 = (Parsing.peek_val __caml_parser_env 1 : 'thunk_tmNoEq_) in Obj.repr( -# 1326 "parse.mly" +# 1336 "parse.mly" (let (_1, _2, t, _4) = ((), (), _3, ()) in ( mk_meta_tac t )) -# 8761 "parse.ml" +# 8792 "parse.ml" : 'aqual)) ; (fun __caml_parser_env -> Obj.repr( -# 1329 "parse.mly" +# 1339 "parse.mly" (let _1 = () in ( Implicit )) -# 8768 "parse.ml" +# 8799 "parse.ml" : 'aqual)) ; (fun __caml_parser_env -> Obj.repr( -# 1332 "parse.mly" +# 1342 "parse.mly" (let _1 = () in ( Equality )) -# 8775 "parse.ml" +# 8806 "parse.ml" : 'aqual)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 1 : 'right_flexible_list_SEMICOLON_noSeqTerm_) in Obj.repr( -# 1337 "parse.mly" +# 1347 "parse.mly" (let (_1, l, _3) = ((), _2, ()) in let t = ( l ) in ( t )) -# 8784 "parse.ml" +# 8815 "parse.ml" : 'binderAttributes)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'separated_nonempty_list_BAR_tuplePattern_) in Obj.repr( -# 1343 "parse.mly" +# 1353 "parse.mly" (let pats = _1 in ( pats )) -# 8792 "parse.ml" +# 8823 "parse.ml" : 'disjunctivePattern)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'separated_nonempty_list_COMMA_constructorPattern_) in Obj.repr( -# 1348 "parse.mly" +# 1358 "parse.mly" (let pats = _1 in ( match pats with | [x] -> x | l -> mk_pattern (PatTuple (l, false)) (rhs parseState 1) )) -# 8800 "parse.ml" +# 8831 "parse.ml" : 'tuplePattern)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 2 : 'constructorPattern) in let _3 = (Parsing.peek_val __caml_parser_env 0 : 'constructorPattern) in Obj.repr( -# 1353 "parse.mly" +# 1363 "parse.mly" (let (pat, _2, pats) = (_1, (), _3) in ( mk_pattern (consPat (rhs parseState 3) pat pats) (rhs2 parseState 1 3) )) -# 8809 "parse.ml" +# 8840 "parse.ml" : 'constructorPattern)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 1 : 'quident) in let _2 = (Parsing.peek_val __caml_parser_env 0 : 'nonempty_list_atomicPattern_) in Obj.repr( -# 1356 "parse.mly" +# 1366 "parse.mly" (let (uid, args) = (_1, _2) in ( let head_pat = mk_pattern (PatName uid) (rhs parseState 1) in mk_pattern (PatApp (head_pat, args)) (rhs2 parseState 1 2) )) -# 8821 "parse.ml" +# 8852 "parse.ml" : 'constructorPattern)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'atomicPattern) in Obj.repr( -# 1362 "parse.mly" +# 1372 "parse.mly" (let pat = _1 in ( pat )) -# 8829 "parse.ml" +# 8860 "parse.ml" : 'constructorPattern)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 4 : 'tuplePattern) in let _4 = (Parsing.peek_val __caml_parser_env 2 : 'simpleArrow) in let _5 = (Parsing.peek_val __caml_parser_env 1 : 'refineOpt) in Obj.repr( -# 1367 "parse.mly" +# 1377 "parse.mly" (let (_1, pat, _3, t, phi_opt, _6) = ((), _2, (), _4, _5, ()) in ( let pos_t = rhs2 parseState 2 4 in let pos = rhs2 parseState 1 6 in mkRefinedPattern pat t true phi_opt pos_t pos )) -# 8843 "parse.ml" +# 8874 "parse.ml" : 'atomicPattern)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 1 : 'loption_separated_nonempty_list_SEMICOLON_tuplePattern__) in Obj.repr( -# 1374 "parse.mly" +# 1384 "parse.mly" (let (_1, xs, _3) = ((), _2, ()) in let pats = ( xs ) in ( mk_pattern (PatList pats) (rhs2 parseState 1 3) )) -# 8852 "parse.ml" +# 8883 "parse.ml" : 'atomicPattern)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 1 : 'right_flexible_nonempty_list_SEMICOLON_fieldPattern_) in Obj.repr( -# 1378 "parse.mly" +# 1388 "parse.mly" (let (_1, record_pat, _3) = ((), _2, ()) in ( mk_pattern (PatRecord record_pat) (rhs2 parseState 1 3) )) -# 8860 "parse.ml" +# 8891 "parse.ml" : 'atomicPattern)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 3 : 'constructorPattern) in let _4 = (Parsing.peek_val __caml_parser_env 1 : 'separated_nonempty_list_COMMA_constructorPattern_) in Obj.repr( -# 1381 "parse.mly" +# 1391 "parse.mly" (let (_1, pat0, _3, pats, _5) = ((), _2, (), _4, ()) in ( mk_pattern (PatTuple(pat0::pats, true)) (rhs2 parseState 1 5) )) -# 8869 "parse.ml" +# 8900 "parse.ml" : 'atomicPattern)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 1 : 'tuplePattern) in Obj.repr( -# 1384 "parse.mly" +# 1394 "parse.mly" (let (_1, pat, _3) = ((), _2, ()) in ( pat )) -# 8877 "parse.ml" +# 8908 "parse.ml" : 'atomicPattern)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'tvar) in Obj.repr( -# 1387 "parse.mly" +# 1397 "parse.mly" (let tv = _1 in ( mk_pattern (PatTvar (tv, None, [])) (rhs parseState 1) )) -# 8885 "parse.ml" +# 8916 "parse.ml" : 'atomicPattern)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 1 : string) in Obj.repr( -# 1390 "parse.mly" +# 1400 "parse.mly" (let (_1, op, _3) = ((), _2, ()) in let op = ( mk_ident (op, rhs parseState 1) ) in ( mk_pattern (PatOp op) (rhs2 parseState 1 3) )) -# 8894 "parse.ml" +# 8925 "parse.ml" : 'atomicPattern)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 1 : 'binop_name) in Obj.repr( -# 1394 "parse.mly" +# 1404 "parse.mly" (let (_1, op, _3) = ((), _2, ()) in let op = ( op ) in ( mk_pattern (PatOp op) (rhs2 parseState 1 3) )) -# 8903 "parse.ml" +# 8934 "parse.ml" : 'atomicPattern)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 1 : string) in Obj.repr( -# 1398 "parse.mly" +# 1408 "parse.mly" (let (_1, op, _3) = ((), _2, ()) in let op = ( mk_ident (op, rhs parseState 1) ) in ( mk_pattern (PatOp op) (rhs2 parseState 1 3) )) -# 8912 "parse.ml" +# 8943 "parse.ml" : 'atomicPattern)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 1 : string) in Obj.repr( -# 1402 "parse.mly" +# 1412 "parse.mly" (let (_1, op, _3) = ((), _2, ()) in let op = let op = ( mk_ident ("and" ^ op, rhs parseState 1) ) in (op) in ( mk_pattern (PatOp op) (rhs2 parseState 1 3) )) -# 8924 "parse.ml" +# 8955 "parse.ml" : 'atomicPattern)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 1 : string) in Obj.repr( -# 1409 "parse.mly" +# 1419 "parse.mly" (let (_1, op, _3) = ((), _2, ()) in let op = let op = ( mk_ident ("let" ^ op, rhs parseState 1) ) in (op) in ( mk_pattern (PatOp op) (rhs2 parseState 1 3) )) -# 8936 "parse.ml" +# 8967 "parse.ml" : 'atomicPattern)) ; (fun __caml_parser_env -> Obj.repr( -# 1416 "parse.mly" +# 1426 "parse.mly" (let _1 = () in ( mk_pattern (PatWild (None, [])) (rhs parseState 1) )) -# 8943 "parse.ml" +# 8974 "parse.ml" : 'atomicPattern)) ; (fun __caml_parser_env -> Obj.repr( -# 1419 "parse.mly" +# 1429 "parse.mly" (let (_1, _2) = ((), ()) in ( mk_pattern (PatWild (Some Implicit, [])) (rhs parseState 1) )) -# 8950 "parse.ml" +# 8981 "parse.ml" : 'atomicPattern)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'constant) in Obj.repr( -# 1422 "parse.mly" +# 1432 "parse.mly" (let c = _1 in ( mk_pattern (PatConst c) (rhs parseState 1) )) -# 8958 "parse.ml" +# 8989 "parse.ml" : 'atomicPattern)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 0 : 'atomicTerm) in Obj.repr( -# 1425 "parse.mly" +# 1435 "parse.mly" (let (_1, q) = ((), _2) in ( mk_pattern (PatVQuote q) (rhs2 parseState 1 2) )) -# 8966 "parse.ml" +# 8997 "parse.ml" : 'atomicPattern)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'aqualifiedWithAttrs_lident_) in Obj.repr( -# 1428 "parse.mly" +# 1438 "parse.mly" (let qual_id = _1 in ( let (aqual, attrs), lid = qual_id in mk_pattern (PatVar (lid, aqual, attrs)) (rhs parseState 1) )) -# 8976 "parse.ml" +# 9007 "parse.ml" : 'atomicPattern)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'quident) in Obj.repr( -# 1433 "parse.mly" +# 1443 "parse.mly" (let uid = _1 in ( mk_pattern (PatName uid) (rhs parseState 1) )) -# 8984 "parse.ml" +# 9015 "parse.ml" : 'atomicPattern)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 2 : 'qlident) in let _3 = (Parsing.peek_val __caml_parser_env 0 : 'tuplePattern) in Obj.repr( -# 1438 "parse.mly" +# 1448 "parse.mly" (let (x, _2, y) = (_1, (), _3) in let p = ( (x, y) ) in ( p )) -# 8994 "parse.ml" +# 9025 "parse.ml" : 'fieldPattern)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'qlident) in Obj.repr( -# 1442 "parse.mly" +# 1452 "parse.mly" (let lid = _1 in ( lid, mk_pattern (PatVar (ident_of_lid lid, None, [])) (rhs parseState 1) )) -# 9002 "parse.ml" +# 9033 "parse.ml" : 'fieldPattern)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 3 : 'lidentOrUnderscore) in let _4 = (Parsing.peek_val __caml_parser_env 1 : 'simpleArrow) in Obj.repr( -# 1447 "parse.mly" +# 1457 "parse.mly" (let (_1, id, _3, t, _5) = ((), _2, (), _4, ()) in ( let r = rhs2 parseState 1 5 in let w = mk_pattern (PatVar (id, Some TypeClassArg, [])) r in let asc = (t, None) in [mk_pattern (PatAscribed(w, asc)) r] )) -# 9015 "parse.ml" +# 9046 "parse.ml" : 'patternOrMultibinder)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 1 : 'simpleArrow) in Obj.repr( -# 1454 "parse.mly" +# 1464 "parse.mly" (let (_1, t, _3) = ((), _2, ()) in ( let r = rhs2 parseState 1 3 in let id = gen r in @@ -9038,15 +9069,15 @@ let p = ( (x, y) ) in let asc = (t, None) in [mk_pattern (PatAscribed(w, asc)) r] )) -# 9028 "parse.ml" +# 9059 "parse.ml" : 'patternOrMultibinder)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'atomicPattern) in Obj.repr( -# 1462 "parse.mly" +# 1472 "parse.mly" (let pat = _1 in ( [pat] )) -# 9036 "parse.ml" +# 9067 "parse.ml" : 'patternOrMultibinder)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 5 : 'aqualifiedWithAttrs_lident_) in @@ -9054,7 +9085,7 @@ let p = ( (x, y) ) in let _5 = (Parsing.peek_val __caml_parser_env 2 : 'simpleArrow) in let _6 = (Parsing.peek_val __caml_parser_env 1 : 'refineOpt) in Obj.repr( -# 1465 "parse.mly" +# 1475 "parse.mly" (let (_1, qual_id0, qual_ids, _4, t, r, _7) = ((), _2, _3, (), _5, _6, ()) in ( let pos = rhs2 parseState 1 7 in @@ -9062,166 +9093,166 @@ let p = ( (x, y) ) in let qual_ids = qual_id0 :: qual_ids in List.map (fun ((aq, attrs), x) -> mkRefinedPattern (mk_pattern (PatVar (x, aq, attrs)) pos) t false r t_pos pos) qual_ids )) -# 9052 "parse.ml" +# 9083 "parse.ml" : 'patternOrMultibinder)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'aqualifiedWithAttrs_lidentOrUnderscore_) in Obj.repr( -# 1475 "parse.mly" +# 1485 "parse.mly" (let aqualifiedWithAttrs_lid = _1 in ( let (q, attrs), lid = aqualifiedWithAttrs_lid in mk_binder_with_attrs (Variable lid) (rhs parseState 1) Type_level q attrs )) -# 9063 "parse.ml" +# 9094 "parse.ml" : 'binder)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'tvar) in Obj.repr( -# 1481 "parse.mly" +# 1491 "parse.mly" (let tv = _1 in ( mk_binder (TVariable tv) (rhs parseState 1) Kind None )) -# 9071 "parse.ml" +# 9102 "parse.ml" : 'binder)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 3 : 'lidentOrUnderscore) in let _4 = (Parsing.peek_val __caml_parser_env 1 : 'simpleArrow) in Obj.repr( -# 1486 "parse.mly" +# 1496 "parse.mly" (let (_1, id, _3, t, _5) = ((), _2, (), _4, ()) in ( let r = rhs2 parseState 1 5 in [mk_binder (Annotated (id, t)) r Type_level (Some TypeClassArg)] )) -# 9082 "parse.ml" +# 9113 "parse.ml" : 'multiBinder)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 1 : 'simpleArrow) in Obj.repr( -# 1491 "parse.mly" +# 1501 "parse.mly" (let (_1, t, _3) = ((), _2, ()) in ( let r = rhs2 parseState 1 3 in let id = gen r in [mk_binder (Annotated (id, t)) r Type_level (Some TypeClassArg)] )) -# 9093 "parse.ml" +# 9124 "parse.ml" : 'multiBinder)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 4 : 'nonempty_list_aqualifiedWithAttrs_lidentOrUnderscore__) in let _4 = (Parsing.peek_val __caml_parser_env 2 : 'simpleArrow) in let _5 = (Parsing.peek_val __caml_parser_env 1 : 'refineOpt) in Obj.repr( -# 1497 "parse.mly" +# 1507 "parse.mly" (let (_1, qual_ids, _3, t, r, _6) = ((), _2, (), _4, _5, ()) in ( let should_bind_var = match qual_ids with | [ _ ] -> true | _ -> false in List.map (fun ((q, attrs), x) -> mkRefinedBinder x t should_bind_var r (rhs2 parseState 1 6) q attrs) qual_ids )) -# 9107 "parse.ml" +# 9138 "parse.ml" : 'multiBinder)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'list___anonymous_4_) in Obj.repr( -# 1506 "parse.mly" +# 1516 "parse.mly" (let bss = _1 in ( flatten bss )) -# 9115 "parse.ml" +# 9146 "parse.ml" : 'binders)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 2 : 'aqual) in let _2 = (Parsing.peek_val __caml_parser_env 1 : 'binderAttributes) in let _3 = (Parsing.peek_val __caml_parser_env 0 : FStar_Ident.ident) in Obj.repr( -# 1511 "parse.mly" +# 1521 "parse.mly" (let (aq, attrs, x) = (_1, _2, _3) in ( (Some aq, attrs), x )) -# 9125 "parse.ml" +# 9156 "parse.ml" : 'aqualifiedWithAttrs_lident_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 1 : 'aqual) in let _2 = (Parsing.peek_val __caml_parser_env 0 : FStar_Ident.ident) in Obj.repr( -# 1514 "parse.mly" +# 1524 "parse.mly" (let (aq, x) = (_1, _2) in ( (Some aq, []), x )) -# 9134 "parse.ml" +# 9165 "parse.ml" : 'aqualifiedWithAttrs_lident_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 1 : 'binderAttributes) in let _2 = (Parsing.peek_val __caml_parser_env 0 : FStar_Ident.ident) in Obj.repr( -# 1517 "parse.mly" +# 1527 "parse.mly" (let (attrs, x) = (_1, _2) in ( (None, attrs), x )) -# 9143 "parse.ml" +# 9174 "parse.ml" : 'aqualifiedWithAttrs_lident_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : FStar_Ident.ident) in Obj.repr( -# 1520 "parse.mly" +# 1530 "parse.mly" (let x = _1 in ( (None, []), x )) -# 9151 "parse.ml" +# 9182 "parse.ml" : 'aqualifiedWithAttrs_lident_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 2 : 'aqual) in let _2 = (Parsing.peek_val __caml_parser_env 1 : 'binderAttributes) in let _3 = (Parsing.peek_val __caml_parser_env 0 : FStar_Ident.ident) in Obj.repr( -# 1525 "parse.mly" +# 1535 "parse.mly" (let (aq, attrs, id) = (_1, _2, _3) in let x = ( id ) in ( (Some aq, attrs), x )) -# 9162 "parse.ml" +# 9193 "parse.ml" : 'aqualifiedWithAttrs_lidentOrOperator_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 4 : 'aqual) in let _2 = (Parsing.peek_val __caml_parser_env 3 : 'binderAttributes) in let _4 = (Parsing.peek_val __caml_parser_env 1 : string) in Obj.repr( -# 1529 "parse.mly" +# 1539 "parse.mly" (let (aq, attrs, _1, op, _3) = (_1, _2, (), _4, ()) in let x = let id = ( mk_ident (op, rhs parseState 1) ) in ( mk_ident (compile_op' (string_of_id id) (range_of_id id), range_of_id id) ) in ( (Some aq, attrs), x )) -# 9176 "parse.ml" +# 9207 "parse.ml" : 'aqualifiedWithAttrs_lidentOrOperator_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 4 : 'aqual) in let _2 = (Parsing.peek_val __caml_parser_env 3 : 'binderAttributes) in let _4 = (Parsing.peek_val __caml_parser_env 1 : 'binop_name) in Obj.repr( -# 1536 "parse.mly" +# 1546 "parse.mly" (let (aq, attrs, _1, op, _3) = (_1, _2, (), _4, ()) in let x = let id = ( op ) in ( mk_ident (compile_op' (string_of_id id) (range_of_id id), range_of_id id) ) in ( (Some aq, attrs), x )) -# 9190 "parse.ml" +# 9221 "parse.ml" : 'aqualifiedWithAttrs_lidentOrOperator_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 4 : 'aqual) in let _2 = (Parsing.peek_val __caml_parser_env 3 : 'binderAttributes) in let _4 = (Parsing.peek_val __caml_parser_env 1 : string) in Obj.repr( -# 1543 "parse.mly" +# 1553 "parse.mly" (let (aq, attrs, _1, op, _3) = (_1, _2, (), _4, ()) in let x = let id = ( mk_ident (op, rhs parseState 1) ) in ( mk_ident (compile_op' (string_of_id id) (range_of_id id), range_of_id id) ) in ( (Some aq, attrs), x )) -# 9204 "parse.ml" +# 9235 "parse.ml" : 'aqualifiedWithAttrs_lidentOrOperator_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 4 : 'aqual) in let _2 = (Parsing.peek_val __caml_parser_env 3 : 'binderAttributes) in let _4 = (Parsing.peek_val __caml_parser_env 1 : string) in Obj.repr( -# 1550 "parse.mly" +# 1560 "parse.mly" (let (aq, attrs, _1, op, _3) = (_1, _2, (), _4, ()) in let x = let id = @@ -9231,14 +9262,14 @@ let x = ( mk_ident (compile_op' (string_of_id id) (range_of_id id), range_of_id id) ) in ( (Some aq, attrs), x )) -# 9221 "parse.ml" +# 9252 "parse.ml" : 'aqualifiedWithAttrs_lidentOrOperator_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 4 : 'aqual) in let _2 = (Parsing.peek_val __caml_parser_env 3 : 'binderAttributes) in let _4 = (Parsing.peek_val __caml_parser_env 1 : string) in Obj.repr( -# 1560 "parse.mly" +# 1570 "parse.mly" (let (aq, attrs, _1, op, _3) = (_1, _2, (), _4, ()) in let x = let id = @@ -9248,62 +9279,62 @@ let x = ( mk_ident (compile_op' (string_of_id id) (range_of_id id), range_of_id id) ) in ( (Some aq, attrs), x )) -# 9238 "parse.ml" +# 9269 "parse.ml" : 'aqualifiedWithAttrs_lidentOrOperator_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 1 : 'aqual) in let _2 = (Parsing.peek_val __caml_parser_env 0 : FStar_Ident.ident) in Obj.repr( -# 1570 "parse.mly" +# 1580 "parse.mly" (let (aq, id) = (_1, _2) in let x = ( id ) in ( (Some aq, []), x )) -# 9248 "parse.ml" +# 9279 "parse.ml" : 'aqualifiedWithAttrs_lidentOrOperator_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 3 : 'aqual) in let _3 = (Parsing.peek_val __caml_parser_env 1 : string) in Obj.repr( -# 1574 "parse.mly" +# 1584 "parse.mly" (let (aq, _1, op, _3) = (_1, (), _3, ()) in let x = let id = ( mk_ident (op, rhs parseState 1) ) in ( mk_ident (compile_op' (string_of_id id) (range_of_id id), range_of_id id) ) in ( (Some aq, []), x )) -# 9261 "parse.ml" +# 9292 "parse.ml" : 'aqualifiedWithAttrs_lidentOrOperator_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 3 : 'aqual) in let _3 = (Parsing.peek_val __caml_parser_env 1 : 'binop_name) in Obj.repr( -# 1581 "parse.mly" +# 1591 "parse.mly" (let (aq, _1, op, _3) = (_1, (), _3, ()) in let x = let id = ( op ) in ( mk_ident (compile_op' (string_of_id id) (range_of_id id), range_of_id id) ) in ( (Some aq, []), x )) -# 9274 "parse.ml" +# 9305 "parse.ml" : 'aqualifiedWithAttrs_lidentOrOperator_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 3 : 'aqual) in let _3 = (Parsing.peek_val __caml_parser_env 1 : string) in Obj.repr( -# 1588 "parse.mly" +# 1598 "parse.mly" (let (aq, _1, op, _3) = (_1, (), _3, ()) in let x = let id = ( mk_ident (op, rhs parseState 1) ) in ( mk_ident (compile_op' (string_of_id id) (range_of_id id), range_of_id id) ) in ( (Some aq, []), x )) -# 9287 "parse.ml" +# 9318 "parse.ml" : 'aqualifiedWithAttrs_lidentOrOperator_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 3 : 'aqual) in let _3 = (Parsing.peek_val __caml_parser_env 1 : string) in Obj.repr( -# 1595 "parse.mly" +# 1605 "parse.mly" (let (aq, _1, op, _3) = (_1, (), _3, ()) in let x = let id = @@ -9313,13 +9344,13 @@ let x = ( mk_ident (compile_op' (string_of_id id) (range_of_id id), range_of_id id) ) in ( (Some aq, []), x )) -# 9303 "parse.ml" +# 9334 "parse.ml" : 'aqualifiedWithAttrs_lidentOrOperator_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 3 : 'aqual) in let _3 = (Parsing.peek_val __caml_parser_env 1 : string) in Obj.repr( -# 1605 "parse.mly" +# 1615 "parse.mly" (let (aq, _1, op, _3) = (_1, (), _3, ()) in let x = let id = @@ -9329,62 +9360,62 @@ let x = ( mk_ident (compile_op' (string_of_id id) (range_of_id id), range_of_id id) ) in ( (Some aq, []), x )) -# 9319 "parse.ml" +# 9350 "parse.ml" : 'aqualifiedWithAttrs_lidentOrOperator_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 1 : 'binderAttributes) in let _2 = (Parsing.peek_val __caml_parser_env 0 : FStar_Ident.ident) in Obj.repr( -# 1615 "parse.mly" +# 1625 "parse.mly" (let (attrs, id) = (_1, _2) in let x = ( id ) in ( (None, attrs), x )) -# 9329 "parse.ml" +# 9360 "parse.ml" : 'aqualifiedWithAttrs_lidentOrOperator_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 3 : 'binderAttributes) in let _3 = (Parsing.peek_val __caml_parser_env 1 : string) in Obj.repr( -# 1619 "parse.mly" +# 1629 "parse.mly" (let (attrs, _1, op, _3) = (_1, (), _3, ()) in let x = let id = ( mk_ident (op, rhs parseState 1) ) in ( mk_ident (compile_op' (string_of_id id) (range_of_id id), range_of_id id) ) in ( (None, attrs), x )) -# 9342 "parse.ml" +# 9373 "parse.ml" : 'aqualifiedWithAttrs_lidentOrOperator_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 3 : 'binderAttributes) in let _3 = (Parsing.peek_val __caml_parser_env 1 : 'binop_name) in Obj.repr( -# 1626 "parse.mly" +# 1636 "parse.mly" (let (attrs, _1, op, _3) = (_1, (), _3, ()) in let x = let id = ( op ) in ( mk_ident (compile_op' (string_of_id id) (range_of_id id), range_of_id id) ) in ( (None, attrs), x )) -# 9355 "parse.ml" +# 9386 "parse.ml" : 'aqualifiedWithAttrs_lidentOrOperator_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 3 : 'binderAttributes) in let _3 = (Parsing.peek_val __caml_parser_env 1 : string) in Obj.repr( -# 1633 "parse.mly" +# 1643 "parse.mly" (let (attrs, _1, op, _3) = (_1, (), _3, ()) in let x = let id = ( mk_ident (op, rhs parseState 1) ) in ( mk_ident (compile_op' (string_of_id id) (range_of_id id), range_of_id id) ) in ( (None, attrs), x )) -# 9368 "parse.ml" +# 9399 "parse.ml" : 'aqualifiedWithAttrs_lidentOrOperator_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 3 : 'binderAttributes) in let _3 = (Parsing.peek_val __caml_parser_env 1 : string) in Obj.repr( -# 1640 "parse.mly" +# 1650 "parse.mly" (let (attrs, _1, op, _3) = (_1, (), _3, ()) in let x = let id = @@ -9394,13 +9425,13 @@ let x = ( mk_ident (compile_op' (string_of_id id) (range_of_id id), range_of_id id) ) in ( (None, attrs), x )) -# 9384 "parse.ml" +# 9415 "parse.ml" : 'aqualifiedWithAttrs_lidentOrOperator_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 3 : 'binderAttributes) in let _3 = (Parsing.peek_val __caml_parser_env 1 : string) in Obj.repr( -# 1650 "parse.mly" +# 1660 "parse.mly" (let (attrs, _1, op, _3) = (_1, (), _3, ()) in let x = let id = @@ -9410,57 +9441,57 @@ let x = ( mk_ident (compile_op' (string_of_id id) (range_of_id id), range_of_id id) ) in ( (None, attrs), x )) -# 9400 "parse.ml" +# 9431 "parse.ml" : 'aqualifiedWithAttrs_lidentOrOperator_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : FStar_Ident.ident) in Obj.repr( -# 1660 "parse.mly" +# 1670 "parse.mly" (let id = _1 in let x = ( id ) in ( (None, []), x )) -# 9409 "parse.ml" +# 9440 "parse.ml" : 'aqualifiedWithAttrs_lidentOrOperator_)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 1 : string) in Obj.repr( -# 1664 "parse.mly" +# 1674 "parse.mly" (let (_1, op, _3) = ((), _2, ()) in let x = let id = ( mk_ident (op, rhs parseState 1) ) in ( mk_ident (compile_op' (string_of_id id) (range_of_id id), range_of_id id) ) in ( (None, []), x )) -# 9421 "parse.ml" +# 9452 "parse.ml" : 'aqualifiedWithAttrs_lidentOrOperator_)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 1 : 'binop_name) in Obj.repr( -# 1671 "parse.mly" +# 1681 "parse.mly" (let (_1, op, _3) = ((), _2, ()) in let x = let id = ( op ) in ( mk_ident (compile_op' (string_of_id id) (range_of_id id), range_of_id id) ) in ( (None, []), x )) -# 9433 "parse.ml" +# 9464 "parse.ml" : 'aqualifiedWithAttrs_lidentOrOperator_)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 1 : string) in Obj.repr( -# 1678 "parse.mly" +# 1688 "parse.mly" (let (_1, op, _3) = ((), _2, ()) in let x = let id = ( mk_ident (op, rhs parseState 1) ) in ( mk_ident (compile_op' (string_of_id id) (range_of_id id), range_of_id id) ) in ( (None, []), x )) -# 9445 "parse.ml" +# 9476 "parse.ml" : 'aqualifiedWithAttrs_lidentOrOperator_)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 1 : string) in Obj.repr( -# 1685 "parse.mly" +# 1695 "parse.mly" (let (_1, op, _3) = ((), _2, ()) in let x = let id = @@ -9470,12 +9501,12 @@ let x = ( mk_ident (compile_op' (string_of_id id) (range_of_id id), range_of_id id) ) in ( (None, []), x )) -# 9460 "parse.ml" +# 9491 "parse.ml" : 'aqualifiedWithAttrs_lidentOrOperator_)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 1 : string) in Obj.repr( -# 1695 "parse.mly" +# 1705 "parse.mly" (let (_1, op, _3) = ((), _2, ()) in let x = let id = @@ -9485,320 +9516,320 @@ let x = ( mk_ident (compile_op' (string_of_id id) (range_of_id id), range_of_id id) ) in ( (None, []), x )) -# 9475 "parse.ml" +# 9506 "parse.ml" : 'aqualifiedWithAttrs_lidentOrOperator_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 2 : 'aqual) in let _2 = (Parsing.peek_val __caml_parser_env 1 : 'binderAttributes) in let _3 = (Parsing.peek_val __caml_parser_env 0 : 'lidentOrUnderscore) in Obj.repr( -# 1707 "parse.mly" +# 1717 "parse.mly" (let (aq, attrs, x) = (_1, _2, _3) in ( (Some aq, attrs), x )) -# 9485 "parse.ml" +# 9516 "parse.ml" : 'aqualifiedWithAttrs_lidentOrUnderscore_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 1 : 'aqual) in let _2 = (Parsing.peek_val __caml_parser_env 0 : 'lidentOrUnderscore) in Obj.repr( -# 1710 "parse.mly" +# 1720 "parse.mly" (let (aq, x) = (_1, _2) in ( (Some aq, []), x )) -# 9494 "parse.ml" +# 9525 "parse.ml" : 'aqualifiedWithAttrs_lidentOrUnderscore_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 1 : 'binderAttributes) in let _2 = (Parsing.peek_val __caml_parser_env 0 : 'lidentOrUnderscore) in Obj.repr( -# 1713 "parse.mly" +# 1723 "parse.mly" (let (attrs, x) = (_1, _2) in ( (None, attrs), x )) -# 9503 "parse.ml" +# 9534 "parse.ml" : 'aqualifiedWithAttrs_lidentOrUnderscore_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'lidentOrUnderscore) in Obj.repr( -# 1716 "parse.mly" +# 1726 "parse.mly" (let x = _1 in ( (None, []), x )) -# 9511 "parse.ml" +# 9542 "parse.ml" : 'aqualifiedWithAttrs_lidentOrUnderscore_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'path_lident_) in Obj.repr( -# 1721 "parse.mly" +# 1731 "parse.mly" (let ids = _1 in ( lid_of_ids ids )) -# 9519 "parse.ml" +# 9550 "parse.ml" : 'qlident)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'path_uident_) in Obj.repr( -# 1726 "parse.mly" +# 1736 "parse.mly" (let ids = _1 in ( lid_of_ids ids )) -# 9527 "parse.ml" +# 9558 "parse.ml" : 'quident)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : FStar_Ident.ident) in Obj.repr( -# 1731 "parse.mly" +# 1741 "parse.mly" (let id = _1 in ( [id] )) -# 9535 "parse.ml" +# 9566 "parse.ml" : 'path_lident_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 2 : 'uident) in let _3 = (Parsing.peek_val __caml_parser_env 0 : 'path_lident_) in Obj.repr( -# 1734 "parse.mly" +# 1744 "parse.mly" (let (uid, _2, p) = (_1, (), _3) in ( uid::p )) -# 9544 "parse.ml" +# 9575 "parse.ml" : 'path_lident_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'uident) in Obj.repr( -# 1739 "parse.mly" +# 1749 "parse.mly" (let id = _1 in ( [id] )) -# 9552 "parse.ml" +# 9583 "parse.ml" : 'path_uident_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 2 : 'uident) in let _3 = (Parsing.peek_val __caml_parser_env 0 : 'path_uident_) in Obj.repr( -# 1742 "parse.mly" +# 1752 "parse.mly" (let (uid, _2, p) = (_1, (), _3) in ( uid::p )) -# 9561 "parse.ml" +# 9592 "parse.ml" : 'path_uident_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : FStar_Ident.ident) in Obj.repr( -# 1747 "parse.mly" +# 1757 "parse.mly" (let x = _1 in ( x )) -# 9569 "parse.ml" +# 9600 "parse.ml" : 'ident)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'uident) in Obj.repr( -# 1750 "parse.mly" +# 1760 "parse.mly" (let x = _1 in ( x )) -# 9577 "parse.ml" +# 9608 "parse.ml" : 'ident)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'qlident) in Obj.repr( -# 1755 "parse.mly" +# 1765 "parse.mly" (let qid = _1 in ( qid )) -# 9585 "parse.ml" +# 9616 "parse.ml" : 'qlidentOrOperator)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 1 : string) in Obj.repr( -# 1758 "parse.mly" +# 1768 "parse.mly" (let (_1, op, _3) = ((), _2, ()) in let id = ( mk_ident (op, rhs parseState 1) ) in ( lid_of_ns_and_id [] (id_of_text (compile_op' (string_of_id id) (range_of_id id))) )) -# 9594 "parse.ml" +# 9625 "parse.ml" : 'qlidentOrOperator)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 1 : 'binop_name) in Obj.repr( -# 1762 "parse.mly" +# 1772 "parse.mly" (let (_1, op, _3) = ((), _2, ()) in let id = ( op ) in ( lid_of_ns_and_id [] (id_of_text (compile_op' (string_of_id id) (range_of_id id))) )) -# 9603 "parse.ml" +# 9634 "parse.ml" : 'qlidentOrOperator)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 1 : string) in Obj.repr( -# 1766 "parse.mly" +# 1776 "parse.mly" (let (_1, op, _3) = ((), _2, ()) in let id = ( mk_ident (op, rhs parseState 1) ) in ( lid_of_ns_and_id [] (id_of_text (compile_op' (string_of_id id) (range_of_id id))) )) -# 9612 "parse.ml" +# 9643 "parse.ml" : 'qlidentOrOperator)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 1 : string) in Obj.repr( -# 1770 "parse.mly" +# 1780 "parse.mly" (let (_1, op, _3) = ((), _2, ()) in let id = let op = ( mk_ident ("and" ^ op, rhs parseState 1) ) in (op) in ( lid_of_ns_and_id [] (id_of_text (compile_op' (string_of_id id) (range_of_id id))) )) -# 9624 "parse.ml" +# 9655 "parse.ml" : 'qlidentOrOperator)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 1 : string) in Obj.repr( -# 1777 "parse.mly" +# 1787 "parse.mly" (let (_1, op, _3) = ((), _2, ()) in let id = let op = ( mk_ident ("let" ^ op, rhs parseState 1) ) in (op) in ( lid_of_ns_and_id [] (id_of_text (compile_op' (string_of_id id) (range_of_id id))) )) -# 9636 "parse.ml" +# 9667 "parse.ml" : 'qlidentOrOperator)) ; (fun __caml_parser_env -> Obj.repr( -# 1786 "parse.mly" +# 1796 "parse.mly" (let _1 = () in (None)) -# 9643 "parse.ml" +# 9674 "parse.ml" : 'matchMaybeOp)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : string) in Obj.repr( -# 1789 "parse.mly" +# 1799 "parse.mly" (let op = _1 in ( Some (mk_ident ("let" ^ op, rhs parseState 1)) )) -# 9651 "parse.ml" +# 9682 "parse.ml" : 'matchMaybeOp)) ; (fun __caml_parser_env -> Obj.repr( -# 1794 "parse.mly" +# 1804 "parse.mly" (let _1 = () in (None)) -# 9658 "parse.ml" +# 9689 "parse.ml" : 'ifMaybeOp)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : string) in Obj.repr( -# 1797 "parse.mly" +# 1807 "parse.mly" (let op = _1 in ( Some (mk_ident ("let" ^ op, rhs parseState 1)) )) -# 9666 "parse.ml" +# 9697 "parse.ml" : 'ifMaybeOp)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : string) in Obj.repr( -# 1802 "parse.mly" +# 1812 "parse.mly" (let id = _1 in ( mk_ident(id, rhs parseState 1))) -# 9674 "parse.ml" +# 9705 "parse.ml" : 'lidentOrUnderscore)) ; (fun __caml_parser_env -> Obj.repr( -# 1805 "parse.mly" +# 1815 "parse.mly" (let _1 = () in ( gen (rhs parseState 1) )) -# 9681 "parse.ml" +# 9712 "parse.ml" : 'lidentOrUnderscore)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : string) in Obj.repr( -# 1810 "parse.mly" +# 1820 "parse.mly" (let id = _1 in ( mk_ident(id, rhs parseState 1))) -# 9689 "parse.ml" +# 9720 "parse.ml" : FStar_Ident.ident)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : string) in Obj.repr( -# 1815 "parse.mly" +# 1825 "parse.mly" (let id = _1 in ( mk_ident(id, rhs parseState 1) )) -# 9697 "parse.ml" +# 9728 "parse.ml" : 'uident)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : string) in Obj.repr( -# 1820 "parse.mly" +# 1830 "parse.mly" (let tv = _1 in ( mk_ident(tv, rhs parseState 1) )) -# 9705 "parse.ml" +# 9736 "parse.ml" : 'tvar)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'atomicTerm) in Obj.repr( -# 1825 "parse.mly" +# 1835 "parse.mly" (let t = _1 in ( mk_term (Abs ([mk_pattern (PatWild (None, [])) (rhs parseState 3)], t)) (rhs parseState 3) Expr )) -# 9713 "parse.ml" +# 9744 "parse.ml" : 'thunk_atomicTerm_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'tmNoEq) in Obj.repr( -# 1830 "parse.mly" +# 1840 "parse.mly" (let t = _1 in ( mk_term (Abs ([mk_pattern (PatWild (None, [])) (rhs parseState 3)], t)) (rhs parseState 3) Expr )) -# 9721 "parse.ml" +# 9752 "parse.ml" : 'thunk_tmNoEq_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'typ) in Obj.repr( -# 1835 "parse.mly" +# 1845 "parse.mly" (let t = _1 in ( mk_term (Abs ([mk_pattern (PatWild (None, [])) (rhs parseState 3)], t)) (rhs parseState 3) Expr )) -# 9729 "parse.ml" +# 9760 "parse.ml" : 'thunk_typ_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'typ) in Obj.repr( -# 1840 "parse.mly" +# 1850 "parse.mly" (let t = _1 in ( let u = mk_term (Const Const_unit) (rhs parseState 3) Expr in let t = mk_term (Seq (u, t)) (rhs parseState 3) Expr in mk_term (Abs ([mk_pattern (PatWild (None, [])) (rhs parseState 3)], t)) (rhs parseState 3) Expr )) -# 9739 "parse.ml" +# 9770 "parse.ml" : 'thunk2_typ_)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 1 : 'tmArrow_tmNoEq_) in let _3 = (Parsing.peek_val __caml_parser_env 0 : 'option___anonymous_5_) in Obj.repr( -# 1847 "parse.mly" +# 1857 "parse.mly" (let (_1, t, tacopt) = ((), _2, _3) in ( t, tacopt )) -# 9748 "parse.ml" +# 9779 "parse.ml" : 'ascribeTyp)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 0 : 'kind) in Obj.repr( -# 1852 "parse.mly" +# 1862 "parse.mly" (let (_1, k) = ((), _2) in ( k )) -# 9756 "parse.ml" +# 9787 "parse.ml" : 'ascribeKind)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'tmArrow_tmNoEq_) in Obj.repr( -# 1857 "parse.mly" +# 1867 "parse.mly" (let t = _1 in ( {t with level=Kind} )) -# 9764 "parse.ml" +# 9795 "parse.ml" : 'kind)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'noSeqTerm) in Obj.repr( -# 1862 "parse.mly" +# 1872 "parse.mly" (let e = _1 in ( e )) -# 9772 "parse.ml" +# 9803 "parse.ml" : FStar_Parser_AST.term)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 2 : 'noSeqTerm) in let _3 = (Parsing.peek_val __caml_parser_env 0 : FStar_Parser_AST.term) in Obj.repr( -# 1865 "parse.mly" +# 1875 "parse.mly" (let (e1, _2, e2) = (_1, (), _3) in ( mk_term (Seq(e1, e2)) (rhs2 parseState 1 3) Expr )) -# 9781 "parse.ml" +# 9812 "parse.ml" : FStar_Parser_AST.term)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 2 : 'noSeqTerm) in let _2 = (Parsing.peek_val __caml_parser_env 1 : string option) in let _3 = (Parsing.peek_val __caml_parser_env 0 : FStar_Parser_AST.term) in Obj.repr( -# 1868 "parse.mly" +# 1878 "parse.mly" (let (e1, op, e2) = (_1, _2, _3) in ( let t = match op with | Some op -> @@ -9810,61 +9841,61 @@ in Bind(gen (rhs parseState 2), e1, e2) in mk_term t (rhs2 parseState 1 3) Expr )) -# 9800 "parse.ml" +# 9831 "parse.ml" : FStar_Parser_AST.term)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 4 : 'lidentOrUnderscore) in let _3 = (Parsing.peek_val __caml_parser_env 2 : 'noSeqTerm) in let _5 = (Parsing.peek_val __caml_parser_env 0 : FStar_Parser_AST.term) in Obj.repr( -# 1880 "parse.mly" +# 1890 "parse.mly" (let (x, _2, e1, _4, e2) = (_1, (), _3, (), _5) in ( log_issue (lhs parseState) (Warning_DeprecatedLightDoNotation, do_notation_deprecation_warning); mk_term (Bind(x, e1, e2)) (rhs2 parseState 1 5) Expr )) -# 9811 "parse.ml" +# 9842 "parse.ml" : FStar_Parser_AST.term)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 2 : 'option___anonymous_6_) in let _3 = (Parsing.peek_val __caml_parser_env 0 : 'tmIff) in Obj.repr( -# 1886 "parse.mly" +# 1896 "parse.mly" (let (as_opt, _2, t) = (_1, (), _3) in (as_opt,t,false)) -# 9820 "parse.ml" +# 9851 "parse.ml" : 'match_returning)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 2 : 'option___anonymous_7_) in let _3 = (Parsing.peek_val __caml_parser_env 0 : 'tmIff) in Obj.repr( -# 1889 "parse.mly" +# 1899 "parse.mly" (let (as_opt, _2, t) = (_1, (), _3) in (as_opt,t,true)) -# 9829 "parse.ml" +# 9860 "parse.ml" : 'match_returning)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'typ) in Obj.repr( -# 1894 "parse.mly" +# 1904 "parse.mly" (let t = _1 in ( t )) -# 9837 "parse.ml" +# 9868 "parse.ml" : 'noSeqTerm)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 3 : 'tmIff) in let _3 = (Parsing.peek_val __caml_parser_env 1 : 'tmIff) in let _4 = (Parsing.peek_val __caml_parser_env 0 : 'option___anonymous_8_) in Obj.repr( -# 1897 "parse.mly" +# 1907 "parse.mly" (let (e, _2, t, tactic_opt) = (_1, (), _3, _4) in ( mk_term (Ascribed(e,{t with level=Expr},tactic_opt,false)) (rhs2 parseState 1 4) Expr )) -# 9847 "parse.ml" +# 9878 "parse.ml" : 'noSeqTerm)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 3 : 'tmIff) in let _3 = (Parsing.peek_val __caml_parser_env 1 : 'tmIff) in let _4 = (Parsing.peek_val __caml_parser_env 0 : 'option___anonymous_9_) in Obj.repr( -# 1900 "parse.mly" +# 1910 "parse.mly" (let (e, _2, t, tactic_opt) = (_1, (), _3, _4) in ( log_issue (lhs parseState) @@ -9872,14 +9903,14 @@ in "Equality type ascriptions is an experimental feature subject to redesign in the future"); mk_term (Ascribed(e,{t with level=Expr},tactic_opt,true)) (rhs2 parseState 1 4) Expr )) -# 9862 "parse.ml" +# 9893 "parse.ml" : 'noSeqTerm)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 5 : 'atomicTermNotQUident) in let _3 = (Parsing.peek_val __caml_parser_env 3 : FStar_Parser_AST.term) in let _6 = (Parsing.peek_val __caml_parser_env 0 : 'noSeqTerm) in Obj.repr( -# 1908 "parse.mly" +# 1918 "parse.mly" (let (e1, _1, e, _3_inlined1, _3, e3) = (_1, (), _3, (), (), _6) in let op_expr = ( mk_ident (".()", rhs parseState 1), e, rhs2 parseState 1 3 ) in ( @@ -9887,14 +9918,14 @@ let op_expr = ( mk_ident (".()", rhs parseState 1), let opid = mk_ident (string_of_id op ^ "<-", range_of_id op) in mk_term (Op(opid, [ e1; e2; e3 ])) (rhs2 parseState 1 4) Expr )) -# 9877 "parse.ml" +# 9908 "parse.ml" : 'noSeqTerm)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 5 : 'atomicTermNotQUident) in let _3 = (Parsing.peek_val __caml_parser_env 3 : FStar_Parser_AST.term) in let _6 = (Parsing.peek_val __caml_parser_env 0 : 'noSeqTerm) in Obj.repr( -# 1916 "parse.mly" +# 1926 "parse.mly" (let (e1, _1, e, _3_inlined1, _3, e3) = (_1, (), _3, (), (), _6) in let op_expr = ( mk_ident (".[]", rhs parseState 1), e, rhs2 parseState 1 3 ) in ( @@ -9902,14 +9933,14 @@ let op_expr = ( mk_ident (".[]", rhs parseState 1), let opid = mk_ident (string_of_id op ^ "<-", range_of_id op) in mk_term (Op(opid, [ e1; e2; e3 ])) (rhs2 parseState 1 4) Expr )) -# 9892 "parse.ml" +# 9923 "parse.ml" : 'noSeqTerm)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 5 : 'atomicTermNotQUident) in let _3 = (Parsing.peek_val __caml_parser_env 3 : FStar_Parser_AST.term) in let _6 = (Parsing.peek_val __caml_parser_env 0 : 'noSeqTerm) in Obj.repr( -# 1924 "parse.mly" +# 1934 "parse.mly" (let (e1, _1, e, _3_inlined1, _3, e3) = (_1, (), _3, (), (), _6) in let op_expr = ( mk_ident (".[||]", rhs parseState 1), e, rhs2 parseState 1 3 ) in ( @@ -9917,14 +9948,14 @@ let op_expr = ( mk_ident (".[||]", rhs pars let opid = mk_ident (string_of_id op ^ "<-", range_of_id op) in mk_term (Op(opid, [ e1; e2; e3 ])) (rhs2 parseState 1 4) Expr )) -# 9907 "parse.ml" +# 9938 "parse.ml" : 'noSeqTerm)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 5 : 'atomicTermNotQUident) in let _3 = (Parsing.peek_val __caml_parser_env 3 : FStar_Parser_AST.term) in let _6 = (Parsing.peek_val __caml_parser_env 0 : 'noSeqTerm) in Obj.repr( -# 1932 "parse.mly" +# 1942 "parse.mly" (let (e1, _1, e, _3_inlined1, _3, e3) = (_1, (), _3, (), (), _6) in let op_expr = ( mk_ident (".(||)", rhs parseState 1), e, rhs2 parseState 1 3 ) in ( @@ -9932,36 +9963,36 @@ let op_expr = ( mk_ident (".(||) let opid = mk_ident (string_of_id op ^ "<-", range_of_id op) in mk_term (Op(opid, [ e1; e2; e3 ])) (rhs2 parseState 1 4) Expr )) -# 9922 "parse.ml" +# 9953 "parse.ml" : 'noSeqTerm)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 0 : 'typ) in Obj.repr( -# 1940 "parse.mly" +# 1950 "parse.mly" (let (_1, t) = ((), _2) in ( mk_term (Requires(t, None)) (rhs2 parseState 1 2) Type_level )) -# 9930 "parse.ml" +# 9961 "parse.ml" : 'noSeqTerm)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 0 : 'typ) in Obj.repr( -# 1943 "parse.mly" +# 1953 "parse.mly" (let (_1, t) = ((), _2) in ( mk_term (Ensures(t, None)) (rhs2 parseState 1 2) Type_level )) -# 9938 "parse.ml" +# 9969 "parse.ml" : 'noSeqTerm)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 0 : 'typ) in Obj.repr( -# 1946 "parse.mly" +# 1956 "parse.mly" (let (_1, t) = ((), _2) in ( mk_term (Decreases (t, None)) (rhs2 parseState 1 2) Type_level )) -# 9946 "parse.ml" +# 9977 "parse.ml" : 'noSeqTerm)) ; (fun __caml_parser_env -> let _3 = (Parsing.peek_val __caml_parser_env 1 : 'noSeqTerm) in Obj.repr( -# 1949 "parse.mly" +# 1959 "parse.mly" (let (_1, _2, t, _4) = ((), (), _3, ()) in ( match t.tm with | App (t1, t2, _) -> @@ -9970,15 +10001,15 @@ let op_expr = ( mk_ident (".(||) | _ -> raise_error (Fatal_SyntaxError, "Syntax error: To use well-founded relations, write e1 e2") (rhs parseState 3) )) -# 9960 "parse.ml" +# 9991 "parse.ml" : 'noSeqTerm)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 0 : 'nonempty_list_atomicTerm_) in Obj.repr( -# 1958 "parse.mly" +# 1968 "parse.mly" (let (_1, es) = ((), _2) in ( mk_term (Attributes es) (rhs2 parseState 1 2) Type_level )) -# 9968 "parse.ml" +# 9999 "parse.ml" : 'noSeqTerm)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 6 : 'ifMaybeOp) in @@ -9987,10 +10018,10 @@ let op_expr = ( mk_ident (".(||) let _5 = (Parsing.peek_val __caml_parser_env 2 : 'noSeqTerm) in let _7 = (Parsing.peek_val __caml_parser_env 0 : 'noSeqTerm) in Obj.repr( -# 1961 "parse.mly" +# 1971 "parse.mly" (let (op, e1, ret_opt, _4, e2, _6, e3) = (_1, _2, _3, (), _5, (), _7) in ( mk_term (If(e1, op, ret_opt, e2, e3)) (rhs2 parseState 1 7) Expr )) -# 9980 "parse.ml" +# 10011 "parse.ml" : 'noSeqTerm)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 4 : 'ifMaybeOp) in @@ -9998,26 +10029,26 @@ let op_expr = ( mk_ident (".(||) let _3 = (Parsing.peek_val __caml_parser_env 2 : 'option_match_returning_) in let _5 = (Parsing.peek_val __caml_parser_env 0 : 'noSeqTerm) in Obj.repr( -# 1964 "parse.mly" +# 1974 "parse.mly" (let (op, e1, ret_opt, _4, e2) = (_1, _2, _3, (), _5) in ( let e3 = mk_term (Const Const_unit) (rhs2 parseState 1 5) Expr in mk_term (If(e1, op, ret_opt, e2, e3)) (rhs2 parseState 1 5) Expr )) -# 9994 "parse.ml" +# 10025 "parse.ml" : 'noSeqTerm)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 2 : FStar_Parser_AST.term) in let _4 = (Parsing.peek_val __caml_parser_env 0 : 'reverse_left_flexible_nonempty_list_BAR_patternBranch_) in Obj.repr( -# 1970 "parse.mly" +# 1980 "parse.mly" (let (_1, e1, _3, xs) = ((), _2, (), _4) in let pbs = ( List.rev xs ) in ( let branches = focusBranches (pbs) (rhs2 parseState 1 4) in mk_term (TryWith(e1, branches)) (rhs2 parseState 1 4) Expr )) -# 10007 "parse.ml" +# 10038 "parse.ml" : 'noSeqTerm)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 4 : 'matchMaybeOp) in @@ -10025,21 +10056,21 @@ let pbs = ( List.rev xs ) in let _3 = (Parsing.peek_val __caml_parser_env 2 : 'option_match_returning_) in let _5 = (Parsing.peek_val __caml_parser_env 0 : 'reverse_left_flexible_list_BAR___anonymous_10_) in Obj.repr( -# 1977 "parse.mly" +# 1987 "parse.mly" (let (op, e, ret_opt, _4, xs) = (_1, _2, _3, (), _5) in let pbs = ( List.rev xs ) in ( let branches = focusBranches pbs (rhs2 parseState 1 5) in mk_term (Match(e, op, ret_opt, branches)) (rhs2 parseState 1 5) Expr )) -# 10022 "parse.ml" +# 10053 "parse.ml" : 'noSeqTerm)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 4 : bool) in let _3 = (Parsing.peek_val __caml_parser_env 2 : FStar_Parser_AST.term) in let _5 = (Parsing.peek_val __caml_parser_env 0 : FStar_Parser_AST.term) in Obj.repr( -# 1984 "parse.mly" +# 1994 "parse.mly" (let (_1, _2, t, _4, e) = (_1, (), _3, (), _5) in ( match t.tm with @@ -10055,7 +10086,7 @@ let pbs = ( List.rev xs ) in or, a record type with `let open e <: t in e'`") (rhs parseState 3) )) -# 10045 "parse.ml" +# 10076 "parse.ml" : 'noSeqTerm)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 5 : bool) in @@ -10064,7 +10095,7 @@ let pbs = ( List.rev xs ) in let _4 = (Parsing.peek_val __caml_parser_env 2 : 'list_attr_letbinding_) in let _6 = (Parsing.peek_val __caml_parser_env 0 : FStar_Parser_AST.term) in Obj.repr( -# 2000 "parse.mly" +# 2010 "parse.mly" (let (_2, q, lb, lbs, _6, e) = (_1, _2, _3, _4, (), _6) in let attrs = ( None ) in ( @@ -10072,7 +10103,7 @@ let attrs = ( None ) in let lbs = focusAttrLetBindings lbs (rhs2 parseState 2 3) in mk_term (Let(q, lbs, e)) (rhs2 parseState 1 6) Expr )) -# 10062 "parse.ml" +# 10093 "parse.ml" : 'noSeqTerm)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 6 : 'attribute) in @@ -10082,7 +10113,7 @@ let attrs = ( None ) in let _5 = (Parsing.peek_val __caml_parser_env 2 : 'list_attr_letbinding_) in let _7 = (Parsing.peek_val __caml_parser_env 0 : FStar_Parser_AST.term) in Obj.repr( -# 2008 "parse.mly" +# 2018 "parse.mly" (let (x, _2, q, lb, lbs, _6, e) = (_1, _2, _3, _4, _5, (), _7) in let attrs = ( Some x ) in ( @@ -10090,7 +10121,7 @@ let attrs = ( Some x ) in let lbs = focusAttrLetBindings lbs (rhs2 parseState 2 3) in mk_term (Let(q, lbs, e)) (rhs2 parseState 1 6) Expr )) -# 10080 "parse.ml" +# 10111 "parse.ml" : 'noSeqTerm)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 4 : string) in @@ -10098,42 +10129,42 @@ let attrs = ( Some x ) in let _3 = (Parsing.peek_val __caml_parser_env 2 : 'list___anonymous_11_) in let _5 = (Parsing.peek_val __caml_parser_env 0 : FStar_Parser_AST.term) in Obj.repr( -# 2016 "parse.mly" +# 2026 "parse.mly" (let (op, b, lbs, _4, e) = (_1, _2, _3, (), _5) in let op = ( mk_ident ("let" ^ op, rhs parseState 1) ) in ( let lbs = (op, b)::lbs in mk_term (LetOperator ( List.map (fun (op, (pat, tm)) -> (op, pat, tm)) lbs , e)) (rhs2 parseState 1 4) Expr )) -# 10095 "parse.ml" +# 10126 "parse.ml" : 'noSeqTerm)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 0 : 'reverse_left_flexible_nonempty_list_BAR_patternBranch_) in Obj.repr( -# 2023 "parse.mly" +# 2033 "parse.mly" (let (_1, xs) = ((), _2) in let pbs = ( List.rev xs ) in ( let branches = focusBranches pbs (rhs2 parseState 1 2) in mk_function branches (lhs parseState) (rhs2 parseState 1 2) )) -# 10107 "parse.ml" +# 10138 "parse.ml" : 'noSeqTerm)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 0 : 'atomicTerm) in Obj.repr( -# 2030 "parse.mly" +# 2040 "parse.mly" (let (_1, e) = ((), _2) in ( let a = set_lid_range assume_lid (rhs parseState 1) in mkExplicitApp (mk_term (Var a) (rhs parseState 1) Expr) [e] (rhs2 parseState 1 2) )) -# 10117 "parse.ml" +# 10148 "parse.ml" : 'noSeqTerm)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 1 : 'atomicTerm) in let _3 = (Parsing.peek_val __caml_parser_env 0 : 'option___anonymous_12_) in Obj.repr( -# 2035 "parse.mly" +# 2045 "parse.mly" (let (_1, e, tactic_opt) = ((), _2, _3) in ( match tactic_opt with @@ -10144,53 +10175,53 @@ let pbs = ( List.rev xs ) in let a = set_lid_range assert_by_tactic_lid (rhs parseState 1) in mkExplicitApp (mk_term (Var a) (rhs parseState 1) Expr) [e; tac] (rhs2 parseState 1 4) )) -# 10134 "parse.ml" +# 10165 "parse.ml" : 'noSeqTerm)) ; (fun __caml_parser_env -> let _3 = (Parsing.peek_val __caml_parser_env 0 : 'thunk_atomicTerm_) in Obj.repr( -# 2046 "parse.mly" +# 2056 "parse.mly" (let (_1, _2, tactic) = ((), (), _3) in ( let a = set_lid_range synth_lid (rhs parseState 1) in mkExplicitApp (mk_term (Var a) (rhs parseState 1) Expr) [tactic] (rhs2 parseState 1 2) )) -# 10145 "parse.ml" +# 10176 "parse.ml" : 'noSeqTerm)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 0 : 'atomicTerm) in Obj.repr( -# 2052 "parse.mly" +# 2062 "parse.mly" (let (_1, tactic) = ((), _2) in ( let a = set_lid_range synth_lid (rhs parseState 1) in mkExplicitApp (mk_term (Var a) (rhs parseState 1) Expr) [tactic] (rhs2 parseState 1 2) )) -# 10156 "parse.ml" +# 10187 "parse.ml" : 'noSeqTerm)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 5 : 'atomicTerm) in let _4 = (Parsing.peek_val __caml_parser_env 3 : 'noSeqTerm) in let _6 = (Parsing.peek_val __caml_parser_env 1 : 'list_calcStep_) in Obj.repr( -# 2058 "parse.mly" +# 2068 "parse.mly" (let (_1, rel, _3, init, _5, steps, _7) = ((), _2, (), _4, (), _6, ()) in ( mk_term (CalcProof (rel, init, steps)) (rhs2 parseState 1 7) Expr )) -# 10168 "parse.ml" +# 10199 "parse.ml" : 'noSeqTerm)) ; (fun __caml_parser_env -> let _3 = (Parsing.peek_val __caml_parser_env 4 : 'binders) in let _5 = (Parsing.peek_val __caml_parser_env 2 : 'noSeqTerm) in let _7 = (Parsing.peek_val __caml_parser_env 0 : 'noSeqTerm) in Obj.repr( -# 2063 "parse.mly" +# 2073 "parse.mly" (let (_1, _2, bs, _4, p, _6, e) = ((), (), _3, (), _5, (), _7) in ( mk_term (IntroForall(bs, p, e)) (rhs2 parseState 1 7) Expr )) -# 10180 "parse.ml" +# 10211 "parse.ml" : 'noSeqTerm)) ; (fun __caml_parser_env -> let _3 = (Parsing.peek_val __caml_parser_env 6 : 'binders) in @@ -10198,14 +10229,14 @@ let pbs = ( List.rev xs ) in let _7 = (Parsing.peek_val __caml_parser_env 2 : 'list_atomicTerm_) in let _9 = (Parsing.peek_val __caml_parser_env 0 : 'noSeqTerm) in Obj.repr( -# 2068 "parse.mly" +# 2078 "parse.mly" (let (_1, _2, bs, _4, p, _6, vs, _8, e) = ((), (), _3, (), _5, (), _7, (), _9) in ( if List.length bs <> List.length vs then raise_error (Fatal_SyntaxError, "Syntax error: expected instantiations for all binders") (rhs parseState 7) else mk_term (IntroExists(bs, p, vs, e)) (rhs2 parseState 1 9) Expr )) -# 10195 "parse.ml" +# 10226 "parse.ml" : 'noSeqTerm)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 6 : 'tmFormula) in @@ -10213,12 +10244,12 @@ let pbs = ( List.rev xs ) in let _6 = (Parsing.peek_val __caml_parser_env 2 : 'singleBinder) in let _8 = (Parsing.peek_val __caml_parser_env 0 : 'noSeqTerm) in Obj.repr( -# 2075 "parse.mly" +# 2085 "parse.mly" (let (_1, p, _3, q, _5, y, _7, e) = ((), _2, (), _4, (), _6, (), _8) in ( mk_term (IntroImplies(p, q, y, e)) (rhs2 parseState 1 8) Expr )) -# 10208 "parse.ml" +# 10239 "parse.ml" : 'noSeqTerm)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 5 : 'tmFormula) in @@ -10226,7 +10257,7 @@ let pbs = ( List.rev xs ) in let _6 = (Parsing.peek_val __caml_parser_env 1 : string) in let _7 = (Parsing.peek_val __caml_parser_env 0 : 'noSeqTerm) in Obj.repr( -# 2080 "parse.mly" +# 2090 "parse.mly" (let (_1, p, _3, q, _5, lr, e) = ((), _2, (), _4, (), _6, _7) in ( let b = @@ -10236,7 +10267,7 @@ let pbs = ( List.rev xs ) in in mk_term (IntroOr(b, p, q, e)) (rhs2 parseState 1 7) Expr )) -# 10226 "parse.ml" +# 10257 "parse.ml" : 'noSeqTerm)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 6 : 'tmConjunction) in @@ -10244,24 +10275,24 @@ let pbs = ( List.rev xs ) in let _6 = (Parsing.peek_val __caml_parser_env 2 : 'noSeqTerm) in let _8 = (Parsing.peek_val __caml_parser_env 0 : 'noSeqTerm) in Obj.repr( -# 2090 "parse.mly" +# 2100 "parse.mly" (let (_1, p, _3, q, _5, e1, _7, e2) = ((), _2, (), _4, (), _6, (), _8) in ( mk_term (IntroAnd(p, q, e1, e2)) (rhs2 parseState 1 8) Expr )) -# 10239 "parse.ml" +# 10270 "parse.ml" : 'noSeqTerm)) ; (fun __caml_parser_env -> let _3 = (Parsing.peek_val __caml_parser_env 4 : 'binders) in let _5 = (Parsing.peek_val __caml_parser_env 2 : 'noSeqTerm) in let _7 = (Parsing.peek_val __caml_parser_env 0 : 'list_atomicTerm_) in Obj.repr( -# 2095 "parse.mly" +# 2105 "parse.mly" (let (_1, _2, xs, _4, p, _6, vs) = ((), (), _3, (), _5, (), _7) in ( mk_term (ElimForall(xs, p, vs)) (rhs2 parseState 1 7) Expr )) -# 10251 "parse.ml" +# 10282 "parse.ml" : 'noSeqTerm)) ; (fun __caml_parser_env -> let _3 = (Parsing.peek_val __caml_parser_env 8 : 'binders) in @@ -10270,24 +10301,24 @@ let pbs = ( List.rev xs ) in let _9 = (Parsing.peek_val __caml_parser_env 2 : 'singleBinder) in let _11 = (Parsing.peek_val __caml_parser_env 0 : 'noSeqTerm) in Obj.repr( -# 2100 "parse.mly" +# 2110 "parse.mly" (let (_1, _2, bs, _4, p, _6, q, _8, y, _10, e) = ((), (), _3, (), _5, (), _7, (), _9, (), _11) in ( mk_term (ElimExists(bs, p, q, y, e)) (rhs2 parseState 1 11) Expr )) -# 10265 "parse.ml" +# 10296 "parse.ml" : 'noSeqTerm)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 4 : 'tmFormula) in let _4 = (Parsing.peek_val __caml_parser_env 2 : 'tmFormula) in let _6 = (Parsing.peek_val __caml_parser_env 0 : 'noSeqTerm) in Obj.repr( -# 2105 "parse.mly" +# 2115 "parse.mly" (let (_1, p, _3, q, _5, e) = ((), _2, (), _4, (), _6) in ( mk_term (ElimImplies(p, q, e)) (rhs2 parseState 1 6) Expr )) -# 10277 "parse.ml" +# 10308 "parse.ml" : 'noSeqTerm)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 12 : 'tmFormula) in @@ -10298,12 +10329,12 @@ let pbs = ( List.rev xs ) in let _12 = (Parsing.peek_val __caml_parser_env 2 : 'singleBinder) in let _14 = (Parsing.peek_val __caml_parser_env 0 : 'noSeqTerm) in Obj.repr( -# 2110 "parse.mly" +# 2120 "parse.mly" (let (_1, p, _3, q, _5, r, _7, x, _9, e1, _11, y, _13, e2) = ((), _2, (), _4, (), _6, (), _8, (), _10, (), _12, (), _14) in ( mk_term (ElimOr(p, q, r, x, e1, y, e2)) (rhs2 parseState 1 14) Expr )) -# 10293 "parse.ml" +# 10324 "parse.ml" : 'noSeqTerm)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 8 : 'tmConjunction) in @@ -10312,56 +10343,56 @@ let pbs = ( List.rev xs ) in let _8 = (Parsing.peek_val __caml_parser_env 2 : 'binders) in let _10 = (Parsing.peek_val __caml_parser_env 0 : 'noSeqTerm) in Obj.repr( -# 2115 "parse.mly" +# 2125 "parse.mly" (let (_1, p, _3, q, _5, r, _7, xs, _9, e) = ((), _2, (), _4, (), _6, (), _8, (), _10) in ( match xs with | [x;y] -> mk_term (ElimAnd(p, q, r, x, y, e)) (rhs2 parseState 1 10) Expr )) -# 10308 "parse.ml" +# 10339 "parse.ml" : 'noSeqTerm)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'binders) in Obj.repr( -# 2123 "parse.mly" +# 2133 "parse.mly" (let bs = _1 in ( match bs with | [b] -> b | _ -> raise_error (Fatal_SyntaxError, "Syntax error: expected a single binder") (rhs parseState 1) )) -# 10320 "parse.ml" +# 10351 "parse.ml" : 'singleBinder)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'binop_name) in Obj.repr( -# 2132 "parse.mly" +# 2142 "parse.mly" (let i = _1 in ( mk_term (Op (i, [])) (rhs parseState 1) Expr )) -# 10328 "parse.ml" +# 10359 "parse.ml" : 'calcRel)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 1 : 'qlident) in Obj.repr( -# 2135 "parse.mly" +# 2145 "parse.mly" (let (_1, id, _3) = ((), _2, ()) in ( mk_term (Var id) (rhs2 parseState 2 4) Un )) -# 10336 "parse.ml" +# 10367 "parse.ml" : 'calcRel)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'atomicTerm) in Obj.repr( -# 2138 "parse.mly" +# 2148 "parse.mly" (let t = _1 in ( t )) -# 10344 "parse.ml" +# 10375 "parse.ml" : 'calcRel)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 5 : 'calcRel) in let _3 = (Parsing.peek_val __caml_parser_env 3 : 'option_term_) in let _5 = (Parsing.peek_val __caml_parser_env 1 : 'noSeqTerm) in Obj.repr( -# 2143 "parse.mly" +# 2153 "parse.mly" (let (rel, _2, justif, _4, next, _6) = (_1, (), _3, (), _5, ()) in ( let justif = @@ -10371,22 +10402,22 @@ let pbs = ( List.rev xs ) in in CalcStep (rel, justif, next) )) -# 10361 "parse.ml" +# 10392 "parse.ml" : 'calcStep)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'simpleTerm) in Obj.repr( -# 2155 "parse.mly" +# 2165 "parse.mly" (let t = _1 in ( t )) -# 10369 "parse.ml" +# 10400 "parse.ml" : 'typ)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 3 : 'binders) in let _4 = (Parsing.peek_val __caml_parser_env 1 : 'trigger) in let _5 = (Parsing.peek_val __caml_parser_env 0 : 'noSeqTerm) in Obj.repr( -# 2158 "parse.mly" +# 2168 "parse.mly" (let (_1, bs, _3, trigger, e) = ((), _2, (), _4, _5) in let q = ( fun x -> QForall x ) in ( @@ -10397,14 +10428,14 @@ let q = ( fun x -> QForall x ) in let idents = idents_of_binders bs (rhs2 parseState 1 3) in mk_term (q (bs, (idents, trigger), e)) (rhs2 parseState 1 5) Formula )) -# 10387 "parse.ml" +# 10418 "parse.ml" : 'typ)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 3 : 'binders) in let _4 = (Parsing.peek_val __caml_parser_env 1 : 'trigger) in let _5 = (Parsing.peek_val __caml_parser_env 0 : 'noSeqTerm) in Obj.repr( -# 2169 "parse.mly" +# 2179 "parse.mly" (let (_1, bs, _3, trigger, e) = ((), _2, (), _4, _5) in let q = ( fun x -> QExists x) in ( @@ -10415,75 +10446,75 @@ let q = ( fun x -> QExists x) in let idents = idents_of_binders bs (rhs2 parseState 1 3) in mk_term (q (bs, (idents, trigger), e)) (rhs2 parseState 1 5) Formula )) -# 10405 "parse.ml" +# 10436 "parse.ml" : 'typ)) ; (fun __caml_parser_env -> Obj.repr( -# 2182 "parse.mly" +# 2192 "parse.mly" ( ( [] )) -# 10411 "parse.ml" +# 10442 "parse.ml" : 'trigger)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 1 : 'disjunctivePats) in Obj.repr( -# 2184 "parse.mly" +# 2194 "parse.mly" (let (_1, pats, _3) = ((), _2, ()) in ( pats )) -# 10419 "parse.ml" +# 10450 "parse.ml" : 'trigger)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'separated_nonempty_list_DISJUNCTION_conjunctivePat_) in Obj.repr( -# 2189 "parse.mly" +# 2199 "parse.mly" (let pats = _1 in ( pats )) -# 10427 "parse.ml" +# 10458 "parse.ml" : 'disjunctivePats)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'separated_nonempty_list_SEMICOLON_appTerm_) in Obj.repr( -# 2194 "parse.mly" +# 2204 "parse.mly" (let pats = _1 in ( pats )) -# 10435 "parse.ml" +# 10466 "parse.ml" : 'conjunctivePat)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'tmIff) in Obj.repr( -# 2199 "parse.mly" +# 2209 "parse.mly" (let e = _1 in ( e )) -# 10443 "parse.ml" +# 10474 "parse.ml" : 'simpleTerm)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 2 : 'nonempty_list_patternOrMultibinder_) in let _4 = (Parsing.peek_val __caml_parser_env 0 : FStar_Parser_AST.term) in Obj.repr( -# 2202 "parse.mly" +# 2212 "parse.mly" (let (_1, pats, _3, e) = ((), _2, (), _4) in ( mk_term (Abs(flatten pats, e)) (rhs2 parseState 1 4) Un )) -# 10452 "parse.ml" +# 10483 "parse.ml" : 'simpleTerm)) ; (fun __caml_parser_env -> Obj.repr( -# 2207 "parse.mly" +# 2217 "parse.mly" (let _1 = () in ( false )) -# 10459 "parse.ml" +# 10490 "parse.ml" : 'maybeFocusArrow)) ; (fun __caml_parser_env -> Obj.repr( -# 2210 "parse.mly" +# 2220 "parse.mly" (let _1 = () in ( true )) -# 10466 "parse.ml" +# 10497 "parse.ml" : 'maybeFocusArrow)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 2 : 'disjunctivePattern) in let _2 = (Parsing.peek_val __caml_parser_env 1 : 'maybeFocusArrow) in let _3 = (Parsing.peek_val __caml_parser_env 0 : FStar_Parser_AST.term) in Obj.repr( -# 2215 "parse.mly" +# 2225 "parse.mly" (let (pat, focus, e) = (_1, _2, _3) in let when_opt = ( None ) in ( @@ -10493,7 +10524,7 @@ let when_opt = ( None ) in in (focus, (pat, when_opt, e)) )) -# 10483 "parse.ml" +# 10514 "parse.ml" : 'patternBranch)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 4 : 'disjunctivePattern) in @@ -10501,7 +10532,7 @@ let when_opt = ( None ) in let _4 = (Parsing.peek_val __caml_parser_env 1 : 'maybeFocusArrow) in let _5 = (Parsing.peek_val __caml_parser_env 0 : FStar_Parser_AST.term) in Obj.repr( -# 2225 "parse.mly" +# 2235 "parse.mly" (let (pat, _1, e_inlined1, focus, e) = (_1, (), _3, _4, _5) in let when_opt = let e = e_inlined1 in @@ -10514,47 +10545,47 @@ in in (focus, (pat, when_opt, e)) )) -# 10504 "parse.ml" +# 10535 "parse.ml" : 'patternBranch)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 2 : 'tmImplies) in let _3 = (Parsing.peek_val __caml_parser_env 0 : 'tmIff) in Obj.repr( -# 2240 "parse.mly" +# 2250 "parse.mly" (let (e1, _2, e2) = (_1, (), _3) in ( mk_term (Op(mk_ident("<==>", rhs parseState 2), [e1; e2])) (rhs2 parseState 1 3) Formula )) -# 10513 "parse.ml" +# 10544 "parse.ml" : 'tmIff)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'tmImplies) in Obj.repr( -# 2243 "parse.mly" +# 2253 "parse.mly" (let e = _1 in ( e )) -# 10521 "parse.ml" +# 10552 "parse.ml" : 'tmIff)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 2 : 'tmArrow_tmFormula_) in let _3 = (Parsing.peek_val __caml_parser_env 0 : 'tmImplies) in Obj.repr( -# 2248 "parse.mly" +# 2258 "parse.mly" (let (e1, _2, e2) = (_1, (), _3) in ( mk_term (Op(mk_ident("==>", rhs parseState 2), [e1; e2])) (rhs2 parseState 1 3) Formula )) -# 10530 "parse.ml" +# 10561 "parse.ml" : 'tmImplies)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'tmArrow_tmFormula_) in Obj.repr( -# 2251 "parse.mly" +# 2261 "parse.mly" (let e = _1 in ( e )) -# 10538 "parse.ml" +# 10569 "parse.ml" : 'tmImplies)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 3 : 'tmFormula) in let _5 = (Parsing.peek_val __caml_parser_env 0 : 'tmArrow_tmFormula_) in Obj.repr( -# 2256 "parse.mly" +# 2266 "parse.mly" (let (_1, t, _3, _2, tgt) = ((), _2, (), (), _5) in let dom = ( ((Some TypeClassArg, []), t) ) in ( @@ -10565,14 +10596,14 @@ let dom = ( ((Some TypeClassArg, []), t) ) in in mk_term (Product([b], tgt)) (rhs2 parseState 1 3) Un )) -# 10555 "parse.ml" +# 10586 "parse.ml" : 'tmArrow_tmFormula_)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 4 : 'aqual) in let _3 = (Parsing.peek_val __caml_parser_env 3 : 'tmFormula) in let _6 = (Parsing.peek_val __caml_parser_env 0 : 'tmArrow_tmFormula_) in Obj.repr( -# 2267 "parse.mly" +# 2277 "parse.mly" (let (_1, q, dom_tm, _5, _2, tgt) = ((), _2, _3, (), (), _6) in let dom = let attrs_opt = ( None ) in @@ -10586,7 +10617,7 @@ in in mk_term (Product([b], tgt)) (rhs2 parseState 1 3) Un )) -# 10576 "parse.ml" +# 10607 "parse.ml" : 'tmArrow_tmFormula_)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 5 : 'aqual) in @@ -10594,7 +10625,7 @@ in let _4 = (Parsing.peek_val __caml_parser_env 3 : 'tmFormula) in let _7 = (Parsing.peek_val __caml_parser_env 0 : 'tmArrow_tmFormula_) in Obj.repr( -# 2281 "parse.mly" +# 2291 "parse.mly" (let (_1, q, x, dom_tm, _5, _2, tgt) = ((), _2, _3, _4, (), (), _7) in let dom = let attrs_opt = ( Some x ) in @@ -10608,13 +10639,13 @@ in in mk_term (Product([b], tgt)) (rhs2 parseState 1 3) Un )) -# 10598 "parse.ml" +# 10629 "parse.ml" : 'tmArrow_tmFormula_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 2 : 'tmFormula) in let _3 = (Parsing.peek_val __caml_parser_env 0 : 'tmArrow_tmFormula_) in Obj.repr( -# 2295 "parse.mly" +# 2305 "parse.mly" (let (dom_tm, _2, tgt) = (_1, (), _3) in let dom = let attrs_opt = ( None ) in @@ -10629,14 +10660,14 @@ in in mk_term (Product([b], tgt)) (rhs2 parseState 1 3) Un )) -# 10619 "parse.ml" +# 10650 "parse.ml" : 'tmArrow_tmFormula_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 3 : 'binderAttributes) in let _2 = (Parsing.peek_val __caml_parser_env 2 : 'tmFormula) in let _4 = (Parsing.peek_val __caml_parser_env 0 : 'tmArrow_tmFormula_) in Obj.repr( -# 2310 "parse.mly" +# 2320 "parse.mly" (let (x, dom_tm, _2, tgt) = (_1, _2, (), _4) in let dom = let attrs_opt = ( Some x ) in @@ -10651,14 +10682,14 @@ in in mk_term (Product([b], tgt)) (rhs2 parseState 1 3) Un )) -# 10641 "parse.ml" +# 10672 "parse.ml" : 'tmArrow_tmFormula_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 3 : 'aqual) in let _2 = (Parsing.peek_val __caml_parser_env 2 : 'tmFormula) in let _4 = (Parsing.peek_val __caml_parser_env 0 : 'tmArrow_tmFormula_) in Obj.repr( -# 2325 "parse.mly" +# 2335 "parse.mly" (let (x, dom_tm, _2, tgt) = (_1, _2, (), _4) in let dom = let attrs_opt = ( None ) in @@ -10673,7 +10704,7 @@ in in mk_term (Product([b], tgt)) (rhs2 parseState 1 3) Un )) -# 10663 "parse.ml" +# 10694 "parse.ml" : 'tmArrow_tmFormula_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 4 : 'aqual) in @@ -10681,7 +10712,7 @@ in let _3 = (Parsing.peek_val __caml_parser_env 2 : 'tmFormula) in let _5 = (Parsing.peek_val __caml_parser_env 0 : 'tmArrow_tmFormula_) in Obj.repr( -# 2340 "parse.mly" +# 2350 "parse.mly" (let (x, x_inlined1, dom_tm, _2, tgt) = (_1, _2, _3, (), _5) in let dom = let attrs_opt = @@ -10699,21 +10730,21 @@ in in mk_term (Product([b], tgt)) (rhs2 parseState 1 3) Un )) -# 10689 "parse.ml" +# 10720 "parse.ml" : 'tmArrow_tmFormula_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'tmFormula) in Obj.repr( -# 2358 "parse.mly" +# 2368 "parse.mly" (let e = _1 in ( e )) -# 10697 "parse.ml" +# 10728 "parse.ml" : 'tmArrow_tmFormula_)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 3 : 'tmNoEq) in let _5 = (Parsing.peek_val __caml_parser_env 0 : 'tmArrow_tmNoEq_) in Obj.repr( -# 2363 "parse.mly" +# 2373 "parse.mly" (let (_1, t, _3, _2, tgt) = ((), _2, (), (), _5) in let dom = ( ((Some TypeClassArg, []), t) ) in ( @@ -10724,14 +10755,14 @@ let dom = ( ((Some TypeClassArg, []), t) ) in in mk_term (Product([b], tgt)) (rhs2 parseState 1 3) Un )) -# 10714 "parse.ml" +# 10745 "parse.ml" : 'tmArrow_tmNoEq_)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 4 : 'aqual) in let _3 = (Parsing.peek_val __caml_parser_env 3 : 'tmNoEq) in let _6 = (Parsing.peek_val __caml_parser_env 0 : 'tmArrow_tmNoEq_) in Obj.repr( -# 2374 "parse.mly" +# 2384 "parse.mly" (let (_1, q, dom_tm, _5, _2, tgt) = ((), _2, _3, (), (), _6) in let dom = let attrs_opt = ( None ) in @@ -10745,7 +10776,7 @@ in in mk_term (Product([b], tgt)) (rhs2 parseState 1 3) Un )) -# 10735 "parse.ml" +# 10766 "parse.ml" : 'tmArrow_tmNoEq_)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 5 : 'aqual) in @@ -10753,7 +10784,7 @@ in let _4 = (Parsing.peek_val __caml_parser_env 3 : 'tmNoEq) in let _7 = (Parsing.peek_val __caml_parser_env 0 : 'tmArrow_tmNoEq_) in Obj.repr( -# 2388 "parse.mly" +# 2398 "parse.mly" (let (_1, q, x, dom_tm, _5, _2, tgt) = ((), _2, _3, _4, (), (), _7) in let dom = let attrs_opt = ( Some x ) in @@ -10767,13 +10798,13 @@ in in mk_term (Product([b], tgt)) (rhs2 parseState 1 3) Un )) -# 10757 "parse.ml" +# 10788 "parse.ml" : 'tmArrow_tmNoEq_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 2 : 'tmNoEq) in let _3 = (Parsing.peek_val __caml_parser_env 0 : 'tmArrow_tmNoEq_) in Obj.repr( -# 2402 "parse.mly" +# 2412 "parse.mly" (let (dom_tm, _2, tgt) = (_1, (), _3) in let dom = let attrs_opt = ( None ) in @@ -10788,14 +10819,14 @@ in in mk_term (Product([b], tgt)) (rhs2 parseState 1 3) Un )) -# 10778 "parse.ml" +# 10809 "parse.ml" : 'tmArrow_tmNoEq_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 3 : 'binderAttributes) in let _2 = (Parsing.peek_val __caml_parser_env 2 : 'tmNoEq) in let _4 = (Parsing.peek_val __caml_parser_env 0 : 'tmArrow_tmNoEq_) in Obj.repr( -# 2417 "parse.mly" +# 2427 "parse.mly" (let (x, dom_tm, _2, tgt) = (_1, _2, (), _4) in let dom = let attrs_opt = ( Some x ) in @@ -10810,14 +10841,14 @@ in in mk_term (Product([b], tgt)) (rhs2 parseState 1 3) Un )) -# 10800 "parse.ml" +# 10831 "parse.ml" : 'tmArrow_tmNoEq_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 3 : 'aqual) in let _2 = (Parsing.peek_val __caml_parser_env 2 : 'tmNoEq) in let _4 = (Parsing.peek_val __caml_parser_env 0 : 'tmArrow_tmNoEq_) in Obj.repr( -# 2432 "parse.mly" +# 2442 "parse.mly" (let (x, dom_tm, _2, tgt) = (_1, _2, (), _4) in let dom = let attrs_opt = ( None ) in @@ -10832,7 +10863,7 @@ in in mk_term (Product([b], tgt)) (rhs2 parseState 1 3) Un )) -# 10822 "parse.ml" +# 10853 "parse.ml" : 'tmArrow_tmNoEq_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 4 : 'aqual) in @@ -10840,7 +10871,7 @@ in let _3 = (Parsing.peek_val __caml_parser_env 2 : 'tmNoEq) in let _5 = (Parsing.peek_val __caml_parser_env 0 : 'tmArrow_tmNoEq_) in Obj.repr( -# 2447 "parse.mly" +# 2457 "parse.mly" (let (x, x_inlined1, dom_tm, _2, tgt) = (_1, _2, _3, (), _5) in let dom = let attrs_opt = @@ -10858,21 +10889,21 @@ in in mk_term (Product([b], tgt)) (rhs2 parseState 1 3) Un )) -# 10848 "parse.ml" +# 10879 "parse.ml" : 'tmArrow_tmNoEq_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'tmNoEq) in Obj.repr( -# 2465 "parse.mly" +# 2475 "parse.mly" (let e = _1 in ( e )) -# 10856 "parse.ml" +# 10887 "parse.ml" : 'tmArrow_tmNoEq_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 2 : 'simpleArrowDomain) in let _3 = (Parsing.peek_val __caml_parser_env 0 : 'simpleArrow) in Obj.repr( -# 2470 "parse.mly" +# 2480 "parse.mly" (let (dom, _2, tgt) = (_1, (), _3) in ( let ((aq_opt, attrs), dom_tm) = dom in @@ -10882,62 +10913,62 @@ in in mk_term (Product([b], tgt)) (rhs2 parseState 1 3) Un )) -# 10872 "parse.ml" +# 10903 "parse.ml" : 'simpleArrow)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'tmEqNoRefinement) in Obj.repr( -# 2480 "parse.mly" +# 2490 "parse.mly" (let e = _1 in ( e )) -# 10880 "parse.ml" +# 10911 "parse.ml" : 'simpleArrow)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 1 : 'tmEqNoRefinement) in Obj.repr( -# 2485 "parse.mly" +# 2495 "parse.mly" (let (_1, t, _3) = ((), _2, ()) in ( ((Some TypeClassArg, []), t) )) -# 10888 "parse.ml" +# 10919 "parse.ml" : 'simpleArrowDomain)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'tmEqNoRefinement) in Obj.repr( -# 2488 "parse.mly" +# 2498 "parse.mly" (let dom_tm = _1 in let attrs_opt = ( None ) in let aq_opt = ( None ) in ( (aq_opt, none_to_empty_list attrs_opt), dom_tm )) -# 10898 "parse.ml" +# 10929 "parse.ml" : 'simpleArrowDomain)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 1 : 'binderAttributes) in let _2 = (Parsing.peek_val __caml_parser_env 0 : 'tmEqNoRefinement) in Obj.repr( -# 2493 "parse.mly" +# 2503 "parse.mly" (let (x, dom_tm) = (_1, _2) in let attrs_opt = ( Some x ) in let aq_opt = ( None ) in ( (aq_opt, none_to_empty_list attrs_opt), dom_tm )) -# 10909 "parse.ml" +# 10940 "parse.ml" : 'simpleArrowDomain)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 1 : 'aqual) in let _2 = (Parsing.peek_val __caml_parser_env 0 : 'tmEqNoRefinement) in Obj.repr( -# 2498 "parse.mly" +# 2508 "parse.mly" (let (x, dom_tm) = (_1, _2) in let attrs_opt = ( None ) in let aq_opt = ( Some x ) in ( (aq_opt, none_to_empty_list attrs_opt), dom_tm )) -# 10920 "parse.ml" +# 10951 "parse.ml" : 'simpleArrowDomain)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 2 : 'aqual) in let _2 = (Parsing.peek_val __caml_parser_env 1 : 'binderAttributes) in let _3 = (Parsing.peek_val __caml_parser_env 0 : 'tmEqNoRefinement) in Obj.repr( -# 2503 "parse.mly" +# 2513 "parse.mly" (let (x, x_inlined1, dom_tm) = (_1, _2, _3) in let attrs_opt = let x = x_inlined1 in @@ -10945,370 +10976,370 @@ let attrs_opt = in let aq_opt = ( Some x ) in ( (aq_opt, none_to_empty_list attrs_opt), dom_tm )) -# 10935 "parse.ml" +# 10966 "parse.ml" : 'simpleArrowDomain)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 2 : 'tmFormula) in let _3 = (Parsing.peek_val __caml_parser_env 0 : 'tmConjunction) in Obj.repr( -# 2513 "parse.mly" +# 2523 "parse.mly" (let (e1, _2, e2) = (_1, (), _3) in ( mk_term (Op(mk_ident("\\/", rhs parseState 2), [e1;e2])) (rhs2 parseState 1 3) Formula )) -# 10944 "parse.ml" +# 10975 "parse.ml" : 'tmFormula)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'tmConjunction) in Obj.repr( -# 2516 "parse.mly" +# 2526 "parse.mly" (let e = _1 in ( e )) -# 10952 "parse.ml" +# 10983 "parse.ml" : 'tmFormula)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 2 : 'tmConjunction) in let _3 = (Parsing.peek_val __caml_parser_env 0 : 'tmTuple) in Obj.repr( -# 2521 "parse.mly" +# 2531 "parse.mly" (let (e1, _2, e2) = (_1, (), _3) in ( mk_term (Op(mk_ident("/\\", rhs parseState 2), [e1;e2])) (rhs2 parseState 1 3) Formula )) -# 10961 "parse.ml" +# 10992 "parse.ml" : 'tmConjunction)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'tmTuple) in Obj.repr( -# 2524 "parse.mly" +# 2534 "parse.mly" (let e = _1 in ( e )) -# 10969 "parse.ml" +# 11000 "parse.ml" : 'tmConjunction)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'separated_nonempty_list_COMMA_tmEq_) in Obj.repr( -# 2529 "parse.mly" +# 2539 "parse.mly" (let el = _1 in ( match el with | [x] -> x | components -> mkTuple components (rhs2 parseState 1 1) )) -# 10981 "parse.ml" +# 11012 "parse.ml" : 'tmTuple)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 2 : 'tmEqWith_appTermNoRecordExp_) in let _3 = (Parsing.peek_val __caml_parser_env 0 : 'tmEqWith_appTermNoRecordExp_) in Obj.repr( -# 2538 "parse.mly" +# 2548 "parse.mly" (let (e1, _2, e2) = (_1, (), _3) in ( mk_term (Op(mk_ident("=", rhs parseState 2), [e1; e2])) (rhs2 parseState 1 3) Un)) -# 10990 "parse.ml" +# 11021 "parse.ml" : 'tmEqWith_appTermNoRecordExp_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 2 : 'tmEqWith_appTermNoRecordExp_) in let _3 = (Parsing.peek_val __caml_parser_env 0 : 'tmEqWith_appTermNoRecordExp_) in Obj.repr( -# 2541 "parse.mly" +# 2551 "parse.mly" (let (e1, _2, e2) = (_1, (), _3) in ( mk_term (Op(mk_ident(":=", rhs parseState 2), [e1; e2])) (rhs2 parseState 1 3) Un)) -# 10999 "parse.ml" +# 11030 "parse.ml" : 'tmEqWith_appTermNoRecordExp_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 2 : 'tmEqWith_appTermNoRecordExp_) in let _3 = (Parsing.peek_val __caml_parser_env 0 : 'tmEqWith_appTermNoRecordExp_) in Obj.repr( -# 2544 "parse.mly" +# 2554 "parse.mly" (let (e1, _2, e2) = (_1, (), _3) in ( mk_term (Op(mk_ident("|>", rhs parseState 2), [e1; e2])) (rhs2 parseState 1 3) Un)) -# 11008 "parse.ml" +# 11039 "parse.ml" : 'tmEqWith_appTermNoRecordExp_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 2 : 'tmEqWith_appTermNoRecordExp_) in let _2 = (Parsing.peek_val __caml_parser_env 1 : string) in let _3 = (Parsing.peek_val __caml_parser_env 0 : 'tmEqWith_appTermNoRecordExp_) in Obj.repr( -# 2547 "parse.mly" +# 2557 "parse.mly" (let (e1, op, e2) = (_1, _2, _3) in let op = ( mk_ident (op, rhs parseState 1) ) in ( mk_term (Op(op, [e1; e2])) (rhs2 parseState 1 3) Un)) -# 11019 "parse.ml" +# 11050 "parse.ml" : 'tmEqWith_appTermNoRecordExp_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 2 : 'tmEqWith_appTermNoRecordExp_) in let _2 = (Parsing.peek_val __caml_parser_env 1 : string) in let _3 = (Parsing.peek_val __caml_parser_env 0 : 'tmEqWith_appTermNoRecordExp_) in Obj.repr( -# 2551 "parse.mly" +# 2561 "parse.mly" (let (e1, op, e2) = (_1, _2, _3) in let op = ( mk_ident (op, rhs parseState 1) ) in ( mk_term (Op(op, [e1; e2])) (rhs2 parseState 1 3) Un)) -# 11030 "parse.ml" +# 11061 "parse.ml" : 'tmEqWith_appTermNoRecordExp_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 2 : 'tmEqWith_appTermNoRecordExp_) in let _2 = (Parsing.peek_val __caml_parser_env 1 : string) in let _3 = (Parsing.peek_val __caml_parser_env 0 : 'tmEqWith_appTermNoRecordExp_) in Obj.repr( -# 2555 "parse.mly" +# 2565 "parse.mly" (let (e1, op, e2) = (_1, _2, _3) in let op = ( mk_ident (op, rhs parseState 1) ) in ( mk_term (Op(op, [e1; e2])) (rhs2 parseState 1 3) Un)) -# 11041 "parse.ml" +# 11072 "parse.ml" : 'tmEqWith_appTermNoRecordExp_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 2 : 'tmEqWith_appTermNoRecordExp_) in let _2 = (Parsing.peek_val __caml_parser_env 1 : string) in let _3 = (Parsing.peek_val __caml_parser_env 0 : 'tmEqWith_appTermNoRecordExp_) in Obj.repr( -# 2559 "parse.mly" +# 2569 "parse.mly" (let (e1, op, e2) = (_1, _2, _3) in let op = ( mk_ident (op, rhs parseState 1) ) in ( mk_term (Op(op, [e1; e2])) (rhs2 parseState 1 3) Un)) -# 11052 "parse.ml" +# 11083 "parse.ml" : 'tmEqWith_appTermNoRecordExp_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 2 : 'tmEqWith_appTermNoRecordExp_) in let _2 = (Parsing.peek_val __caml_parser_env 1 : string) in let _3 = (Parsing.peek_val __caml_parser_env 0 : 'tmEqWith_appTermNoRecordExp_) in Obj.repr( -# 2563 "parse.mly" +# 2573 "parse.mly" (let (e1, op, e2) = (_1, _2, _3) in let op = ( mk_ident (op, rhs parseState 1) ) in ( mk_term (Op(op, [e1; e2])) (rhs2 parseState 1 3) Un)) -# 11063 "parse.ml" +# 11094 "parse.ml" : 'tmEqWith_appTermNoRecordExp_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 2 : 'tmEqWith_appTermNoRecordExp_) in let _2 = (Parsing.peek_val __caml_parser_env 1 : string) in let _3 = (Parsing.peek_val __caml_parser_env 0 : 'tmEqWith_appTermNoRecordExp_) in Obj.repr( -# 2567 "parse.mly" +# 2577 "parse.mly" (let (e1, op, e2) = (_1, _2, _3) in let op = ( mk_ident (op, rhs parseState 1) ) in ( mk_term (Op(op, [e1; e2])) (rhs2 parseState 1 3) Un)) -# 11074 "parse.ml" +# 11105 "parse.ml" : 'tmEqWith_appTermNoRecordExp_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 2 : 'tmEqWith_appTermNoRecordExp_) in let _3 = (Parsing.peek_val __caml_parser_env 0 : 'tmEqWith_appTermNoRecordExp_) in Obj.repr( -# 2571 "parse.mly" +# 2581 "parse.mly" (let (e1, _2, e2) = (_1, (), _3) in ( mk_term (Op(mk_ident("-", rhs parseState 2), [e1; e2])) (rhs2 parseState 1 3) Un)) -# 11083 "parse.ml" +# 11114 "parse.ml" : 'tmEqWith_appTermNoRecordExp_)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 0 : 'tmEqWith_appTermNoRecordExp_) in Obj.repr( -# 2574 "parse.mly" +# 2584 "parse.mly" (let (_1, e) = ((), _2) in ( mk_uminus e (rhs parseState 1) (rhs2 parseState 1 2) Expr )) -# 11091 "parse.ml" +# 11122 "parse.ml" : 'tmEqWith_appTermNoRecordExp_)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 0 : 'tmEqWith_appTermNoRecordExp_) in Obj.repr( -# 2577 "parse.mly" +# 2587 "parse.mly" (let (_1, e) = ((), _2) in ( mk_term (Quote (e, Dynamic)) (rhs2 parseState 1 3) Un )) -# 11099 "parse.ml" +# 11130 "parse.ml" : 'tmEqWith_appTermNoRecordExp_)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 0 : 'tmEqWith_appTermNoRecordExp_) in Obj.repr( -# 2580 "parse.mly" +# 2590 "parse.mly" (let (_1, e) = ((), _2) in ( mk_term (Quote (e, Static)) (rhs2 parseState 1 3) Un )) -# 11107 "parse.ml" +# 11138 "parse.ml" : 'tmEqWith_appTermNoRecordExp_)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 0 : 'atomicTerm) in Obj.repr( -# 2583 "parse.mly" +# 2593 "parse.mly" (let (_1, e) = ((), _2) in ( let q = mk_term (Quote (e, Dynamic)) (rhs2 parseState 1 3) Un in mk_term (Antiquote q) (rhs2 parseState 1 3) Un )) -# 11116 "parse.ml" +# 11147 "parse.ml" : 'tmEqWith_appTermNoRecordExp_)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 0 : 'atomicTerm) in Obj.repr( -# 2587 "parse.mly" +# 2597 "parse.mly" (let (_1, e) = ((), _2) in ( mk_term (Antiquote e) (rhs2 parseState 1 3) Un )) -# 11124 "parse.ml" +# 11155 "parse.ml" : 'tmEqWith_appTermNoRecordExp_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'tmNoEqWith_appTermNoRecordExp_) in Obj.repr( -# 2590 "parse.mly" +# 2600 "parse.mly" (let e = _1 in ( e )) -# 11132 "parse.ml" +# 11163 "parse.ml" : 'tmEqWith_appTermNoRecordExp_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 2 : 'tmEqWith_tmRefinement_) in let _3 = (Parsing.peek_val __caml_parser_env 0 : 'tmEqWith_tmRefinement_) in Obj.repr( -# 2595 "parse.mly" +# 2605 "parse.mly" (let (e1, _2, e2) = (_1, (), _3) in ( mk_term (Op(mk_ident("=", rhs parseState 2), [e1; e2])) (rhs2 parseState 1 3) Un)) -# 11141 "parse.ml" +# 11172 "parse.ml" : 'tmEqWith_tmRefinement_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 2 : 'tmEqWith_tmRefinement_) in let _3 = (Parsing.peek_val __caml_parser_env 0 : 'tmEqWith_tmRefinement_) in Obj.repr( -# 2598 "parse.mly" +# 2608 "parse.mly" (let (e1, _2, e2) = (_1, (), _3) in ( mk_term (Op(mk_ident(":=", rhs parseState 2), [e1; e2])) (rhs2 parseState 1 3) Un)) -# 11150 "parse.ml" +# 11181 "parse.ml" : 'tmEqWith_tmRefinement_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 2 : 'tmEqWith_tmRefinement_) in let _3 = (Parsing.peek_val __caml_parser_env 0 : 'tmEqWith_tmRefinement_) in Obj.repr( -# 2601 "parse.mly" +# 2611 "parse.mly" (let (e1, _2, e2) = (_1, (), _3) in ( mk_term (Op(mk_ident("|>", rhs parseState 2), [e1; e2])) (rhs2 parseState 1 3) Un)) -# 11159 "parse.ml" +# 11190 "parse.ml" : 'tmEqWith_tmRefinement_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 2 : 'tmEqWith_tmRefinement_) in let _2 = (Parsing.peek_val __caml_parser_env 1 : string) in let _3 = (Parsing.peek_val __caml_parser_env 0 : 'tmEqWith_tmRefinement_) in Obj.repr( -# 2604 "parse.mly" +# 2614 "parse.mly" (let (e1, op, e2) = (_1, _2, _3) in let op = ( mk_ident (op, rhs parseState 1) ) in ( mk_term (Op(op, [e1; e2])) (rhs2 parseState 1 3) Un)) -# 11170 "parse.ml" +# 11201 "parse.ml" : 'tmEqWith_tmRefinement_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 2 : 'tmEqWith_tmRefinement_) in let _2 = (Parsing.peek_val __caml_parser_env 1 : string) in let _3 = (Parsing.peek_val __caml_parser_env 0 : 'tmEqWith_tmRefinement_) in Obj.repr( -# 2608 "parse.mly" +# 2618 "parse.mly" (let (e1, op, e2) = (_1, _2, _3) in let op = ( mk_ident (op, rhs parseState 1) ) in ( mk_term (Op(op, [e1; e2])) (rhs2 parseState 1 3) Un)) -# 11181 "parse.ml" +# 11212 "parse.ml" : 'tmEqWith_tmRefinement_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 2 : 'tmEqWith_tmRefinement_) in let _2 = (Parsing.peek_val __caml_parser_env 1 : string) in let _3 = (Parsing.peek_val __caml_parser_env 0 : 'tmEqWith_tmRefinement_) in Obj.repr( -# 2612 "parse.mly" +# 2622 "parse.mly" (let (e1, op, e2) = (_1, _2, _3) in let op = ( mk_ident (op, rhs parseState 1) ) in ( mk_term (Op(op, [e1; e2])) (rhs2 parseState 1 3) Un)) -# 11192 "parse.ml" +# 11223 "parse.ml" : 'tmEqWith_tmRefinement_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 2 : 'tmEqWith_tmRefinement_) in let _2 = (Parsing.peek_val __caml_parser_env 1 : string) in let _3 = (Parsing.peek_val __caml_parser_env 0 : 'tmEqWith_tmRefinement_) in Obj.repr( -# 2616 "parse.mly" +# 2626 "parse.mly" (let (e1, op, e2) = (_1, _2, _3) in let op = ( mk_ident (op, rhs parseState 1) ) in ( mk_term (Op(op, [e1; e2])) (rhs2 parseState 1 3) Un)) -# 11203 "parse.ml" +# 11234 "parse.ml" : 'tmEqWith_tmRefinement_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 2 : 'tmEqWith_tmRefinement_) in let _2 = (Parsing.peek_val __caml_parser_env 1 : string) in let _3 = (Parsing.peek_val __caml_parser_env 0 : 'tmEqWith_tmRefinement_) in Obj.repr( -# 2620 "parse.mly" +# 2630 "parse.mly" (let (e1, op, e2) = (_1, _2, _3) in let op = ( mk_ident (op, rhs parseState 1) ) in ( mk_term (Op(op, [e1; e2])) (rhs2 parseState 1 3) Un)) -# 11214 "parse.ml" +# 11245 "parse.ml" : 'tmEqWith_tmRefinement_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 2 : 'tmEqWith_tmRefinement_) in let _2 = (Parsing.peek_val __caml_parser_env 1 : string) in let _3 = (Parsing.peek_val __caml_parser_env 0 : 'tmEqWith_tmRefinement_) in Obj.repr( -# 2624 "parse.mly" +# 2634 "parse.mly" (let (e1, op, e2) = (_1, _2, _3) in let op = ( mk_ident (op, rhs parseState 1) ) in ( mk_term (Op(op, [e1; e2])) (rhs2 parseState 1 3) Un)) -# 11225 "parse.ml" +# 11256 "parse.ml" : 'tmEqWith_tmRefinement_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 2 : 'tmEqWith_tmRefinement_) in let _3 = (Parsing.peek_val __caml_parser_env 0 : 'tmEqWith_tmRefinement_) in Obj.repr( -# 2628 "parse.mly" +# 2638 "parse.mly" (let (e1, _2, e2) = (_1, (), _3) in ( mk_term (Op(mk_ident("-", rhs parseState 2), [e1; e2])) (rhs2 parseState 1 3) Un)) -# 11234 "parse.ml" +# 11265 "parse.ml" : 'tmEqWith_tmRefinement_)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 0 : 'tmEqWith_tmRefinement_) in Obj.repr( -# 2631 "parse.mly" +# 2641 "parse.mly" (let (_1, e) = ((), _2) in ( mk_uminus e (rhs parseState 1) (rhs2 parseState 1 2) Expr )) -# 11242 "parse.ml" +# 11273 "parse.ml" : 'tmEqWith_tmRefinement_)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 0 : 'tmEqWith_tmRefinement_) in Obj.repr( -# 2634 "parse.mly" +# 2644 "parse.mly" (let (_1, e) = ((), _2) in ( mk_term (Quote (e, Dynamic)) (rhs2 parseState 1 3) Un )) -# 11250 "parse.ml" +# 11281 "parse.ml" : 'tmEqWith_tmRefinement_)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 0 : 'tmEqWith_tmRefinement_) in Obj.repr( -# 2637 "parse.mly" +# 2647 "parse.mly" (let (_1, e) = ((), _2) in ( mk_term (Quote (e, Static)) (rhs2 parseState 1 3) Un )) -# 11258 "parse.ml" +# 11289 "parse.ml" : 'tmEqWith_tmRefinement_)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 0 : 'atomicTerm) in Obj.repr( -# 2640 "parse.mly" +# 2650 "parse.mly" (let (_1, e) = ((), _2) in ( let q = mk_term (Quote (e, Dynamic)) (rhs2 parseState 1 3) Un in mk_term (Antiquote q) (rhs2 parseState 1 3) Un )) -# 11267 "parse.ml" +# 11298 "parse.ml" : 'tmEqWith_tmRefinement_)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 0 : 'atomicTerm) in Obj.repr( -# 2644 "parse.mly" +# 2654 "parse.mly" (let (_1, e) = ((), _2) in ( mk_term (Antiquote e) (rhs2 parseState 1 3) Un )) -# 11275 "parse.ml" +# 11306 "parse.ml" : 'tmEqWith_tmRefinement_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'tmNoEqWith_tmRefinement_) in Obj.repr( -# 2647 "parse.mly" +# 2657 "parse.mly" (let e = _1 in ( e )) -# 11283 "parse.ml" +# 11314 "parse.ml" : 'tmEqWith_tmRefinement_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 2 : 'tmNoEqWith_appTermNoRecordExp_) in let _3 = (Parsing.peek_val __caml_parser_env 0 : 'tmNoEqWith_appTermNoRecordExp_) in Obj.repr( -# 2652 "parse.mly" +# 2662 "parse.mly" (let (e1, _2, e2) = (_1, (), _3) in ( consTerm (rhs parseState 2) e1 e2 )) -# 11292 "parse.ml" +# 11323 "parse.ml" : 'tmNoEqWith_appTermNoRecordExp_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 2 : 'tmNoEqWith_appTermNoRecordExp_) in let _3 = (Parsing.peek_val __caml_parser_env 0 : 'tmNoEqWith_appTermNoRecordExp_) in Obj.repr( -# 2655 "parse.mly" +# 2665 "parse.mly" (let (e1, _2, e2) = (_1, (), _3) in ( let dom = @@ -11327,85 +11358,85 @@ let op = ( mk_ident (op, rhs parseState 1) ) in in mk_term (Sum(dom, res)) (rhs2 parseState 1 3) Type_level )) -# 11317 "parse.ml" +# 11348 "parse.ml" : 'tmNoEqWith_appTermNoRecordExp_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 2 : 'tmNoEqWith_appTermNoRecordExp_) in let _2 = (Parsing.peek_val __caml_parser_env 1 : string) in let _3 = (Parsing.peek_val __caml_parser_env 0 : 'tmNoEqWith_appTermNoRecordExp_) in Obj.repr( -# 2674 "parse.mly" +# 2684 "parse.mly" (let (e1, op, e2) = (_1, _2, _3) in ( mk_term (Op(mk_ident(op, rhs parseState 2), [e1; e2])) (rhs2 parseState 1 3) Un)) -# 11327 "parse.ml" +# 11358 "parse.ml" : 'tmNoEqWith_appTermNoRecordExp_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 4 : 'tmNoEqWith_appTermNoRecordExp_) in let _3 = (Parsing.peek_val __caml_parser_env 2 : 'tmNoEqWith_appTermNoRecordExp_) in let _5 = (Parsing.peek_val __caml_parser_env 0 : 'tmNoEqWith_appTermNoRecordExp_) in Obj.repr( -# 2677 "parse.mly" +# 2687 "parse.mly" (let (e1, _2, op, _4, e2) = (_1, (), _3, (), _5) in ( mkApp op [ e1, Infix; e2, Nothing ] (rhs2 parseState 1 5) )) -# 11337 "parse.ml" +# 11368 "parse.ml" : 'tmNoEqWith_appTermNoRecordExp_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 2 : 'tmNoEqWith_appTermNoRecordExp_) in let _2 = (Parsing.peek_val __caml_parser_env 1 : string) in let _3 = (Parsing.peek_val __caml_parser_env 0 : 'tmNoEqWith_appTermNoRecordExp_) in Obj.repr( -# 2680 "parse.mly" +# 2690 "parse.mly" (let (e1, op, e2) = (_1, _2, _3) in ( mk_term (Op(mk_ident(op, rhs parseState 2), [e1; e2])) (rhs2 parseState 1 3) Un)) -# 11347 "parse.ml" +# 11378 "parse.ml" : 'tmNoEqWith_appTermNoRecordExp_)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 1 : 'recordExp) in Obj.repr( -# 2683 "parse.mly" +# 2693 "parse.mly" (let (_1, e, _3) = ((), _2, ()) in ( e )) -# 11355 "parse.ml" +# 11386 "parse.ml" : 'tmNoEqWith_appTermNoRecordExp_)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 0 : 'atomicTerm) in Obj.repr( -# 2686 "parse.mly" +# 2696 "parse.mly" (let (_1, e) = ((), _2) in ( mk_term (VQuote e) (rhs2 parseState 1 3) Un )) -# 11363 "parse.ml" +# 11394 "parse.ml" : 'tmNoEqWith_appTermNoRecordExp_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 1 : string) in let _2 = (Parsing.peek_val __caml_parser_env 0 : 'atomicTerm) in Obj.repr( -# 2689 "parse.mly" +# 2699 "parse.mly" (let (op, e) = (_1, _2) in ( mk_term (Op(mk_ident (op, rhs parseState 1), [e])) (rhs2 parseState 1 2) Formula )) -# 11372 "parse.ml" +# 11403 "parse.ml" : 'tmNoEqWith_appTermNoRecordExp_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'appTermNoRecordExp) in Obj.repr( -# 2692 "parse.mly" +# 2702 "parse.mly" (let e = _1 in ( e )) -# 11380 "parse.ml" +# 11411 "parse.ml" : 'tmNoEqWith_appTermNoRecordExp_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 2 : 'tmNoEqWith_tmRefinement_) in let _3 = (Parsing.peek_val __caml_parser_env 0 : 'tmNoEqWith_tmRefinement_) in Obj.repr( -# 2697 "parse.mly" +# 2707 "parse.mly" (let (e1, _2, e2) = (_1, (), _3) in ( consTerm (rhs parseState 2) e1 e2 )) -# 11389 "parse.ml" +# 11420 "parse.ml" : 'tmNoEqWith_tmRefinement_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 2 : 'tmNoEqWith_tmRefinement_) in let _3 = (Parsing.peek_val __caml_parser_env 0 : 'tmNoEqWith_tmRefinement_) in Obj.repr( -# 2700 "parse.mly" +# 2710 "parse.mly" (let (e1, _2, e2) = (_1, (), _3) in ( let dom = @@ -11424,237 +11455,237 @@ let op = ( mk_ident (op, rhs parseState 1) ) in in mk_term (Sum(dom, res)) (rhs2 parseState 1 3) Type_level )) -# 11414 "parse.ml" +# 11445 "parse.ml" : 'tmNoEqWith_tmRefinement_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 2 : 'tmNoEqWith_tmRefinement_) in let _2 = (Parsing.peek_val __caml_parser_env 1 : string) in let _3 = (Parsing.peek_val __caml_parser_env 0 : 'tmNoEqWith_tmRefinement_) in Obj.repr( -# 2719 "parse.mly" +# 2729 "parse.mly" (let (e1, op, e2) = (_1, _2, _3) in ( mk_term (Op(mk_ident(op, rhs parseState 2), [e1; e2])) (rhs2 parseState 1 3) Un)) -# 11424 "parse.ml" +# 11455 "parse.ml" : 'tmNoEqWith_tmRefinement_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 4 : 'tmNoEqWith_tmRefinement_) in let _3 = (Parsing.peek_val __caml_parser_env 2 : 'tmNoEqWith_tmRefinement_) in let _5 = (Parsing.peek_val __caml_parser_env 0 : 'tmNoEqWith_tmRefinement_) in Obj.repr( -# 2722 "parse.mly" +# 2732 "parse.mly" (let (e1, _2, op, _4, e2) = (_1, (), _3, (), _5) in ( mkApp op [ e1, Infix; e2, Nothing ] (rhs2 parseState 1 5) )) -# 11434 "parse.ml" +# 11465 "parse.ml" : 'tmNoEqWith_tmRefinement_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 2 : 'tmNoEqWith_tmRefinement_) in let _2 = (Parsing.peek_val __caml_parser_env 1 : string) in let _3 = (Parsing.peek_val __caml_parser_env 0 : 'tmNoEqWith_tmRefinement_) in Obj.repr( -# 2725 "parse.mly" +# 2735 "parse.mly" (let (e1, op, e2) = (_1, _2, _3) in ( mk_term (Op(mk_ident(op, rhs parseState 2), [e1; e2])) (rhs2 parseState 1 3) Un)) -# 11444 "parse.ml" +# 11475 "parse.ml" : 'tmNoEqWith_tmRefinement_)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 1 : 'recordExp) in Obj.repr( -# 2728 "parse.mly" +# 2738 "parse.mly" (let (_1, e, _3) = ((), _2, ()) in ( e )) -# 11452 "parse.ml" +# 11483 "parse.ml" : 'tmNoEqWith_tmRefinement_)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 0 : 'atomicTerm) in Obj.repr( -# 2731 "parse.mly" +# 2741 "parse.mly" (let (_1, e) = ((), _2) in ( mk_term (VQuote e) (rhs2 parseState 1 3) Un )) -# 11460 "parse.ml" +# 11491 "parse.ml" : 'tmNoEqWith_tmRefinement_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 1 : string) in let _2 = (Parsing.peek_val __caml_parser_env 0 : 'atomicTerm) in Obj.repr( -# 2734 "parse.mly" +# 2744 "parse.mly" (let (op, e) = (_1, _2) in ( mk_term (Op(mk_ident (op, rhs parseState 1), [e])) (rhs2 parseState 1 2) Formula )) -# 11469 "parse.ml" +# 11500 "parse.ml" : 'tmNoEqWith_tmRefinement_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'tmRefinement) in Obj.repr( -# 2737 "parse.mly" +# 2747 "parse.mly" (let e = _1 in ( e )) -# 11477 "parse.ml" +# 11508 "parse.ml" : 'tmNoEqWith_tmRefinement_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : string) in Obj.repr( -# 2742 "parse.mly" +# 2752 "parse.mly" (let o = _1 in ( mk_ident (o, rhs parseState 1) )) -# 11485 "parse.ml" +# 11516 "parse.ml" : 'binop_name)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : string) in Obj.repr( -# 2745 "parse.mly" +# 2755 "parse.mly" (let o = _1 in ( mk_ident (o, rhs parseState 1) )) -# 11493 "parse.ml" +# 11524 "parse.ml" : 'binop_name)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : string) in Obj.repr( -# 2748 "parse.mly" +# 2758 "parse.mly" (let o = _1 in ( mk_ident (o, rhs parseState 1) )) -# 11501 "parse.ml" +# 11532 "parse.ml" : 'binop_name)) ; (fun __caml_parser_env -> Obj.repr( -# 2751 "parse.mly" +# 2761 "parse.mly" (let o = () in ( mk_ident ("=", rhs parseState 1) )) -# 11508 "parse.ml" +# 11539 "parse.ml" : 'binop_name)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : string) in Obj.repr( -# 2754 "parse.mly" +# 2764 "parse.mly" (let o = _1 in ( mk_ident (o, rhs parseState 1) )) -# 11516 "parse.ml" +# 11547 "parse.ml" : 'binop_name)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : string) in Obj.repr( -# 2757 "parse.mly" +# 2767 "parse.mly" (let o = _1 in ( mk_ident (o, rhs parseState 1) )) -# 11524 "parse.ml" +# 11555 "parse.ml" : 'binop_name)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : string) in Obj.repr( -# 2760 "parse.mly" +# 2770 "parse.mly" (let o = _1 in ( mk_ident (o, rhs parseState 1) )) -# 11532 "parse.ml" +# 11563 "parse.ml" : 'binop_name)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : string) in Obj.repr( -# 2763 "parse.mly" +# 2773 "parse.mly" (let o = _1 in ( mk_ident (o, rhs parseState 1) )) -# 11540 "parse.ml" +# 11571 "parse.ml" : 'binop_name)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : string) in Obj.repr( -# 2766 "parse.mly" +# 2776 "parse.mly" (let o = _1 in ( mk_ident (o, rhs parseState 1) )) -# 11548 "parse.ml" +# 11579 "parse.ml" : 'binop_name)) ; (fun __caml_parser_env -> Obj.repr( -# 2769 "parse.mly" +# 2779 "parse.mly" (let o = () in ( mk_ident ("==>", rhs parseState 1) )) -# 11555 "parse.ml" +# 11586 "parse.ml" : 'binop_name)) ; (fun __caml_parser_env -> Obj.repr( -# 2772 "parse.mly" +# 2782 "parse.mly" (let o = () in ( mk_ident ("/\\", rhs parseState 1) )) -# 11562 "parse.ml" +# 11593 "parse.ml" : 'binop_name)) ; (fun __caml_parser_env -> Obj.repr( -# 2775 "parse.mly" +# 2785 "parse.mly" (let o = () in ( mk_ident ("\\/", rhs parseState 1) )) -# 11569 "parse.ml" +# 11600 "parse.ml" : 'binop_name)) ; (fun __caml_parser_env -> Obj.repr( -# 2778 "parse.mly" +# 2788 "parse.mly" (let o = () in ( mk_ident ("<==>", rhs parseState 1) )) -# 11576 "parse.ml" +# 11607 "parse.ml" : 'binop_name)) ; (fun __caml_parser_env -> Obj.repr( -# 2781 "parse.mly" +# 2791 "parse.mly" (let o = () in ( mk_ident ("|>", rhs parseState 1) )) -# 11583 "parse.ml" +# 11614 "parse.ml" : 'binop_name)) ; (fun __caml_parser_env -> Obj.repr( -# 2784 "parse.mly" +# 2794 "parse.mly" (let o = () in ( mk_ident (":=", rhs parseState 1) )) -# 11590 "parse.ml" +# 11621 "parse.ml" : 'binop_name)) ; (fun __caml_parser_env -> Obj.repr( -# 2787 "parse.mly" +# 2797 "parse.mly" (let o = () in ( mk_ident ("::", rhs parseState 1) )) -# 11597 "parse.ml" +# 11628 "parse.ml" : 'binop_name)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : string) in Obj.repr( -# 2790 "parse.mly" +# 2800 "parse.mly" (let o = _1 in ( mk_ident (o, rhs parseState 1) )) -# 11605 "parse.ml" +# 11636 "parse.ml" : 'binop_name)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : string) in Obj.repr( -# 2793 "parse.mly" +# 2803 "parse.mly" (let o = _1 in ( mk_ident (o, rhs parseState 1) )) -# 11613 "parse.ml" +# 11644 "parse.ml" : 'binop_name)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'tmEqWith_appTermNoRecordExp_) in Obj.repr( -# 2798 "parse.mly" +# 2808 "parse.mly" (let e = _1 in ( e )) -# 11621 "parse.ml" +# 11652 "parse.ml" : 'tmEqNoRefinement)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'tmEqWith_tmRefinement_) in Obj.repr( -# 2803 "parse.mly" +# 2813 "parse.mly" (let e = _1 in ( e )) -# 11629 "parse.ml" +# 11660 "parse.ml" : 'tmEq)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'tmNoEqWith_tmRefinement_) in Obj.repr( -# 2808 "parse.mly" +# 2818 "parse.mly" (let e = _1 in ( e )) -# 11637 "parse.ml" +# 11668 "parse.ml" : 'tmNoEq)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 3 : 'lidentOrUnderscore) in let _3 = (Parsing.peek_val __caml_parser_env 1 : 'appTermNoRecordExp) in let _4 = (Parsing.peek_val __caml_parser_env 0 : 'refineOpt) in Obj.repr( -# 2813 "parse.mly" +# 2823 "parse.mly" (let (id, _2, e, phi_opt) = (_1, (), _3, _4) in ( let t = match phi_opt with @@ -11662,326 +11693,326 @@ let op = ( mk_ident (op, rhs parseState 1) ) in | Some phi -> Refine(mk_binder (Annotated(id, e)) (rhs2 parseState 1 3) Type_level None, phi) in mk_term t (rhs2 parseState 1 4) Type_level )) -# 11652 "parse.ml" +# 11683 "parse.ml" : 'tmRefinement)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'appTerm) in Obj.repr( -# 2821 "parse.mly" +# 2831 "parse.mly" (let e = _1 in ( e )) -# 11660 "parse.ml" +# 11691 "parse.ml" : 'tmRefinement)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'option___anonymous_13_) in Obj.repr( -# 2826 "parse.mly" +# 2836 "parse.mly" (let phi_opt = _1 in (phi_opt)) -# 11668 "parse.ml" +# 11699 "parse.ml" : 'refineOpt)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'right_flexible_nonempty_list_SEMICOLON_simpleDef_) in Obj.repr( -# 2831 "parse.mly" +# 2841 "parse.mly" (let record_fields = _1 in ( mk_term (Record (None, record_fields)) (rhs parseState 1) Expr )) -# 11676 "parse.ml" +# 11707 "parse.ml" : 'recordExp)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 2 : 'appTerm) in let _3 = (Parsing.peek_val __caml_parser_env 0 : 'right_flexible_nonempty_list_SEMICOLON_simpleDef_) in Obj.repr( -# 2834 "parse.mly" +# 2844 "parse.mly" (let (e, _2, record_fields) = (_1, (), _3) in ( mk_term (Record (Some e, record_fields)) (rhs2 parseState 1 3) Expr )) -# 11685 "parse.ml" +# 11716 "parse.ml" : 'recordExp)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 2 : 'qlidentOrOperator) in let _3 = (Parsing.peek_val __caml_parser_env 0 : 'noSeqTerm) in Obj.repr( -# 2839 "parse.mly" +# 2849 "parse.mly" (let (x, _2, y) = (_1, (), _3) in let e = ( (x, y) ) in ( e )) -# 11695 "parse.ml" +# 11726 "parse.ml" : 'simpleDef)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'qlidentOrOperator) in Obj.repr( -# 2843 "parse.mly" +# 2853 "parse.mly" (let lid = _1 in ( lid, mk_term (Name (lid_of_ids [ ident_of_lid lid ])) (rhs parseState 1) Un )) -# 11703 "parse.ml" +# 11734 "parse.ml" : 'simpleDef)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 1 : 'indexingTerm) in let _2 = (Parsing.peek_val __caml_parser_env 0 : 'list___anonymous_14_) in Obj.repr( -# 2848 "parse.mly" +# 2858 "parse.mly" (let (head, args) = (_1, _2) in let t = ( mkApp head (map (fun (x,y) -> (y,x)) args) (rhs2 parseState 1 2) ) in (t)) -# 11713 "parse.ml" +# 11744 "parse.ml" : 'appTerm)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 1 : 'indexingTerm) in let _2 = (Parsing.peek_val __caml_parser_env 0 : 'list_argTerm_) in Obj.repr( -# 2854 "parse.mly" +# 2864 "parse.mly" (let (head, args) = (_1, _2) in let t = ( mkApp head (map (fun (x,y) -> (y,x)) args) (rhs2 parseState 1 2) ) in (t)) -# 11723 "parse.ml" +# 11754 "parse.ml" : 'appTermNoRecordExp)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'indexingTerm) in Obj.repr( -# 2860 "parse.mly" +# 2870 "parse.mly" (let y = _1 in let x = let x = ( Nothing ) in ( (x, y) ) in ( x )) -# 11735 "parse.ml" +# 11766 "parse.ml" : 'argTerm)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 0 : 'indexingTerm) in Obj.repr( -# 2867 "parse.mly" +# 2877 "parse.mly" (let (_1, y) = ((), _2) in let x = let x = ( Hash ) in ( (x, y) ) in ( x )) -# 11747 "parse.ml" +# 11778 "parse.ml" : 'argTerm)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'universe) in Obj.repr( -# 2874 "parse.mly" +# 2884 "parse.mly" (let u = _1 in ( u )) -# 11755 "parse.ml" +# 11786 "parse.ml" : 'argTerm)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 1 : 'atomicTermNotQUident) in let _2 = (Parsing.peek_val __caml_parser_env 0 : 'nonempty_list_dotOperator_) in Obj.repr( -# 2879 "parse.mly" +# 2889 "parse.mly" (let (e1, op_exprs) = (_1, _2) in ( List.fold_left (fun e1 (op, e2, r) -> mk_term (Op(op, [ e1; e2 ])) (union_ranges e1.range r) Expr) e1 op_exprs )) -# 11768 "parse.ml" +# 11799 "parse.ml" : 'indexingTerm)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'atomicTerm) in Obj.repr( -# 2886 "parse.mly" +# 2896 "parse.mly" (let e = _1 in ( e )) -# 11776 "parse.ml" +# 11807 "parse.ml" : 'indexingTerm)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'atomicTermNotQUident) in Obj.repr( -# 2891 "parse.mly" +# 2901 "parse.mly" (let x = _1 in ( x )) -# 11784 "parse.ml" +# 11815 "parse.ml" : 'atomicTerm)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'atomicTermQUident) in Obj.repr( -# 2894 "parse.mly" +# 2904 "parse.mly" (let x = _1 in ( x )) -# 11792 "parse.ml" +# 11823 "parse.ml" : 'atomicTerm)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'opPrefixTerm_atomicTermQUident_) in Obj.repr( -# 2897 "parse.mly" +# 2907 "parse.mly" (let x = _1 in ( x )) -# 11800 "parse.ml" +# 11831 "parse.ml" : 'atomicTerm)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'quident) in Obj.repr( -# 2902 "parse.mly" +# 2912 "parse.mly" (let id = _1 in ( let t = Name id in let e = mk_term t (rhs parseState 1) Un in e )) -# 11812 "parse.ml" +# 11843 "parse.ml" : 'atomicTermQUident)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 3 : 'quident) in let _3 = (Parsing.peek_val __caml_parser_env 1 : FStar_Parser_AST.term) in Obj.repr( -# 2909 "parse.mly" +# 2919 "parse.mly" (let (id, _2, t, _4) = (_1, (), _3, ()) in ( mk_term (LetOpen (id, t)) (rhs2 parseState 1 4) Expr )) -# 11823 "parse.ml" +# 11854 "parse.ml" : 'atomicTermQUident)) ; (fun __caml_parser_env -> Obj.repr( -# 2916 "parse.mly" +# 2926 "parse.mly" (let _1 = () in ( mk_term Wild (rhs parseState 1) Un )) -# 11830 "parse.ml" +# 11861 "parse.ml" : 'atomicTermNotQUident)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'tvar) in Obj.repr( -# 2919 "parse.mly" +# 2929 "parse.mly" (let tv = _1 in ( mk_term (Tvar tv) (rhs parseState 1) Type_level )) -# 11838 "parse.ml" +# 11869 "parse.ml" : 'atomicTermNotQUident)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'constant) in Obj.repr( -# 2922 "parse.mly" +# 2932 "parse.mly" (let c = _1 in ( mk_term (Const c) (rhs parseState 1) Expr )) -# 11846 "parse.ml" +# 11877 "parse.ml" : 'atomicTermNotQUident)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'opPrefixTerm_atomicTermNotQUident_) in Obj.repr( -# 2925 "parse.mly" +# 2935 "parse.mly" (let x = _1 in ( x )) -# 11854 "parse.ml" +# 11885 "parse.ml" : 'atomicTermNotQUident)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 1 : string) in Obj.repr( -# 2928 "parse.mly" +# 2938 "parse.mly" (let (_1, op, _3) = ((), _2, ()) in let op = ( mk_ident (op, rhs parseState 1) ) in ( mk_term (Op(op, [])) (rhs2 parseState 1 3) Un )) -# 11863 "parse.ml" +# 11894 "parse.ml" : 'atomicTermNotQUident)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 1 : 'binop_name) in Obj.repr( -# 2932 "parse.mly" +# 2942 "parse.mly" (let (_1, op, _3) = ((), _2, ()) in let op = ( op ) in ( mk_term (Op(op, [])) (rhs2 parseState 1 3) Un )) -# 11872 "parse.ml" +# 11903 "parse.ml" : 'atomicTermNotQUident)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 1 : string) in Obj.repr( -# 2936 "parse.mly" +# 2946 "parse.mly" (let (_1, op, _3) = ((), _2, ()) in let op = ( mk_ident (op, rhs parseState 1) ) in ( mk_term (Op(op, [])) (rhs2 parseState 1 3) Un )) -# 11881 "parse.ml" +# 11912 "parse.ml" : 'atomicTermNotQUident)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 1 : string) in Obj.repr( -# 2940 "parse.mly" +# 2950 "parse.mly" (let (_1, op, _3) = ((), _2, ()) in let op = let op = ( mk_ident ("and" ^ op, rhs parseState 1) ) in (op) in ( mk_term (Op(op, [])) (rhs2 parseState 1 3) Un )) -# 11893 "parse.ml" +# 11924 "parse.ml" : 'atomicTermNotQUident)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 1 : string) in Obj.repr( -# 2947 "parse.mly" +# 2957 "parse.mly" (let (_1, op, _3) = ((), _2, ()) in let op = let op = ( mk_ident ("let" ^ op, rhs parseState 1) ) in (op) in ( mk_term (Op(op, [])) (rhs2 parseState 1 3) Un )) -# 11905 "parse.ml" +# 11936 "parse.ml" : 'atomicTermNotQUident)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 3 : 'tmEq) in let _4 = (Parsing.peek_val __caml_parser_env 1 : 'separated_nonempty_list_COMMA_tmEq_) in Obj.repr( -# 2954 "parse.mly" +# 2964 "parse.mly" (let (_1, e0, _3, el, _5) = ((), _2, (), _4, ()) in ( mkDTuple (e0::el) (rhs2 parseState 1 5) )) -# 11914 "parse.ml" +# 11945 "parse.ml" : 'atomicTermNotQUident)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 1 : 'projectionLHS) in let _2 = (Parsing.peek_val __caml_parser_env 0 : 'list___anonymous_15_) in Obj.repr( -# 2957 "parse.mly" +# 2967 "parse.mly" (let (e, field_projs) = (_1, _2) in ( fold_left (fun e lid -> mk_term (Project(e, lid)) (rhs2 parseState 1 2) Expr ) e field_projs )) -# 11923 "parse.ml" +# 11954 "parse.ml" : 'atomicTermNotQUident)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 1 : FStar_Parser_AST.term) in Obj.repr( -# 2960 "parse.mly" +# 2970 "parse.mly" (let (_1, e, _3) = ((), _2, ()) in ( e )) -# 11931 "parse.ml" +# 11962 "parse.ml" : 'atomicTermNotQUident)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 1 : string) in let _2 = (Parsing.peek_val __caml_parser_env 0 : 'atomicTermNotQUident) in Obj.repr( -# 2965 "parse.mly" +# 2975 "parse.mly" (let (op, e) = (_1, _2) in ( mk_term (Op(mk_ident(op, rhs parseState 1), [e])) (rhs2 parseState 1 2) Expr )) -# 11940 "parse.ml" +# 11971 "parse.ml" : 'opPrefixTerm_atomicTermNotQUident_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 1 : string) in let _2 = (Parsing.peek_val __caml_parser_env 0 : 'atomicTermQUident) in Obj.repr( -# 2970 "parse.mly" +# 2980 "parse.mly" (let (op, e) = (_1, _2) in ( mk_term (Op(mk_ident(op, rhs parseState 1), [e])) (rhs2 parseState 1 2) Expr )) -# 11949 "parse.ml" +# 11980 "parse.ml" : 'opPrefixTerm_atomicTermQUident_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'qidentWithTypeArgs_qlident_option_fsTypeArgs__) in Obj.repr( -# 2975 "parse.mly" +# 2985 "parse.mly" (let e = _1 in ( e )) -# 11957 "parse.ml" +# 11988 "parse.ml" : 'projectionLHS)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'qidentWithTypeArgs_quident_some_fsTypeArgs__) in Obj.repr( -# 2978 "parse.mly" +# 2988 "parse.mly" (let e = _1 in ( e )) -# 11965 "parse.ml" +# 11996 "parse.ml" : 'projectionLHS)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 2 : FStar_Parser_AST.term) in let _3 = (Parsing.peek_val __caml_parser_env 1 : 'option_pair_hasSort_simpleTerm__) in Obj.repr( -# 2981 "parse.mly" +# 2991 "parse.mly" (let (_1, e, sort_opt, _4) = ((), _2, _3, ()) in ( (* Note: we have to keep the parentheses here. Consider t * u * v. This @@ -11996,12 +12027,12 @@ in | Some (level, t) -> mk_term (Ascribed(e,{t with level=level},None,false)) (rhs2 parseState 1 4) level in mk_term (Paren e1) (rhs2 parseState 1 4) (e.level) )) -# 11986 "parse.ml" +# 12017 "parse.ml" : 'projectionLHS)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 1 : 'right_flexible_list_SEMICOLON_noSeqTerm_) in Obj.repr( -# 2996 "parse.mly" +# 3006 "parse.mly" (let (_1, l, _3) = ((), _2, ()) in let es = ( l ) in ( @@ -12009,65 +12040,65 @@ let es = ( l ) in let pos = (rhs2 parseState 1 3) in mkExplicitApp (mk_term (Var (array_of_list_lid)) pos Expr) [l] pos )) -# 11999 "parse.ml" +# 12030 "parse.ml" : 'projectionLHS)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 1 : 'right_flexible_list_SEMICOLON_noSeqTerm_) in Obj.repr( -# 3004 "parse.mly" +# 3014 "parse.mly" (let (_1, l, _3) = ((), _2, ()) in let es = ( l ) in ( mkConsList (rhs2 parseState 1 3) es )) -# 12008 "parse.ml" +# 12039 "parse.ml" : 'projectionLHS)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 1 : 'right_flexible_list_SEMICOLON_noSeqTerm_) in Obj.repr( -# 3008 "parse.mly" +# 3018 "parse.mly" (let (_1, l, _3) = ((), _2, ()) in let es = ( l ) in ( mk_term (LexList es) (rhs2 parseState 1 3) Type_level )) -# 12017 "parse.ml" +# 12048 "parse.ml" : 'projectionLHS)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 1 : 'loption_separated_nonempty_list_COMMA_appTerm__) in Obj.repr( -# 3012 "parse.mly" +# 3022 "parse.mly" (let (_1, xs, _3) = ((), _2, ()) in let es = ( xs ) in ( mkRefSet (rhs2 parseState 1 3) es )) -# 12026 "parse.ml" +# 12057 "parse.ml" : 'projectionLHS)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 2 : 'quident) in let _3 = (Parsing.peek_val __caml_parser_env 0 : FStar_Ident.ident) in Obj.repr( -# 3016 "parse.mly" +# 3026 "parse.mly" (let (ns, _2, id) = (_1, (), _3) in ( mk_term (Projector (ns, id)) (rhs2 parseState 1 3) Expr )) -# 12035 "parse.ml" +# 12066 "parse.ml" : 'projectionLHS)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 1 : 'quident) in Obj.repr( -# 3019 "parse.mly" +# 3029 "parse.mly" (let (lid, _2) = (_1, ()) in ( mk_term (Discrim lid) (rhs2 parseState 1 2) Un )) -# 12043 "parse.ml" +# 12074 "parse.ml" : 'projectionLHS)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 1 : 'separated_nonempty_list_COMMA_atomicTerm_) in Obj.repr( -# 3024 "parse.mly" +# 3034 "parse.mly" (let (_1, targs, _3) = ((), _2, ()) in (targs)) -# 12051 "parse.ml" +# 12082 "parse.ml" : 'fsTypeArgs)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 1 : 'qlident) in let _2 = (Parsing.peek_val __caml_parser_env 0 : 'option_fsTypeArgs_) in Obj.repr( -# 3029 "parse.mly" +# 3039 "parse.mly" (let (id, targs_opt) = (_1, _2) in ( let t = if is_name id then Name id else Var id in @@ -12076,13 +12107,13 @@ let es = ( xs ) in | None -> e | Some targs -> mkFsTypApp e targs (rhs2 parseState 1 2) )) -# 12066 "parse.ml" +# 12097 "parse.ml" : 'qidentWithTypeArgs_qlident_option_fsTypeArgs__)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 1 : 'quident) in let _2 = (Parsing.peek_val __caml_parser_env 0 : 'some_fsTypeArgs_) in Obj.repr( -# 3040 "parse.mly" +# 3050 "parse.mly" (let (id, targs_opt) = (_1, _2) in ( let t = if is_name id then Name id else Var id in @@ -12091,203 +12122,203 @@ let es = ( xs ) in | None -> e | Some targs -> mkFsTypApp e targs (rhs2 parseState 1 2) )) -# 12081 "parse.ml" +# 12112 "parse.ml" : 'qidentWithTypeArgs_quident_some_fsTypeArgs__)) ; (fun __caml_parser_env -> Obj.repr( -# 3051 "parse.mly" +# 3061 "parse.mly" (let _1 = () in ( Type_level )) -# 12088 "parse.ml" +# 12119 "parse.ml" : 'hasSort)) ; (fun __caml_parser_env -> Obj.repr( -# 3056 "parse.mly" +# 3066 "parse.mly" (let _1 = () in ( Const_unit )) -# 12095 "parse.ml" +# 12126 "parse.ml" : 'constant)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : string * bool) in Obj.repr( -# 3059 "parse.mly" +# 3069 "parse.mly" (let n = _1 in ( if snd n then log_issue (lhs parseState) (Error_OutOfRange, "This number is outside the allowable range for representable integer constants"); Const_int (fst n, None) )) -# 12107 "parse.ml" +# 12138 "parse.ml" : 'constant)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : char) in Obj.repr( -# 3066 "parse.mly" +# 3076 "parse.mly" (let c = _1 in ( Const_char c )) -# 12115 "parse.ml" +# 12146 "parse.ml" : 'constant)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : string) in Obj.repr( -# 3069 "parse.mly" +# 3079 "parse.mly" (let s = _1 in ( Const_string (s,lhs(parseState)) )) -# 12123 "parse.ml" +# 12154 "parse.ml" : 'constant)) ; (fun __caml_parser_env -> Obj.repr( -# 3072 "parse.mly" +# 3082 "parse.mly" (let _1 = () in ( Const_bool true )) -# 12130 "parse.ml" +# 12161 "parse.ml" : 'constant)) ; (fun __caml_parser_env -> Obj.repr( -# 3075 "parse.mly" +# 3085 "parse.mly" (let _1 = () in ( Const_bool false )) -# 12137 "parse.ml" +# 12168 "parse.ml" : 'constant)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : string) in Obj.repr( -# 3078 "parse.mly" +# 3088 "parse.mly" (let r = _1 in ( Const_real r )) -# 12145 "parse.ml" +# 12176 "parse.ml" : 'constant)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : string) in Obj.repr( -# 3081 "parse.mly" +# 3091 "parse.mly" (let n = _1 in ( Const_int (n, Some (Unsigned, Int8)) )) -# 12153 "parse.ml" +# 12184 "parse.ml" : 'constant)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : string * bool) in Obj.repr( -# 3084 "parse.mly" +# 3094 "parse.mly" (let n = _1 in ( if snd n then log_issue (lhs(parseState)) (Error_OutOfRange, "This number is outside the allowable range for 8-bit signed integers"); Const_int (fst n, Some (Signed, Int8)) )) -# 12165 "parse.ml" +# 12196 "parse.ml" : 'constant)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : string) in Obj.repr( -# 3091 "parse.mly" +# 3101 "parse.mly" (let n = _1 in ( Const_int (n, Some (Unsigned, Int16)) )) -# 12173 "parse.ml" +# 12204 "parse.ml" : 'constant)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : string * bool) in Obj.repr( -# 3094 "parse.mly" +# 3104 "parse.mly" (let n = _1 in ( if snd n then log_issue (lhs(parseState)) (Error_OutOfRange, "This number is outside the allowable range for 16-bit signed integers"); Const_int (fst n, Some (Signed, Int16)) )) -# 12185 "parse.ml" +# 12216 "parse.ml" : 'constant)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : string) in Obj.repr( -# 3101 "parse.mly" +# 3111 "parse.mly" (let n = _1 in ( Const_int (n, Some (Unsigned, Int32)) )) -# 12193 "parse.ml" +# 12224 "parse.ml" : 'constant)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : string * bool) in Obj.repr( -# 3104 "parse.mly" +# 3114 "parse.mly" (let n = _1 in ( if snd n then log_issue (lhs(parseState)) (Error_OutOfRange, "This number is outside the allowable range for 32-bit signed integers"); Const_int (fst n, Some (Signed, Int32)) )) -# 12205 "parse.ml" +# 12236 "parse.ml" : 'constant)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : string) in Obj.repr( -# 3111 "parse.mly" +# 3121 "parse.mly" (let n = _1 in ( Const_int (n, Some (Unsigned, Int64)) )) -# 12213 "parse.ml" +# 12244 "parse.ml" : 'constant)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : string * bool) in Obj.repr( -# 3114 "parse.mly" +# 3124 "parse.mly" (let n = _1 in ( if snd n then log_issue (lhs(parseState)) (Error_OutOfRange, "This number is outside the allowable range for 64-bit signed integers"); Const_int (fst n, Some (Signed, Int64)) )) -# 12225 "parse.ml" +# 12256 "parse.ml" : 'constant)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : string) in Obj.repr( -# 3121 "parse.mly" +# 3131 "parse.mly" (let n = _1 in ( Const_int (n, Some (Unsigned, Sizet)) )) -# 12233 "parse.ml" +# 12264 "parse.ml" : 'constant)) ; (fun __caml_parser_env -> Obj.repr( -# 3124 "parse.mly" +# 3134 "parse.mly" (let _1 = () in ( Const_reify )) -# 12240 "parse.ml" +# 12271 "parse.ml" : 'constant)) ; (fun __caml_parser_env -> Obj.repr( -# 3127 "parse.mly" +# 3137 "parse.mly" (let _1 = () in ( Const_range_of )) -# 12247 "parse.ml" +# 12278 "parse.ml" : 'constant)) ; (fun __caml_parser_env -> Obj.repr( -# 3130 "parse.mly" +# 3140 "parse.mly" (let _1 = () in ( Const_set_range_of )) -# 12254 "parse.ml" +# 12285 "parse.ml" : 'constant)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 0 : 'atomicUniverse) in Obj.repr( -# 3135 "parse.mly" +# 3145 "parse.mly" (let (_1, ua) = ((), _2) in ( (UnivApp, ua) )) -# 12262 "parse.ml" +# 12293 "parse.ml" : 'universe)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'atomicUniverse) in Obj.repr( -# 3140 "parse.mly" +# 3150 "parse.mly" (let ua = _1 in ( ua )) -# 12270 "parse.ml" +# 12301 "parse.ml" : 'universeFrom)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 2 : 'universeFrom) in let _2 = (Parsing.peek_val __caml_parser_env 1 : string) in let _3 = (Parsing.peek_val __caml_parser_env 0 : 'universeFrom) in Obj.repr( -# 3143 "parse.mly" +# 3153 "parse.mly" (let (u1, op_plus, u2) = (_1, _2, _3) in ( if op_plus <> "+" @@ -12295,13 +12326,13 @@ let es = ( xs ) in ^ "The only allowed operator in that context is +.")); mk_term (Op(mk_ident (op_plus, rhs parseState 2), [u1 ; u2])) (rhs2 parseState 1 3) Expr )) -# 12285 "parse.ml" +# 12316 "parse.ml" : 'universeFrom)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 1 : 'ident) in let _2 = (Parsing.peek_val __caml_parser_env 0 : 'nonempty_list_atomicUniverse_) in Obj.repr( -# 3151 "parse.mly" +# 3161 "parse.mly" (let (max, us) = (_1, _2) in ( if string_of_id max <> string_of_lid max_lid @@ -12311,325 +12342,327 @@ let es = ( xs ) in let max = mk_term (Var (lid_of_ids [max])) (rhs parseState 1) Expr in mkApp max (map (fun u -> u, Nothing) us) (rhs2 parseState 1 2) )) -# 12301 "parse.ml" +# 12332 "parse.ml" : 'universeFrom)) ; (fun __caml_parser_env -> Obj.repr( -# 3163 "parse.mly" +# 3173 "parse.mly" (let _1 = () in ( mk_term Wild (rhs parseState 1) Expr )) -# 12308 "parse.ml" +# 12339 "parse.ml" : 'atomicUniverse)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : string * bool) in Obj.repr( -# 3166 "parse.mly" +# 3176 "parse.mly" (let n = _1 in ( if snd n then log_issue (lhs(parseState)) (Error_OutOfRange, "This number is outside the allowable range for representable integer constants"); mk_term (Const (Const_int (fst n, None))) (rhs parseState 1) Expr )) -# 12320 "parse.ml" +# 12351 "parse.ml" : 'atomicUniverse)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : FStar_Ident.ident) in Obj.repr( -# 3173 "parse.mly" +# 3183 "parse.mly" (let u = _1 in ( mk_term (Uvar u) (range_of_id u) Expr )) -# 12328 "parse.ml" +# 12359 "parse.ml" : 'atomicUniverse)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 1 : 'universeFrom) in Obj.repr( -# 3176 "parse.mly" +# 3186 "parse.mly" (let (_1, u, _3) = ((), _2, ()) in ( u (*mk_term (Paren u) (rhs2 parseState 1 3) Expr*) )) -# 12336 "parse.ml" +# 12367 "parse.ml" : 'atomicUniverse)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 1 : 'warn_error) in Obj.repr( -# 3181 "parse.mly" +# 3191 "parse.mly" (let (e, _2) = (_1, ()) in ( e )) -# 12344 "parse.ml" +# 12375 "parse.ml" : (FStar_Errors.flag * string) list)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 1 : 'flag) in let _2 = (Parsing.peek_val __caml_parser_env 0 : 'range) in Obj.repr( -# 3186 "parse.mly" +# 3196 "parse.mly" (let (f, r) = (_1, _2) in ( [(f, r)] )) -# 12353 "parse.ml" +# 12384 "parse.ml" : 'warn_error)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 2 : 'flag) in let _2 = (Parsing.peek_val __caml_parser_env 1 : 'range) in let _3 = (Parsing.peek_val __caml_parser_env 0 : 'warn_error) in Obj.repr( -# 3189 "parse.mly" +# 3199 "parse.mly" (let (f, r, e) = (_1, _2, _3) in ( (f, r) :: e )) -# 12363 "parse.ml" +# 12394 "parse.ml" : 'warn_error)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : string) in Obj.repr( -# 3194 "parse.mly" +# 3204 "parse.mly" (let op = _1 in ( if op = "@" then CAlwaysError else failwith (format1 "unexpected token %s in warn-error list" op))) -# 12371 "parse.ml" +# 12402 "parse.ml" : 'flag)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : string) in Obj.repr( -# 3197 "parse.mly" +# 3207 "parse.mly" (let op = _1 in ( if op = "+" then CWarning else failwith (format1 "unexpected token %s in warn-error list" op))) -# 12379 "parse.ml" +# 12410 "parse.ml" : 'flag)) ; (fun __caml_parser_env -> Obj.repr( -# 3200 "parse.mly" +# 3210 "parse.mly" (let _1 = () in ( CSilent )) -# 12386 "parse.ml" +# 12417 "parse.ml" : 'flag)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : string * bool) in Obj.repr( -# 3205 "parse.mly" +# 3215 "parse.mly" (let i = _1 in ( format2 "%s..%s" (fst i) (fst i) )) -# 12394 "parse.ml" +# 12425 "parse.ml" : 'range)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : string) in Obj.repr( -# 3208 "parse.mly" +# 3218 "parse.mly" (let r = _1 in ( r )) -# 12402 "parse.ml" +# 12433 "parse.ml" : 'range)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : string) in Obj.repr( -# 3213 "parse.mly" +# 3223 "parse.mly" (let s = _1 in ( s )) -# 12410 "parse.ml" +# 12441 "parse.ml" : 'string)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'fsTypeArgs) in Obj.repr( -# 3218 "parse.mly" +# 3228 "parse.mly" (let x = _1 in ( Some x )) -# 12418 "parse.ml" +# 12449 "parse.ml" : 'some_fsTypeArgs_)) ; (fun __caml_parser_env -> Obj.repr( -# 3223 "parse.mly" +# 3233 "parse.mly" ( ( [] )) -# 12424 "parse.ml" +# 12455 "parse.ml" : 'right_flexible_list_SEMICOLON_fieldPattern_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'fieldPattern) in Obj.repr( -# 3225 "parse.mly" +# 3235 "parse.mly" (let x = _1 in ( [x] )) -# 12432 "parse.ml" +# 12463 "parse.ml" : 'right_flexible_list_SEMICOLON_fieldPattern_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 2 : 'fieldPattern) in let _3 = (Parsing.peek_val __caml_parser_env 0 : 'right_flexible_list_SEMICOLON_fieldPattern_) in Obj.repr( -# 3228 "parse.mly" +# 3238 "parse.mly" (let (x, _2, xs) = (_1, (), _3) in ( x :: xs )) -# 12441 "parse.ml" +# 12472 "parse.ml" : 'right_flexible_list_SEMICOLON_fieldPattern_)) ; (fun __caml_parser_env -> Obj.repr( -# 3233 "parse.mly" +# 3243 "parse.mly" ( ( [] )) -# 12447 "parse.ml" +# 12478 "parse.ml" : 'right_flexible_list_SEMICOLON_noSeqTerm_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'noSeqTerm) in Obj.repr( -# 3235 "parse.mly" +# 3245 "parse.mly" (let x = _1 in ( [x] )) -# 12455 "parse.ml" +# 12486 "parse.ml" : 'right_flexible_list_SEMICOLON_noSeqTerm_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 2 : 'noSeqTerm) in let _3 = (Parsing.peek_val __caml_parser_env 0 : 'right_flexible_list_SEMICOLON_noSeqTerm_) in Obj.repr( -# 3238 "parse.mly" +# 3248 "parse.mly" (let (x, _2, xs) = (_1, (), _3) in ( x :: xs )) -# 12464 "parse.ml" +# 12495 "parse.ml" : 'right_flexible_list_SEMICOLON_noSeqTerm_)) ; (fun __caml_parser_env -> Obj.repr( -# 3243 "parse.mly" +# 3253 "parse.mly" ( ( [] )) -# 12470 "parse.ml" +# 12501 "parse.ml" : 'right_flexible_list_SEMICOLON_recordFieldDecl_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'recordFieldDecl) in Obj.repr( -# 3245 "parse.mly" +# 3255 "parse.mly" (let x = _1 in ( [x] )) -# 12478 "parse.ml" +# 12509 "parse.ml" : 'right_flexible_list_SEMICOLON_recordFieldDecl_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 2 : 'recordFieldDecl) in let _3 = (Parsing.peek_val __caml_parser_env 0 : 'right_flexible_list_SEMICOLON_recordFieldDecl_) in Obj.repr( -# 3248 "parse.mly" +# 3258 "parse.mly" (let (x, _2, xs) = (_1, (), _3) in ( x :: xs )) -# 12487 "parse.ml" +# 12518 "parse.ml" : 'right_flexible_list_SEMICOLON_recordFieldDecl_)) ; (fun __caml_parser_env -> Obj.repr( -# 3253 "parse.mly" +# 3263 "parse.mly" ( ( [] )) -# 12493 "parse.ml" +# 12524 "parse.ml" : 'right_flexible_list_SEMICOLON_simpleDef_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'simpleDef) in Obj.repr( -# 3255 "parse.mly" +# 3265 "parse.mly" (let x = _1 in ( [x] )) -# 12501 "parse.ml" +# 12532 "parse.ml" : 'right_flexible_list_SEMICOLON_simpleDef_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 2 : 'simpleDef) in let _3 = (Parsing.peek_val __caml_parser_env 0 : 'right_flexible_list_SEMICOLON_simpleDef_) in Obj.repr( -# 3258 "parse.mly" +# 3268 "parse.mly" (let (x, _2, xs) = (_1, (), _3) in ( x :: xs )) -# 12510 "parse.ml" +# 12541 "parse.ml" : 'right_flexible_list_SEMICOLON_simpleDef_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'fieldPattern) in Obj.repr( -# 3263 "parse.mly" +# 3273 "parse.mly" (let x = _1 in ( [x] )) -# 12518 "parse.ml" +# 12549 "parse.ml" : 'right_flexible_nonempty_list_SEMICOLON_fieldPattern_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 2 : 'fieldPattern) in let _3 = (Parsing.peek_val __caml_parser_env 0 : 'right_flexible_list_SEMICOLON_fieldPattern_) in Obj.repr( -# 3266 "parse.mly" +# 3276 "parse.mly" (let (x, _2, xs) = (_1, (), _3) in ( x :: xs )) -# 12527 "parse.ml" +# 12558 "parse.ml" : 'right_flexible_nonempty_list_SEMICOLON_fieldPattern_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'recordFieldDecl) in Obj.repr( -# 3271 "parse.mly" +# 3281 "parse.mly" (let x = _1 in ( [x] )) -# 12535 "parse.ml" +# 12566 "parse.ml" : 'right_flexible_nonempty_list_SEMICOLON_recordFieldDecl_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 2 : 'recordFieldDecl) in let _3 = (Parsing.peek_val __caml_parser_env 0 : 'right_flexible_list_SEMICOLON_recordFieldDecl_) in Obj.repr( -# 3274 "parse.mly" +# 3284 "parse.mly" (let (x, _2, xs) = (_1, (), _3) in ( x :: xs )) -# 12544 "parse.ml" +# 12575 "parse.ml" : 'right_flexible_nonempty_list_SEMICOLON_recordFieldDecl_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'simpleDef) in Obj.repr( -# 3279 "parse.mly" +# 3289 "parse.mly" (let x = _1 in ( [x] )) -# 12552 "parse.ml" +# 12583 "parse.ml" : 'right_flexible_nonempty_list_SEMICOLON_simpleDef_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 2 : 'simpleDef) in let _3 = (Parsing.peek_val __caml_parser_env 0 : 'right_flexible_list_SEMICOLON_simpleDef_) in Obj.repr( -# 3282 "parse.mly" +# 3292 "parse.mly" (let (x, _2, xs) = (_1, (), _3) in ( x :: xs )) -# 12561 "parse.ml" +# 12592 "parse.ml" : 'right_flexible_nonempty_list_SEMICOLON_simpleDef_)) ; (fun __caml_parser_env -> Obj.repr( -# 3287 "parse.mly" +# 3297 "parse.mly" ( ( [] )) -# 12567 "parse.ml" +# 12598 "parse.ml" : 'reverse_left_flexible_list_BAR___anonymous_10_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'patternBranch) in Obj.repr( -# 3289 "parse.mly" +# 3299 "parse.mly" (let pb = _1 in let x = (pb) in ( [x] )) -# 12576 "parse.ml" +# 12607 "parse.ml" : 'reverse_left_flexible_list_BAR___anonymous_10_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 2 : 'reverse_left_flexible_list_BAR___anonymous_10_) in let _3 = (Parsing.peek_val __caml_parser_env 0 : 'patternBranch) in Obj.repr( -# 3293 "parse.mly" +# 3303 "parse.mly" (let (xs, _2, pb) = (_1, (), _3) in let x = (pb) in ( x :: xs )) -# 12586 "parse.ml" +# 12617 "parse.ml" : 'reverse_left_flexible_list_BAR___anonymous_10_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 0 : 'patternBranch) in Obj.repr( -# 3299 "parse.mly" +# 3309 "parse.mly" (let x = _1 in let _1 = ( None ) in ( [x] )) -# 12595 "parse.ml" +# 12626 "parse.ml" : 'reverse_left_flexible_nonempty_list_BAR_patternBranch_)) ; (fun __caml_parser_env -> let _2 = (Parsing.peek_val __caml_parser_env 0 : 'patternBranch) in Obj.repr( -# 3303 "parse.mly" +# 3313 "parse.mly" (let (x_inlined1, x) = ((), _2) in let _1 = let x = x_inlined1 in ( Some x ) in ( [x] )) -# 12607 "parse.ml" +# 12638 "parse.ml" : 'reverse_left_flexible_nonempty_list_BAR_patternBranch_)) ; (fun __caml_parser_env -> let _1 = (Parsing.peek_val __caml_parser_env 2 : 'reverse_left_flexible_nonempty_list_BAR_patternBranch_) in let _3 = (Parsing.peek_val __caml_parser_env 0 : 'patternBranch) in Obj.repr( -# 3310 "parse.mly" +# 3320 "parse.mly" (let (xs, _2, x) = (_1, (), _3) in ( x :: xs )) -# 12616 "parse.ml" +# 12647 "parse.ml" : 'reverse_left_flexible_nonempty_list_BAR_patternBranch_)) (* Entry inputFragment *) ; (fun __caml_parser_env -> raise (Parsing.YYexit (Parsing.peek_val __caml_parser_env 0))) +(* Entry oneDeclOrEOF *) +; (fun __caml_parser_env -> raise (Parsing.YYexit (Parsing.peek_val __caml_parser_env 0))) (* Entry term *) ; (fun __caml_parser_env -> raise (Parsing.YYexit (Parsing.peek_val __caml_parser_env 0))) (* Entry warn_error_list *) @@ -12654,8 +12687,10 @@ let yytables = Parsing.names_block=yynames_block } let inputFragment (lexfun : Lexing.lexbuf -> token) (lexbuf : Lexing.lexbuf) = (Parsing.yyparse yytables 1 lexfun lexbuf : FStar_Parser_AST.inputFragment) +let oneDeclOrEOF (lexfun : Lexing.lexbuf -> token) (lexbuf : Lexing.lexbuf) = + (Parsing.yyparse yytables 2 lexfun lexbuf : FStar_Parser_AST.decl option) let term (lexfun : Lexing.lexbuf -> token) (lexbuf : Lexing.lexbuf) = - (Parsing.yyparse yytables 2 lexfun lexbuf : FStar_Parser_AST.term) + (Parsing.yyparse yytables 3 lexfun lexbuf : FStar_Parser_AST.term) let warn_error_list (lexfun : Lexing.lexbuf -> token) (lexbuf : Lexing.lexbuf) = - (Parsing.yyparse yytables 3 lexfun lexbuf : (FStar_Errors.flag * string) list) + (Parsing.yyparse yytables 4 lexfun lexbuf : (FStar_Errors.flag * string) list) ;; diff --git a/ocaml/fstar-lib/generated/FStar_Syntax_Resugar.ml b/ocaml/fstar-lib/generated/FStar_Syntax_Resugar.ml index 7d12de91093..834731c2b82 100644 --- a/ocaml/fstar-lib/generated/FStar_Syntax_Resugar.ml +++ b/ocaml/fstar-lib/generated/FStar_Syntax_Resugar.ml @@ -2792,7 +2792,8 @@ let (resugar_sigelt' : { FStar_Parser_AST.msource = src; FStar_Parser_AST.mdest = dst; - FStar_Parser_AST.lift_op = op + FStar_Parser_AST.lift_op = op; + FStar_Parser_AST.braced = false }) in FStar_Pervasives_Native.Some uu___ | FStar_Syntax_Syntax.Sig_effect_abbrev (lid, vs, bs, c, flags) -> diff --git a/ocaml/fstar-lib/generated/FStar_Universal.ml b/ocaml/fstar-lib/generated/FStar_Universal.ml index a6e8251849d..e1a8a4a7d6b 100644 --- a/ocaml/fstar-lib/generated/FStar_Universal.ml +++ b/ocaml/fstar-lib/generated/FStar_Universal.ml @@ -708,7 +708,8 @@ let (init_env : FStar_Parser_Dep.deps -> FStar_TypeChecker_Env.env) = let (tc_one_fragment : FStar_Syntax_Syntax.modul FStar_Pervasives_Native.option -> FStar_TypeChecker_Env.env_t -> - FStar_Parser_ParseIt.input_frag -> + (FStar_Parser_ParseIt.input_frag, FStar_Parser_AST.decl) + FStar_Pervasives.either -> (FStar_Syntax_Syntax.modul FStar_Pervasives_Native.option * FStar_TypeChecker_Env.env)) = @@ -749,100 +750,110 @@ let (tc_one_fragment : uu___5) -> d | uu___ -> FStar_Compiler_Range.dummyRange in - let uu___ = FStar_Parser_Driver.parse_fragment frag in - match uu___ with - | FStar_Parser_Driver.Empty -> (curmod, env) - | FStar_Parser_Driver.Decls [] -> (curmod, env) - | FStar_Parser_Driver.Modul ast_modul -> + let check_module_name_declaration ast_modul = + let uu___ = let uu___1 = - let uu___2 = - FStar_ToSyntax_Interleave.interleave_module ast_modul false in - FStar_Compiler_Effect.op_Less_Bar (with_dsenv_of_tcenv env) - uu___2 in - (match uu___1 with - | (ast_modul1, env1) -> - let uu___2 = - let uu___3 = - FStar_ToSyntax_ToSyntax.partial_ast_modul_to_modul - curmod ast_modul1 in - FStar_Compiler_Effect.op_Less_Bar - (with_dsenv_of_tcenv env1) uu___3 in - (match uu___2 with - | (modul, env2) -> - ((let uu___4 = - let uu___5 = acceptable_mod_name modul in - Prims.op_Negation uu___5 in - if uu___4 - then - let msg = - let uu___5 = - let uu___6 = fname env2 in - FStar_Parser_Dep.module_name_of_file uu___6 in - FStar_Compiler_Util.format1 - "Interactive mode only supports a single module at the top-level. Expected module %s" - uu___5 in - FStar_Errors.raise_error - (FStar_Errors.Fatal_NonSingletonTopLevelModule, - msg) (range_of_first_mod_decl ast_modul1) - else ()); - (let uu___4 = - let uu___5 = + FStar_ToSyntax_Interleave.interleave_module ast_modul false in + FStar_Compiler_Effect.op_Less_Bar (with_dsenv_of_tcenv env) + uu___1 in + match uu___ with + | (ast_modul1, env1) -> + let uu___1 = + let uu___2 = + FStar_ToSyntax_ToSyntax.partial_ast_modul_to_modul curmod + ast_modul1 in + FStar_Compiler_Effect.op_Less_Bar (with_dsenv_of_tcenv env1) + uu___2 in + (match uu___1 with + | (modul, env2) -> + ((let uu___3 = + let uu___4 = acceptable_mod_name modul in + Prims.op_Negation uu___4 in + if uu___3 + then + let msg = + let uu___4 = + let uu___5 = fname env2 in + FStar_Parser_Dep.module_name_of_file uu___5 in + FStar_Compiler_Util.format1 + "Interactive mode only supports a single module at the top-level. Expected module %s" + uu___4 in + FStar_Errors.raise_error + (FStar_Errors.Fatal_NonSingletonTopLevelModule, msg) + (range_of_first_mod_decl ast_modul1) + else ()); + (let uu___3 = + let uu___4 = + FStar_Syntax_DsEnv.syntax_only + env2.FStar_TypeChecker_Env.dsenv in + if uu___4 + then (modul, env2) + else FStar_TypeChecker_Tc.tc_partial_modul env2 modul in + match uu___3 with + | (modul1, env3) -> + ((FStar_Pervasives_Native.Some modul1), env3)))) in + let check_decls ast_decls = + match curmod with + | FStar_Pervasives_Native.None -> + let uu___ = FStar_Compiler_List.hd ast_decls in + (match uu___ with + | { FStar_Parser_AST.d = uu___1; + FStar_Parser_AST.drange = rng; + FStar_Parser_AST.quals = uu___2; + FStar_Parser_AST.attrs = uu___3;_} -> + FStar_Errors.raise_error + (FStar_Errors.Fatal_ModuleFirstStatement, + "First statement must be a module declaration") rng) + | FStar_Pervasives_Native.Some modul -> + let uu___ = + FStar_Compiler_Util.fold_map + (fun env1 -> + fun a_decl -> + let uu___1 = + let uu___2 = + FStar_ToSyntax_Interleave.prefix_with_interface_decls + modul.FStar_Syntax_Syntax.name a_decl in + FStar_Compiler_Effect.op_Less_Bar + (with_dsenv_of_tcenv env1) uu___2 in + match uu___1 with | (decls, env2) -> (env2, decls)) + env ast_decls in + (match uu___ with + | (env1, ast_decls_l) -> + let uu___1 = + let uu___2 = + FStar_ToSyntax_ToSyntax.decls_to_sigelts + (FStar_Compiler_List.flatten ast_decls_l) in + FStar_Compiler_Effect.op_Less_Bar + (with_dsenv_of_tcenv env1) uu___2 in + (match uu___1 with + | (sigelts, env2) -> + let uu___2 = + let uu___3 = FStar_Syntax_DsEnv.syntax_only env2.FStar_TypeChecker_Env.dsenv in - if uu___5 - then (modul, env2) + if uu___3 + then (modul, [], env2) else - FStar_TypeChecker_Tc.tc_partial_modul env2 modul in - match uu___4 with - | (modul1, env3) -> - ((FStar_Pervasives_Native.Some modul1), env3))))) - | FStar_Parser_Driver.Decls ast_decls -> - (match curmod with - | FStar_Pervasives_Native.None -> - let uu___1 = FStar_Compiler_List.hd ast_decls in - (match uu___1 with - | { FStar_Parser_AST.d = uu___2; - FStar_Parser_AST.drange = rng; - FStar_Parser_AST.quals = uu___3; - FStar_Parser_AST.attrs = uu___4;_} -> - FStar_Errors.raise_error - (FStar_Errors.Fatal_ModuleFirstStatement, - "First statement must be a module declaration") rng) - | FStar_Pervasives_Native.Some modul -> - let uu___1 = - FStar_Compiler_Util.fold_map - (fun env1 -> - fun a_decl -> - let uu___2 = - let uu___3 = - FStar_ToSyntax_Interleave.prefix_with_interface_decls - modul.FStar_Syntax_Syntax.name a_decl in - FStar_Compiler_Effect.op_Less_Bar - (with_dsenv_of_tcenv env1) uu___3 in - match uu___2 with | (decls, env2) -> (env2, decls)) - env ast_decls in - (match uu___1 with - | (env1, ast_decls_l) -> - let uu___2 = - let uu___3 = - FStar_ToSyntax_ToSyntax.decls_to_sigelts - (FStar_Compiler_List.flatten ast_decls_l) in - FStar_Compiler_Effect.op_Less_Bar - (with_dsenv_of_tcenv env1) uu___3 in - (match uu___2 with - | (sigelts, env2) -> - let uu___3 = - let uu___4 = - FStar_Syntax_DsEnv.syntax_only - env2.FStar_TypeChecker_Env.dsenv in - if uu___4 - then (modul, [], env2) - else - FStar_TypeChecker_Tc.tc_more_partial_modul - env2 modul sigelts in - (match uu___3 with - | (modul1, uu___4, env3) -> - ((FStar_Pervasives_Native.Some modul1), env3))))) + FStar_TypeChecker_Tc.tc_more_partial_modul env2 + modul sigelts in + (match uu___2 with + | (modul1, uu___3, env3) -> + ((FStar_Pervasives_Native.Some modul1), env3)))) in + match frag with + | FStar_Pervasives.Inr d -> + (match d.FStar_Parser_AST.d with + | FStar_Parser_AST.TopLevelModule lid -> + check_module_name_declaration + (FStar_Parser_AST.Module (lid, [d])) + | uu___ -> check_decls [d]) + | FStar_Pervasives.Inl frag1 -> + let uu___ = FStar_Parser_Driver.parse_fragment frag1 in + (match uu___ with + | FStar_Parser_Driver.Empty -> (curmod, env) + | FStar_Parser_Driver.Decls [] -> (curmod, env) + | FStar_Parser_Driver.Modul ast_modul -> + check_module_name_declaration ast_modul + | FStar_Parser_Driver.Decls ast_decls -> check_decls ast_decls) let (load_interface_decls : FStar_TypeChecker_Env.env -> Prims.string -> FStar_TypeChecker_Env.env_t) = fun env -> diff --git a/ocaml/fstar-tests/generated/FStar_Tests_Pars.ml b/ocaml/fstar-tests/generated/FStar_Tests_Pars.ml index 334b8a7b2a8..3c1b7969b0a 100644 --- a/ocaml/fstar-tests/generated/FStar_Tests_Pars.ml +++ b/ocaml/fstar-tests/generated/FStar_Tests_Pars.ml @@ -565,7 +565,7 @@ let (pars_and_tc_fragment : Prims.string -> unit) = let uu___4 = FStar_Compiler_Effect.op_Bang test_mod_ref in FStar_Universal.tc_one_fragment uu___4 tcenv - (frag, FStar_Pervasives_Native.None) in + (FStar_Pervasives.Inl frag) in (match uu___3 with | (test_mod', tcenv') -> (FStar_Compiler_Effect.op_Colon_Equals diff --git a/src/Makefile.boot b/src/Makefile.boot index ed0ab280af4..a3d12a1486f 100644 --- a/src/Makefile.boot +++ b/src/Makefile.boot @@ -40,7 +40,9 @@ EXTRACT_MODULES=FStar.Pervasives FStar.Common FStar.Compiler.Range FStar.Thunk FStar.Interactive.CompletionTable \ FStar.Interactive.JsonHelper FStar.Interactive.QueryHelper \ FStar.Interactive.PushHelper FStar.Interactive.Lsp \ - FStar.Interactive.Ide FStar.Interactive.Legacy \ + FStar.Interactive.Ide FStar.Interactive.Ide.Types \ + FStar.Interactive.ReplState \ + FStar.Interactive.Incremental FStar.Interactive.Legacy \ FStar.CheckedFiles FStar.Universal FStar.Prettyprint \ FStar.Main FStar.Compiler.List FStar.Compiler.Option \ FStar.Compiler.Dyn diff --git a/src/fstar/FStar.Interactive.Ide.Types.fst b/src/fstar/FStar.Interactive.Ide.Types.fst new file mode 100755 index 00000000000..7ad07536aeb --- /dev/null +++ b/src/fstar/FStar.Interactive.Ide.Types.fst @@ -0,0 +1,217 @@ +(* + Copyright 2008-2016 Nikhil Swamy and Microsoft Research + + Licensed under the Apache License, Version 2.0 (the "License"); + you may not use this file except in compliance with the License. + You may obtain a copy of the License at + + http://www.apache.org/licenses/LICENSE-2.0 + + Unless required by applicable law or agreed to in writing, software + distributed under the License is distributed on an "AS IS" BASIS, + WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or implied. + See the License for the specific language governing permissions and + limitations under the License. +*) + +module FStar.Interactive.Ide.Types +open FStar.Pervasives +open FStar.Compiler.Effect +open FStar.Compiler.List +open FStar +open FStar.Compiler +open FStar.Compiler.Range +open FStar.Compiler.Util +open FStar.Getopt +open FStar.Ident +open FStar.Errors +open FStar.Interactive.JsonHelper +open FStar.Interactive.QueryHelper +open FStar.Interactive.PushHelper + +open FStar.Universal +open FStar.TypeChecker.Env +open FStar.TypeChecker.Common +open FStar.Interactive +open FStar.Parser.ParseIt + +module SS = FStar.Syntax.Syntax +module DsEnv = FStar.Syntax.DsEnv +module TcErr = FStar.TypeChecker.Err +module TcEnv = FStar.TypeChecker.Env +module CTable = FStar.Interactive.CompletionTable +module QH = FStar.Interactive.QueryHelper + +let initial_range = + Range.mk_range "" (Range.mk_pos 1 0) (Range.mk_pos 1 0) + + +(*****************************************) +(* Reading queries and writing responses *) +(*****************************************) + +let js_pushkind s : push_kind = match js_str s with + | "syntax" -> SyntaxCheck + | "lax" -> LaxCheck + | "full" -> FullCheck + | _ -> js_fail "push_kind" s + +let js_reductionrule s = match js_str s with + | "beta" -> FStar.TypeChecker.Env.Beta + | "delta" -> FStar.TypeChecker.Env.UnfoldUntil SS.delta_constant + | "iota" -> FStar.TypeChecker.Env.Iota + | "zeta" -> FStar.TypeChecker.Env.Zeta + | "reify" -> FStar.TypeChecker.Env.Reify + | "pure-subterms" -> FStar.TypeChecker.Env.PureSubtermsWithinComputations + | _ -> js_fail "reduction rule" s + +type completion_context = +| CKCode +| CKOption of bool (* #set-options (false) or #reset-options (true) *) +| CKModuleOrNamespace of bool (* modules *) * bool (* namespaces *) + +let js_optional_completion_context k = + match k with + | None -> CKCode + | Some k -> + match js_str k with + | "symbol" // Backwards compatibility + | "code" -> CKCode + | "set-options" -> CKOption false + | "reset-options" -> CKOption true + | "open" + | "let-open" -> CKModuleOrNamespace (true, true) + | "include" + | "module-alias" -> CKModuleOrNamespace (true, false) + | _ -> + js_fail "completion context (code, set-options, reset-options, \ +open, let-open, include, module-alias)" k + +type lookup_context = +| LKSymbolOnly +| LKModule +| LKOption +| LKCode + +let js_optional_lookup_context k = + match k with + | None -> LKSymbolOnly // Backwards-compatible default + | Some k -> + match js_str k with + | "symbol-only" -> LKSymbolOnly + | "code" -> LKCode + | "set-options" + | "reset-options" -> LKOption + | "open" + | "let-open" + | "include" + | "module-alias" -> LKModule + | _ -> + js_fail "lookup context (symbol-only, code, set-options, reset-options, \ +open, let-open, include, module-alias)" k + +type position = string * int * int + + +type push_query = + { push_kind: push_kind; + push_line: int; + push_column: int; + push_peek_only: bool; + push_code_or_decl: either string FStar.Parser.AST.decl} + +type query' = +| Exit +| DescribeProtocol +| DescribeRepl +| Segment of string (* File contents *) +| Pop +| Push of push_query +| VfsAdd of option string (* fname *) * string (* contents *) +| AutoComplete of string * completion_context +| Lookup of string * lookup_context * option position * list string +| Compute of string * option (list FStar.TypeChecker.Env.step) +| Search of string +| GenericError of string +| ProtocolViolation of string +| FullBuffer of string +and query = { qq: query'; qid: string } + +let push_query_to_string pq = + let pk = + match pq.push_kind with + | SyntaxCheck -> "SyntaxCheck" + | LaxCheck -> "LaxCheck" + | FullCheck -> "FullCheck" + in + let code_or_decl = + match pq.push_code_or_decl with + | Inl code -> code + | Inr decl -> FStar.Parser.AST.decl_to_string decl + in + FStar.Compiler.Util.format "{ push_kind = %s; push_line = %s; \ + push_column = %s; push_peek_only = %s; push_code_or_decl = %s }" + [pk; string_of_int pq.push_line; + string_of_int pq.push_column; + string_of_bool pq.push_peek_only; + code_or_decl] + +let query_to_string q = match q.qq with +| Exit -> "Exit" +| DescribeProtocol -> "DescribeProtocol" +| DescribeRepl -> "DescribeRepl" +| Segment _ -> "Segment" +| Pop -> "Pop" +| Push pq -> "(Push " ^ push_query_to_string pq ^ ")" +| VfsAdd _ -> "VfsAdd" +| AutoComplete _ -> "AutoComplete" +| Lookup _ -> "Lookup" +| Compute _ -> "Compute" +| Search _ -> "Search" +| GenericError _ -> "GenericError" +| ProtocolViolation _ -> "ProtocolViolation" +| FullBuffer _ -> "FullBuffer" + + + +let query_needs_current_module = function + | Exit | DescribeProtocol | DescribeRepl | Segment _ + | Pop | Push { push_peek_only = false } | VfsAdd _ + | GenericError _ | ProtocolViolation _ | FullBuffer _ -> false + | Push _ | AutoComplete _ | Lookup _ | Compute _ | Search _ -> true + +let interactive_protocol_vernum = 2 + +let interactive_protocol_features = + ["autocomplete"; "autocomplete/context"; + "compute"; "compute/reify"; "compute/pure-subterms"; + "describe-protocol"; "describe-repl"; "exit"; + "lookup"; "lookup/context"; "lookup/documentation"; "lookup/definition"; + "peek"; "pop"; "push"; "search"; "segment"; + "vfs-add"; "tactic-ranges"; "interrupt"; "progress"; "full-buffer"] + +type query_status = | QueryOK | QueryNOK | QueryViolatesProtocol + + +let json_of_issue_level i = + JsonStr (match i with + | ENotImplemented -> "not-implemented" + | EInfo -> "info" + | EWarning -> "warning" + | EError -> "error") + +let json_of_issue issue = + JsonAssoc <| + [("level", json_of_issue_level issue.issue_level)] + @(match issue.issue_number with + | None -> [] + | Some n -> [("number", JsonInt n)]) + @[("message", JsonStr (issue_message issue)); + ("ranges", JsonList + ((match issue.issue_range with + | None -> [] + | Some r -> [json_of_use_range r]) @ + (match issue.issue_range with + | Some r when def_range r <> use_range r -> + [json_of_def_range r] + | _ -> [])))] diff --git a/src/fstar/FStar.Interactive.Ide.fst b/src/fstar/FStar.Interactive.Ide.fst index 2ff32742ee2..619ee24cd02 100644 --- a/src/fstar/FStar.Interactive.Ide.fst +++ b/src/fstar/FStar.Interactive.Ide.fst @@ -28,13 +28,16 @@ open FStar.Errors open FStar.Interactive.JsonHelper open FStar.Interactive.QueryHelper open FStar.Interactive.PushHelper +open FStar.Interactive.ReplState +open FStar.Interactive.Ide.Types +module BU = FStar.Compiler.Util open FStar.Universal open FStar.TypeChecker.Env open FStar.TypeChecker.Common open FStar.Interactive open FStar.Parser.ParseIt - +open FStar.Interactive.Ide.Types module SS = FStar.Syntax.Syntax module DsEnv = FStar.Syntax.DsEnv module TcErr = FStar.TypeChecker.Err @@ -73,28 +76,7 @@ let with_captured_errors' env sigint_handler f = let with_captured_errors env sigint_handler f = if Options.trace_error () then f env else with_captured_errors' env sigint_handler f - -(*************************) -(* REPL tasks and states *) -(*************************) - -let t0 = Util.now () - -(** Create a timed_fname with a dummy modtime **) -let dummy_tf_of_fname fname = - { tf_fname = fname; - tf_modtime = t0 } - -let string_of_timed_fname { tf_fname = fname; tf_modtime = modtime } = - if modtime = t0 then Util.format1 "{ %s }" fname - else Util.format2 "{ %s; %s }" fname (string_of_time modtime) - -type push_query = - { push_kind: push_kind; - push_code: string; - push_line: int; push_column: int; - push_peek_only: bool } - + (** Tasks describing each snapshot of the REPL state. **) type env_t = TcEnv.env @@ -112,17 +94,6 @@ let nothing_left_to_pop st = (* Dependency checks *) (*********************) -let string_of_repl_task = function - | LDInterleaved (intf, impl) -> - Util.format2 "LDInterleaved (%s, %s)" (string_of_timed_fname intf) (string_of_timed_fname impl) - | LDSingle intf_or_impl -> - Util.format1 "LDSingle %s" (string_of_timed_fname intf_or_impl) - | LDInterfaceOfCurrentFile intf -> - Util.format1 "LDInterfaceOfCurrentFile %s" (string_of_timed_fname intf) - | PushFragment frag -> - Util.format1 "PushFragment { code = %s }" frag.frag_text - | Noop -> "Noop {}" - (** Push, run `task`, and pop if it fails. If `must_rollback` is set, always pop. Returns a pair: a boolean indicating @@ -218,106 +189,6 @@ let run_repl_ld_transactions (st: repl_state) (tasks: list repl_task) aux st tasks (List.rev st.repl_deps_stack) -(*****************************************) -(* Reading queries and writing responses *) -(*****************************************) - -let js_pushkind s : push_kind = match js_str s with - | "syntax" -> SyntaxCheck - | "lax" -> LaxCheck - | "full" -> FullCheck - | _ -> js_fail "push_kind" s - -let js_reductionrule s = match js_str s with - | "beta" -> FStar.TypeChecker.Env.Beta - | "delta" -> FStar.TypeChecker.Env.UnfoldUntil SS.delta_constant - | "iota" -> FStar.TypeChecker.Env.Iota - | "zeta" -> FStar.TypeChecker.Env.Zeta - | "reify" -> FStar.TypeChecker.Env.Reify - | "pure-subterms" -> FStar.TypeChecker.Env.PureSubtermsWithinComputations - | _ -> js_fail "reduction rule" s - -type completion_context = -| CKCode -| CKOption of bool (* #set-options (false) or #reset-options (true) *) -| CKModuleOrNamespace of bool (* modules *) * bool (* namespaces *) - -let js_optional_completion_context k = - match k with - | None -> CKCode - | Some k -> - match js_str k with - | "symbol" // Backwards compatibility - | "code" -> CKCode - | "set-options" -> CKOption false - | "reset-options" -> CKOption true - | "open" - | "let-open" -> CKModuleOrNamespace (true, true) - | "include" - | "module-alias" -> CKModuleOrNamespace (true, false) - | _ -> - js_fail "completion context (code, set-options, reset-options, \ -open, let-open, include, module-alias)" k - -type lookup_context = -| LKSymbolOnly -| LKModule -| LKOption -| LKCode - -let js_optional_lookup_context k = - match k with - | None -> LKSymbolOnly // Backwards-compatible default - | Some k -> - match js_str k with - | "symbol-only" -> LKSymbolOnly - | "code" -> LKCode - | "set-options" - | "reset-options" -> LKOption - | "open" - | "let-open" - | "include" - | "module-alias" -> LKModule - | _ -> - js_fail "lookup context (symbol-only, code, set-options, reset-options, \ -open, let-open, include, module-alias)" k - -type position = string * int * int - -type query' = -| Exit -| DescribeProtocol -| DescribeRepl -| Segment of string (* File contents *) -| Pop -| Push of push_query -| VfsAdd of option string (* fname *) * string (* contents *) -| AutoComplete of string * completion_context -| Lookup of string * lookup_context * option position * list string -| Compute of string * option (list FStar.TypeChecker.Env.step) -| Search of string -| GenericError of string -| ProtocolViolation of string -and query = { qq: query'; qid: string } - -let query_needs_current_module = function - | Exit | DescribeProtocol | DescribeRepl | Segment _ - | Pop | Push { push_peek_only = false } | VfsAdd _ - | GenericError _ | ProtocolViolation _ -> false - | Push _ | AutoComplete _ | Lookup _ | Compute _ | Search _ -> true - -let interactive_protocol_vernum = 2 - -let interactive_protocol_features = - ["autocomplete"; "autocomplete/context"; - "compute"; "compute/reify"; "compute/pure-subterms"; - "describe-protocol"; "describe-repl"; "exit"; - "lookup"; "lookup/context"; "lookup/documentation"; "lookup/definition"; - "peek"; "pop"; "push"; "search"; "segment"; - "vfs-add"; "tactic-ranges"; "interrupt"; "progress"] - -type query_status = | QueryOK | QueryNOK | QueryViolatesProtocol - let wrap_js_failure qid expected got = { qid = qid; qq = ProtocolViolation (Util.format2 "JSON decoding failed: expected %s, got %s" @@ -350,10 +221,11 @@ let unpack_interactive_query json = | "describe-repl" -> DescribeRepl | "segment" -> Segment (arg "code" |> js_str) | "peek" | "push" -> Push ({ push_kind = arg "kind" |> js_pushkind; - push_code = arg "code" |> js_str; + push_code_or_decl = Inl (arg "code" |> js_str); push_line = arg "line" |> js_int; push_column = arg "column" |> js_int; push_peek_only = query = "peek" }) + | "full-buffer" -> FullBuffer (arg "code" |> js_str) | "autocomplete" -> AutoComplete (arg "partial-symbol" |> js_str, try_arg "context" |> js_optional_completion_context) | "lookup" -> Lookup (arg "symbol" |> js_str, @@ -396,29 +268,6 @@ let read_interactive_query stream : query = let json_of_opt json_of_a opt_a = Util.dflt JsonNull (Util.map_option json_of_a opt_a) -let json_of_issue_level i = - JsonStr (match i with - | ENotImplemented -> "not-implemented" - | EInfo -> "info" - | EWarning -> "warning" - | EError -> "error") - -let json_of_issue issue = - JsonAssoc <| - [("level", json_of_issue_level issue.issue_level)] - @(match issue.issue_number with - | None -> [] - | Some n -> [("number", JsonInt n)]) - @[("message", JsonStr (issue_message issue)); - ("ranges", JsonList - ((match issue.issue_range with - | None -> [] - | Some r -> [json_of_use_range r]) @ - (match issue.issue_range with - | Some r when def_range r <> use_range r -> - [json_of_def_range r] - | _ -> [])))] - let alist_of_symbol_lookup_result lr = [("name", JsonStr lr.slr_name); ("defined-at", json_of_opt json_of_def_range lr.slr_def_range); @@ -660,6 +509,9 @@ let write_progress stage contents_alist = let js_contents = ("stage", stage) :: contents_alist in write_json (json_of_message "progress" (JsonAssoc js_contents)) +let write_error contents = + write_json (json_of_message "error" (JsonAssoc contents)) + let write_repl_ld_task_progress task = match task with | LDInterleaved (_, tf) | LDSingle tf | LDInterfaceOfCurrentFile tf -> @@ -686,26 +538,65 @@ let rephrase_dependency_error issue = format1 "Error while computing or loading dependencies:\n%s" issue.issue_msg } -let run_push_without_deps st query = +let write_full_buffer_fragment_progress (di:either FStar.Parser.AST.decl (list issue)) = + match di with + | Inl d -> + write_progress (Some "full-buffer-fragment-ok") + ["ranges", json_of_def_range d.FStar.Parser.AST.drange] + | Inr issues -> + let qid = + match !repl_current_qid with + | None -> "unknown" + | Some q -> q + in + write_json (json_of_response qid QueryNOK (JsonList (List.map json_of_issue issues))) + + +let run_push_without_deps st query + : (query_status & json) & either repl_state int = let set_nosynth_flag st flag = { st with repl_env = { st.repl_env with nosynth = flag } } in - let { push_code = text; push_line = line; push_column = column; - push_peek_only = peek_only; push_kind = push_kind } = query in + let { push_code_or_decl = code_or_decl; + push_line = line; + push_column = column; + push_peek_only = peek_only; + push_kind = push_kind } = query in - let frag = { frag_fname = ""; frag_text = text; frag_line = line; frag_col = column } in let _ = if FStar.Options.ide_id_info_off() then TcEnv.toggle_id_info st.repl_env false else TcEnv.toggle_id_info st.repl_env true in + let frag = + match code_or_decl with + | Inl text -> + Inl { frag_fname = ""; frag_text = text; frag_line = line; frag_col = column } + | Inr decl -> + Inr decl + in let st = set_nosynth_flag st peek_only in let success, st = run_repl_transaction st push_kind peek_only (PushFragment frag) in let st = set_nosynth_flag st false in let status = if success || peek_only then QueryOK else QueryNOK in - let json_errors = JsonList (collect_errors () |> List.map json_of_issue) in + let errs = collect_errors () in + let has_error = + List.existsb + (fun i -> + match i.issue_level with + | EError | ENotImplemented -> true + | _ -> false) + errs + in + let _ = + match code_or_decl with + | Inr d when not has_error -> + write_full_buffer_fragment_progress (Inl d) + | _ -> () + in + let json_errors = JsonList (errs |> List.map json_of_issue) in let st = if success then { st with repl_line = line; repl_column = column } else st in ((status, json_errors), Inl st) @@ -843,8 +734,8 @@ let run_with_parsed_and_tc_term st term line column continuation = | _ -> None in let parse frag = - match FStar.Parser.ParseIt.parse (FStar.Parser.ParseIt.Toplevel frag) with - | FStar.Parser.ParseIt.ASTFragment (Inr decls, _) -> Some decls + match FStar.Parser.ParseIt.parse (FStar.Parser.ParseIt.Incremental frag) with + | FStar.Parser.ParseIt.IncrementalFragment (decls, _, _err) -> Some decls | _ -> None in let desugar env decls = @@ -997,21 +888,28 @@ let run_search st search_str = with InvalidSearch s -> (QueryNOK, JsonStr s) in (results, Inl st) -let run_query st (q: query') : (query_status * json) * either repl_state int = - match q with - | Exit -> run_exit st - | DescribeProtocol -> run_describe_protocol st - | DescribeRepl -> run_describe_repl st - | GenericError message -> run_generic_error st message - | ProtocolViolation query -> run_protocol_violation st query - | Segment c -> run_segment st c - | VfsAdd (fname, contents) -> run_vfs_add st fname contents - | Push pquery -> run_push st pquery - | Pop -> run_pop st - | AutoComplete (search_term, context) -> run_autocomplete st search_term context - | Lookup (symbol, context, pos_opt, rq_info) -> run_lookup st symbol context pos_opt rq_info - | Compute (term, rules) -> run_compute st term rules - | Search term -> run_search st term +let as_json_list (q: (query_status & json) & either repl_state int) + : (query_status & list json) & either repl_state int + = let (q, j), s = q in + (q, [j]), s + +let run_query_result = (query_status * list json) * either repl_state int + +let rec fold_query (f:repl_state -> query -> run_query_result) + (l:list query) + (st:repl_state) + (responses: list json) + : run_query_result + = match l with + | [] -> (QueryOK, responses), Inl st + | q::l -> + let (status, resp), st' = f st q in + let responses = responses @ resp in + match status, st' with + | QueryOK, Inl st -> + fold_query f l st responses + | _ -> + (status, responses), st' let validate_query st (q: query) : query = match q.qq with @@ -1022,16 +920,43 @@ let validate_query st (q: query) : query = { qid = q.qid; qq = GenericError "Current module unset" } | _ -> q -let validate_and_run_query st query = + +let rec run_query st (q: query) : (query_status * list json) * either repl_state int = + match q.qq with + | Exit -> as_json_list (run_exit st) + | DescribeProtocol -> as_json_list (run_describe_protocol st) + | DescribeRepl -> as_json_list (run_describe_repl st) + | GenericError message -> as_json_list (run_generic_error st message) + | ProtocolViolation query -> as_json_list (run_protocol_violation st query) + | Segment c -> as_json_list (run_segment st c) + | VfsAdd (fname, contents) -> as_json_list (run_vfs_add st fname contents) + | Push pquery -> as_json_list (run_push st pquery) + | Pop -> as_json_list (run_pop st) + | FullBuffer code -> + let queries = + FStar.Interactive.Incremental.run_full_buffer st q.qid code write_full_buffer_fragment_progress + in + fold_query validate_and_run_query queries st [] + | AutoComplete (search_term, context) -> + as_json_list (run_autocomplete st search_term context) + | Lookup (symbol, context, pos_opt, rq_info) -> + as_json_list (run_lookup st symbol context pos_opt rq_info) + | Compute (term, rules) -> + as_json_list (run_compute st term rules) + | Search term -> + as_json_list (run_search st term) +and validate_and_run_query st query = let query = validate_query st query in repl_current_qid := Some query.qid; - run_query st query.qq + if Options.debug_any () + then BU.print2 "Running query %s: %s\n" query.qid (query_to_string query); + run_query st query (** This is the body of the JavaScript port's main loop. **) let js_repl_eval st query = - let (status, response), st_opt = validate_and_run_query st query in - let js_response = json_of_response query.qid status response in - js_response, st_opt + let (status, responses), st_opt = validate_and_run_query st query in + let js_responses = List.map (json_of_response query.qid status) responses in + js_responses, st_opt let js_repl_eval_js st query_js = js_repl_eval st (deserialize_interactive_query query_js) @@ -1039,7 +964,7 @@ let js_repl_eval_js st query_js = let js_repl_eval_str st query_str = let js_response, st_opt = js_repl_eval st (parse_interactive_query query_str) in - (Util.string_of_json js_response), st_opt + (List.map Util.string_of_json js_response), st_opt (** This too is called from FStar.js **) let js_repl_init_opts () = @@ -1057,9 +982,14 @@ let js_repl_init_opts () = (** This is the main loop for the desktop version **) let rec go st : int = + if Options.debug_any () + then ( + FStar.Compiler.Util.print1 "Repl stack is:\n%s\n" + (string_of_repl_stack (!repl_stack)) + ); let query = read_interactive_query st.repl_stdin in - let (status, response), state_opt = validate_and_run_query st query in - write_response query.qid status response; + let (status, responses), state_opt = validate_and_run_query st query in + List.iter (write_response query.qid status) responses; match state_opt with | Inl st' -> go st' | Inr exitcode -> exitcode @@ -1091,8 +1021,6 @@ let install_ide_mode_hooks printer = FStar.Compiler.Util.set_printer (interactive_printer printer); FStar.Errors.set_handler interactive_error_handler -let initial_range = - Range.mk_range "" (Range.mk_pos 1 0) (Range.mk_pos 1 0) let build_initial_repl_state (filename: string) = let env = init_env FStar.Parser.Dep.empty_deps in diff --git a/src/fstar/FStar.Interactive.Incremental.fst b/src/fstar/FStar.Interactive.Incremental.fst new file mode 100755 index 00000000000..130d9992b3c --- /dev/null +++ b/src/fstar/FStar.Interactive.Incremental.fst @@ -0,0 +1,217 @@ +(* + Copyright 2023 Nikhil Swamy and Microsoft Research + + Licensed under the Apache License, Version 2.0 (the "License"); + you may not use this file except in compliance with the License. + You may obtain a copy of the License at + + http://www.apache.org/licenses/LICENSE-2.0 + + Unless required by applicable law or agreed to in writing, software + distributed under the License is distributed on an "AS IS" BASIS, + WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or implied. + See the License for the specific language governing permissions and + limitations under the License. +*) + +module FStar.Interactive.Incremental +open FStar.Pervasives +open FStar.Compiler.Effect +open FStar.Compiler.List +open FStar +open FStar.Compiler +open FStar.Compiler.Range +open FStar.Compiler.Util +open FStar.Getopt +open FStar.Ident +open FStar.Errors +open FStar.Interactive.JsonHelper +open FStar.Interactive.QueryHelper +open FStar.Interactive.PushHelper +open FStar.Universal +open FStar.TypeChecker.Env +open FStar.TypeChecker.Common +open FStar.Interactive +open FStar.Parser.ParseIt +module SS = FStar.Syntax.Syntax +module DsEnv = FStar.Syntax.DsEnv +module TcErr = FStar.TypeChecker.Err +module TcEnv = FStar.TypeChecker.Env +module CTable = FStar.Interactive.CompletionTable +open FStar.Interactive.Ide.Types +open FStar.Interactive.ReplState +module P = FStar.Parser.ParseIt +module BU = FStar.Compiler.Util +open FStar.Parser.AST +open FStar.Parser.AST.Comparison + +let qid = string & int +let qst a = qid -> a & qid +let return (x:'a) : qst 'a = fun q -> x, q +let (let!) (f:qst 'a) (g: 'a -> qst 'b) + : qst 'b + = fun q -> let x, q' = f q in + g x q' + +let run_qst (f:qst 'a) (q:string) + : 'a + = fst (f (q, 0)) + + +let rec map (f:'a -> qst 'b) (l:list 'a) + : qst (list 'b) + = match l with + | [] -> return [] + | hd::tl -> + let! hd = f hd in + let! tl = map f tl in + return (hd :: tl) + +let shift_qid (q:qid) (i:int) = fst q, snd q + i + +let next_qid + : qst qid + = fun q -> let q = shift_qid q 1 in + q, q + +let get_qid + : qst qid + = fun q -> q, q + +let as_query (q:query') + : qst query + = let! (qid_prefix, i) = next_qid in + return + { + qq=q; + qid=qid_prefix ^ "." ^ string_of_int i + } + +let push_decl (d:decl) + : qst query + = let open FStar.Compiler.Range in + let pq = { + push_kind = FullCheck; + push_line = line_of_pos (start_of_range d.drange); + push_column = col_of_pos (start_of_range d.drange); + push_peek_only = false; + push_code_or_decl = Inr d + } in + as_query (Push pq) + +let push_decls (ds:list decl) + : qst (list query) + = map push_decl ds + +let pop_entries (e:list repl_stack_entry_t) + : qst (list query) + = map (fun _ -> as_query Pop) e + +let response_success (d:decl) + : qst json + = let! (q, _) = get_qid in + let contents = JsonAssoc (["ranges", json_of_def_range d.drange]) in + return (JsonAssoc [("kind", JsonStr "response"); + ("query-id", JsonStr q); + ("status", JsonStr "success"); + ("contents", contents)]) + +let inspect_repl_stack (s:repl_stack_t) + (ds:list decl) + (write_full_buffer_fragment_progress: either decl (list issue) -> unit) + : qst (list query) // & list json) + = let entries = List.rev s in + match BU.prefix_until + (function (_, (PushFragment _, _)) -> true | _ -> false) + entries + with + | None -> + let! ds = push_decls ds in + return ds + + | Some (prefix, first_push, rest) -> + let entries = first_push :: rest in + let repl_task (_, (p, _)) = p in + let rec matching_prefix entries ds + : qst (list query) + = match entries, ds with + | [], [] -> + return [] + + | e::entries, d::ds -> ( + match repl_task e with + | Noop -> + matching_prefix entries (d::ds) + | PushFragment (Inl frag) -> + let! pops = pop_entries (e::entries) in + let! pushes = push_decls (d::ds) in + return (pops @ pushes) + | PushFragment (Inr d') -> + if eq_decl d d' + then ( + write_full_buffer_fragment_progress (Inl d); + matching_prefix entries ds + ) + else let! pops = pop_entries (e::entries) in + let! pushes = push_decls (d::ds) in + return (pops @ pushes) + ) + + | [], ds -> + let! pushes = push_decls ds in + return pushes + + | es, [] -> + let! pops = pop_entries es in + return pops + in + matching_prefix entries ds + +let run_full_buffer (st:repl_state) + (qid:string) + (code:string) + (write_full_buffer_fragment_progress: either decl (list issue) -> unit) + : list query + = let parse_result = + P.parse (Incremental { + frag_fname = Range.file_of_range initial_range; + frag_text = code; + frag_line = Range.line_of_pos (Range.start_of_range initial_range); + frag_col = Range.col_of_pos (Range.start_of_range initial_range); + }) + in + let log_syntax_issues err = + match err with + | None -> () + | Some (raw_error, msg, range) -> + let _, _, num = FStar.Errors.lookup raw_error in + let issue = { + issue_msg = msg; + issue_level = EError; + issue_range = Some range; + issue_number = Some num; + issue_ctx = [] + } in + write_full_buffer_fragment_progress (Inr [issue]) + in + let qs = + match parse_result with + | IncrementalFragment (decls, _, err_opt) -> + let queries = + run_qst (inspect_repl_stack (!repl_stack) decls write_full_buffer_fragment_progress) qid + in + log_syntax_issues err_opt; + if Options.debug_any() + then ( + BU.print1 "Generating queries\n%s\n" + (String.concat "\n" (List.map query_to_string queries)) + ); + queries + + | ParseError err -> + log_syntax_issues (Some err); + [] + | _ -> + failwith "Unexpected parse result" + in + qs diff --git a/src/fstar/FStar.Interactive.JsonHelper.fst b/src/fstar/FStar.Interactive.JsonHelper.fst index 93f15a292fd..f4aa44d5905 100644 --- a/src/fstar/FStar.Interactive.JsonHelper.fst +++ b/src/fstar/FStar.Interactive.JsonHelper.fst @@ -257,3 +257,4 @@ let js_diag (fname: string) (msg: string) (r: option Range.range) : assoct = let js_diag_clear (fname: string) : assoct = [("method", JsonStr "textDocument/publishDiagnostics"); ("params", JsonAssoc [("uri", JsonStr (path_to_uri fname)); ("diagnostics", JsonList [])])] + diff --git a/src/fstar/FStar.Interactive.JsonHelper.fsti b/src/fstar/FStar.Interactive.JsonHelper.fsti index a53dddc08ba..6493a6d24dd 100644 --- a/src/fstar/FStar.Interactive.JsonHelper.fsti +++ b/src/fstar/FStar.Interactive.JsonHelper.fsti @@ -22,13 +22,6 @@ open FStar open FStar.Compiler open FStar.Errors open FStar.Compiler.Util -open FStar.Compiler.Range -open FStar.TypeChecker.Env - -module U = FStar.Compiler.Util -module PI = FStar.Parser.ParseIt -module TcEnv = FStar.TypeChecker.Env -module CTable = FStar.Interactive.CompletionTable // Type of an associative array type assoct = list (string * json) @@ -116,36 +109,6 @@ type lquery = type lsp_query = { query_id: option int; q: lquery } -(* Types concerning repl *) -type repl_depth_t = TcEnv.tcenv_depth_t * int -type optmod_t = option Syntax.Syntax.modul - -type timed_fname = - { tf_fname: string; - tf_modtime: time } - -(** Every snapshot pushed in the repl stack is annotated with one of these. The -``LD``-prefixed (“Load Dependency”) onces are useful when loading or updating -dependencies, as they carry enough information to determine whether a dependency -is stale. **) -type repl_task = - | LDInterleaved of timed_fname * timed_fname (* (interface * implementation) *) - | LDSingle of timed_fname (* interface or implementation *) - | LDInterfaceOfCurrentFile of timed_fname (* interface *) - | PushFragment of PI.input_frag (* code fragment *) - | Noop (* Used by compute *) - -type repl_state = { repl_line: int; repl_column: int; repl_fname: string; - repl_deps_stack: repl_stack_t; - repl_curmod: optmod_t; - repl_env: TcEnv.env; - repl_stdin: stream_reader; - repl_names: CTable.table } -and repl_stack_t = list repl_stack_entry_t -and repl_stack_entry_t = repl_depth_t * (repl_task * repl_state) - -// Global repl_state, keeping state of different buffers -type grepl_state = { grepl_repls: U.psmap repl_state; grepl_stdin: stream_reader } type error_code = | ParseError @@ -199,3 +162,4 @@ val js_diag : string -> string -> option Range.range -> assoct // Build an empty JSON diagnostic; used for clearing diagnostic val js_diag_clear : string -> assoct + diff --git a/src/fstar/FStar.Interactive.Legacy.fst b/src/fstar/FStar.Interactive.Legacy.fst index fbfeaa83f59..c73303ec079 100644 --- a/src/fstar/FStar.Interactive.Legacy.fst +++ b/src/fstar/FStar.Interactive.Legacy.fst @@ -72,7 +72,7 @@ let push_with_kind env lax restore_cmd_line_options msg = let check_frag (env:TcEnv.env) curmod frag = try - let m, env = tc_one_fragment curmod env frag in + let m, env = tc_one_fragment curmod env (Inl frag) in Some (m, env, FStar.Errors.get_err_count()) with | FStar.Errors.Error(e, msg, r, ctx) when not ((Options.trace_error())) -> diff --git a/src/fstar/FStar.Interactive.Lsp.fst b/src/fstar/FStar.Interactive.Lsp.fst index d3678e24c8b..c8eac86789f 100644 --- a/src/fstar/FStar.Interactive.Lsp.fst +++ b/src/fstar/FStar.Interactive.Lsp.fst @@ -24,6 +24,7 @@ open FStar.Compiler.Range open FStar.Errors open FStar.Universal open FStar.Interactive.JsonHelper +open FStar.Interactive.ReplState module U = FStar.Compiler.Util module QH = FStar.Interactive.QueryHelper diff --git a/src/fstar/FStar.Interactive.PushHelper.fst b/src/fstar/FStar.Interactive.PushHelper.fst index 1286593d7ae..3088970a2b9 100644 --- a/src/fstar/FStar.Interactive.PushHelper.fst +++ b/src/fstar/FStar.Interactive.PushHelper.fst @@ -30,6 +30,7 @@ open FStar.Universal open FStar.Parser.ParseIt open FStar.TypeChecker.Env open FStar.Interactive.JsonHelper +open FStar.Interactive.ReplState module U = FStar.Compiler.Util module SS = FStar.Syntax.Syntax @@ -370,5 +371,5 @@ let full_lax text st = match ld_deps st with | Inl (st, deps) -> let names = add_module_completions st.repl_fname deps st.repl_names in - repl_tx ({ st with repl_names = names }) LaxCheck (PushFragment frag) + repl_tx ({ st with repl_names = names }) LaxCheck (PushFragment (Inl frag)) | Inr st -> None, st diff --git a/src/fstar/FStar.Interactive.PushHelper.fsti b/src/fstar/FStar.Interactive.PushHelper.fsti index b708b2797a5..10221554b7c 100644 --- a/src/fstar/FStar.Interactive.PushHelper.fsti +++ b/src/fstar/FStar.Interactive.PushHelper.fsti @@ -26,6 +26,7 @@ open FStar.Compiler.Util open FStar.Ident open FStar.TypeChecker.Env open FStar.Interactive.JsonHelper +open FStar.Interactive.ReplState module DsEnv = FStar.Syntax.DsEnv module CTable = FStar.Interactive.CompletionTable diff --git a/src/fstar/FStar.Interactive.QueryHelper.fst b/src/fstar/FStar.Interactive.QueryHelper.fst index ce5d202c4f2..3983ccf0aa3 100644 --- a/src/fstar/FStar.Interactive.QueryHelper.fst +++ b/src/fstar/FStar.Interactive.QueryHelper.fst @@ -28,6 +28,7 @@ open FStar.Compiler.Util open FStar.TypeChecker.Env open FStar.TypeChecker.Common open FStar.Interactive.JsonHelper +open FStar.Interactive.ReplState open FStar.Interactive.CompletionTable module U = FStar.Compiler.Util diff --git a/src/fstar/FStar.Interactive.QueryHelper.fsti b/src/fstar/FStar.Interactive.QueryHelper.fsti index 1dcb0d7eb80..85552a5e9ed 100644 --- a/src/fstar/FStar.Interactive.QueryHelper.fsti +++ b/src/fstar/FStar.Interactive.QueryHelper.fsti @@ -18,11 +18,13 @@ * queries; this file collects helpers for them *) module FStar.Interactive.QueryHelper -open FStar open FStar.Compiler +open FStar +open FStar.Compiler open FStar.Compiler.Range open FStar.Compiler.Util open FStar.TypeChecker.Env open FStar.Interactive.JsonHelper +open FStar.Interactive.ReplState module TcErr = FStar.TypeChecker.Err module TcEnv = FStar.TypeChecker.Env diff --git a/src/fstar/FStar.Interactive.ReplState.fst b/src/fstar/FStar.Interactive.ReplState.fst new file mode 100755 index 00000000000..ccb73adc1ba --- /dev/null +++ b/src/fstar/FStar.Interactive.ReplState.fst @@ -0,0 +1,121 @@ +(* + Copyright 2019 Microsoft Research + + Licensed under the Apache License, Version 2.0 (the "License"); + you may not use this file except in compliance with the License. + You may obtain a copy of the License at + + http://www.apache.org/licenses/LICENSE-2.0 + + Unless required by applicable law or agreed to in writing, software + distributed under the License is distributed on an "AS IS" BASIS, + WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or implied. + See the License for the specific language governing permissions and + limitations under the License. +*) + +module FStar.Interactive.ReplState +open FStar +open FStar.Compiler +open FStar.Errors +open FStar.Compiler.Util +open FStar.Compiler.Range +open FStar.TypeChecker.Env + +module U = FStar.Compiler.Util +module TcEnv = FStar.TypeChecker.Env +module PI = FStar.Parser.ParseIt +module CTable = FStar.Interactive.CompletionTable + + +(* Types concerning repl *) +type repl_depth_t = TcEnv.tcenv_depth_t * int +type optmod_t = option Syntax.Syntax.modul + +type timed_fname = + { tf_fname: string; + tf_modtime: time } + +(** Every snapshot pushed in the repl stack is annotated with one of these. The +``LD``-prefixed (“Load Dependency”) onces are useful when loading or updating +dependencies, as they carry enough information to determine whether a dependency +is stale. **) +type repl_task = + | LDInterleaved of timed_fname * timed_fname (* (interface * implementation) *) + | LDSingle of timed_fname (* interface or implementation *) + | LDInterfaceOfCurrentFile of timed_fname (* interface *) + | PushFragment of either PI.input_frag FStar.Parser.AST.decl (* code fragment *) + | Noop (* Used by compute *) + +type repl_state = { repl_line: int; repl_column: int; repl_fname: string; + repl_deps_stack: repl_stack_t; + repl_curmod: optmod_t; + repl_env: TcEnv.env; + repl_stdin: stream_reader; + repl_names: CTable.table } +and repl_stack_t = list repl_stack_entry_t +and repl_stack_entry_t = repl_depth_t * (repl_task * repl_state) + +// Global repl_state, keeping state of different buffers +type grepl_state = { grepl_repls: U.psmap repl_state; grepl_stdin: stream_reader } + + +(*************************) +(* REPL tasks and states *) +(*************************) + +let t0 = Util.now () + +(** Create a timed_fname with a dummy modtime **) +let dummy_tf_of_fname fname = + { tf_fname = fname; + tf_modtime = t0 } + +let string_of_timed_fname { tf_fname = fname; tf_modtime = modtime } = + if modtime = t0 then Util.format1 "{ %s }" fname + else Util.format2 "{ %s; %s }" fname (string_of_time modtime) + +let string_of_repl_task = function + | LDInterleaved (intf, impl) -> + Util.format2 "LDInterleaved (%s, %s)" (string_of_timed_fname intf) (string_of_timed_fname impl) + | LDSingle intf_or_impl -> + Util.format1 "LDSingle %s" (string_of_timed_fname intf_or_impl) + | LDInterfaceOfCurrentFile intf -> + Util.format1 "LDInterfaceOfCurrentFile %s" (string_of_timed_fname intf) + | PushFragment (Inl frag) -> + Util.format1 "PushFragment { code = %s }" frag.frag_text + | PushFragment (Inr d) -> + Util.format1 "PushFragment { decl = %s }" (FStar.Parser.AST.decl_to_string d) + | Noop -> "Noop {}" + +module BU = FStar.Compiler.Util + +let string_of_repl_stack_entry + : repl_stack_entry_t -> string + = fun ((depth, i), (task, state)) -> + BU.format "{depth=%s; task=%s}" + [string_of_int i; + string_of_repl_task task] + + +let string_of_repl_stack s = + String.concat ";\n\t\t" + (List.map string_of_repl_stack_entry s) + +let repl_state_to_string (r:repl_state) + : string + = BU.format + "{\n\t\ + repl_line=%s;\n\t\ + repl_column=%s;\n\t\ + repl_fname=%s;\n\t\ + repl_cur_mod=%s;\n\t\ + repl_deps_stack={%s}\n\ + }" + [string_of_int r.repl_line; + string_of_int r.repl_column; + r.repl_fname; + (match r.repl_curmod with + | None -> "None" + | Some m -> Ident.string_of_lid m.name); + string_of_repl_stack r.repl_deps_stack] diff --git a/src/fstar/FStar.Universal.fst b/src/fstar/FStar.Universal.fst index 186abdf7b8f..24d51ebdcf1 100644 --- a/src/fstar/FStar.Universal.fst +++ b/src/fstar/FStar.Universal.fst @@ -174,35 +174,32 @@ let tc_one_fragment curmod (env:TcEnv.env_t) frag = | Parser.AST.Interface (_, { Parser.AST.drange = d } :: _, _) -> d | _ -> Range.dummyRange in - match Parser.Driver.parse_fragment frag with - | Parser.Driver.Empty - | Parser.Driver.Decls [] -> - (curmod, env) - - | Parser.Driver.Modul ast_modul -> - (* It may seem surprising that this function, whose name indicates that - it type-checks a fragment, can actually parse an entire module. - Actually, this is an abuse, and just means that we're type-checking the - first chunk. *) - let ast_modul, env = - with_dsenv_of_tcenv env <| FStar.ToSyntax.Interleave.interleave_module ast_modul false in - let modul, env = - with_dsenv_of_tcenv env <| Desugar.partial_ast_modul_to_modul curmod ast_modul in - if not (acceptable_mod_name modul) then - begin - let msg : string = - BU.format1 "Interactive mode only supports a single module at the top-level. Expected module %s" - (Parser.Dep.module_name_of_file (fname env)) - in - Errors.raise_error (Errors.Fatal_NonSingletonTopLevelModule, msg) - (range_of_first_mod_decl ast_modul) - end; - let modul, env = - if DsEnv.syntax_only env.dsenv then modul, env - else Tc.tc_partial_modul env modul - in - (Some modul, env) - | Parser.Driver.Decls ast_decls -> + let check_module_name_declaration ast_modul = + (* It may seem surprising that this function, whose name indicates that + it type-checks a fragment, can actually parse an entire module. + Actually, this is an abuse, and just means that we're type-checking the + first chunk. *) + let ast_modul, env = + with_dsenv_of_tcenv env <| FStar.ToSyntax.Interleave.interleave_module ast_modul false in + let modul, env = + with_dsenv_of_tcenv env <| Desugar.partial_ast_modul_to_modul curmod ast_modul in + if not (acceptable_mod_name modul) then + begin + let msg : string = + BU.format1 "Interactive mode only supports a single module at the top-level. Expected module %s" + (Parser.Dep.module_name_of_file (fname env)) + in + Errors.raise_error (Errors.Fatal_NonSingletonTopLevelModule, msg) + (range_of_first_mod_decl ast_modul) + end; + let modul, env = + if DsEnv.syntax_only env.dsenv then modul, env + else Tc.tc_partial_modul env modul + in + (Some modul, env) + in + + let check_decls ast_decls = match curmod with | None -> let { Parser.AST.drange = rng } = List.hd ast_decls in @@ -222,7 +219,29 @@ let tc_one_fragment curmod (env:TcEnv.env_t) frag = let modul, _, env = if DsEnv.syntax_only env.dsenv then (modul, [], env) else Tc.tc_more_partial_modul env modul sigelts in (Some modul, env) - + in + match frag with + | Inr d -> ( + match d.d with + | FStar.Parser.AST.TopLevelModule lid -> + check_module_name_declaration (FStar.Parser.AST.Module(lid, [d])) + | _ -> + check_decls [d] + ) + + | Inl frag -> ( + match Parser.Driver.parse_fragment frag with + | Parser.Driver.Empty + | Parser.Driver.Decls [] -> + (curmod, env) + + | Parser.Driver.Modul ast_modul -> + check_module_name_declaration ast_modul + + | Parser.Driver.Decls ast_decls -> + check_decls ast_decls + ) + let load_interface_decls env interface_file_name : TcEnv.env_t = let r = Pars.parse (Pars.Filename interface_file_name) in match r with diff --git a/src/fstar/FStar.Universal.fsti b/src/fstar/FStar.Universal.fsti index 64e11f2dc90..d802a79db9e 100644 --- a/src/fstar/FStar.Universal.fsti +++ b/src/fstar/FStar.Universal.fsti @@ -45,7 +45,7 @@ val core_check: TcEnv.core_check_t val tc_one_fragment : option Syntax.modul -> TcEnv.env_t -> - FStar.Parser.ParseIt.input_frag -> + either FStar.Parser.ParseIt.input_frag FStar.Parser.AST.decl -> option Syntax.modul * TcEnv.env (* Load an interface file into the dsenv. *) diff --git a/src/tests/FStar.Tests.Pars.fst b/src/tests/FStar.Tests.Pars.fst index c26c73f2250..27c8ce9cd4b 100644 --- a/src/tests/FStar.Tests.Pars.fst +++ b/src/tests/FStar.Tests.Pars.fst @@ -164,7 +164,7 @@ let pars_and_tc_fragment (s:string) = let tcenv = init() in let frag = frag_of_text s in try - let test_mod', tcenv' = FStar.Universal.tc_one_fragment !test_mod_ref tcenv (frag, None) in + let test_mod', tcenv' = FStar.Universal.tc_one_fragment !test_mod_ref tcenv (Inl frag) in test_mod_ref := test_mod'; tcenv_ref := Some tcenv'; let n = get_err_count () in From 45fb20f977ed99b65aeb30d4a722d3b72dc868cb Mon Sep 17 00:00:00 2001 From: Nikhil Swamy Date: Fri, 10 Mar 2023 14:29:30 -0800 Subject: [PATCH 07/48] a test for the full-buffer IDE message --- tests/ide/emacs/FullBuffer.fst | 3 +++ tests/ide/emacs/FullBuffer.full-buffer.in | 2 ++ tests/ide/emacs/FullBuffer.full-buffer.out.expected | 6 ++++++ 3 files changed, 11 insertions(+) create mode 100755 tests/ide/emacs/FullBuffer.fst create mode 100755 tests/ide/emacs/FullBuffer.full-buffer.in create mode 100644 tests/ide/emacs/FullBuffer.full-buffer.out.expected diff --git a/tests/ide/emacs/FullBuffer.fst b/tests/ide/emacs/FullBuffer.fst new file mode 100755 index 00000000000..eaa42fcb598 --- /dev/null +++ b/tests/ide/emacs/FullBuffer.fst @@ -0,0 +1,3 @@ +module FullBuffer +let f x = match x with | Some x -> true | None -> false +let test y = if Some? y then f y else true diff --git a/tests/ide/emacs/FullBuffer.full-buffer.in b/tests/ide/emacs/FullBuffer.full-buffer.in new file mode 100755 index 00000000000..497765a9d64 --- /dev/null +++ b/tests/ide/emacs/FullBuffer.full-buffer.in @@ -0,0 +1,2 @@ +{"query-id":"1","query":"vfs-add","args":{"filename":null,"contents":"module FullBuffer\nlet f x = match x with | Some x -> true | None -> false\nlet test y = if Some? y then f y else true\n"}} +{"query-id":"2","query":"full-buffer","args":{"kind":"full","code":"module FullBuffer\nlet f x = match x with | Some x -> true | None -> false\nlet test y = if Some? y then f y else true\nlet test2 : bool = 0\n","line":1,"column":0}} diff --git a/tests/ide/emacs/FullBuffer.full-buffer.out.expected b/tests/ide/emacs/FullBuffer.full-buffer.out.expected new file mode 100644 index 00000000000..9977c0c2e51 --- /dev/null +++ b/tests/ide/emacs/FullBuffer.full-buffer.out.expected @@ -0,0 +1,6 @@ +{"kind": "protocol-info", "rest": "[...]"} +{"kind": "response", "query-id": "1", "response": null, "status": "success"} +{"kind": "response", "query-id": "2", "response": [], "status": "failure"} +{"kind": "response", "query-id": "2", "response": [], "status": "failure"} +{"kind": "response", "query-id": "2", "response": [], "status": "failure"} +{"kind": "response", "query-id": "2", "response": [{"level": "error", "message": "Expected expression of type \"bool\"; got expression \"0\" of type \"int\"", "number": 189, "ranges": [{"beg": [4, 19], "end": [4, 20], "fname": ""}]}], "status": "failure"} From d7e29888b35a2a3455bb4a6f2b9aba261601d1ed Mon Sep 17 00:00:00 2001 From: Nikhil Swamy Date: Fri, 10 Mar 2023 22:52:36 -0800 Subject: [PATCH 08/48] supporting vscode symbol resolution on hover; and status queries on the repl stack --- .../generated/FStar_Interactive_Ide.ml | 261 ++++++++++-------- .../generated/FStar_Interactive_Ide_Types.ml | 11 +- .../FStar_Interactive_Incremental.ml | 154 ++++++----- src/fstar/FStar.Interactive.Ide.Types.fst | 6 +- src/fstar/FStar.Interactive.Ide.fst | 48 ++-- src/fstar/FStar.Interactive.Incremental.fst | 7 +- src/fstar/FStar.Interactive.Incremental.fsti | 1 + 7 files changed, 269 insertions(+), 219 deletions(-) diff --git a/ocaml/fstar-lib/generated/FStar_Interactive_Ide.ml b/ocaml/fstar-lib/generated/FStar_Interactive_Ide.ml index cdd3c3b616a..db4c490b5de 100644 --- a/ocaml/fstar-lib/generated/FStar_Interactive_Ide.ml +++ b/ocaml/fstar-lib/generated/FStar_Interactive_Ide.ml @@ -319,6 +319,20 @@ let (unpack_interactive_query : | FStar_Pervasives_Native.Some (FStar_Compiler_Util.JsonNull) -> FStar_Pervasives_Native.None | other -> other in + let read_position err loc = + let uu___1 = + let uu___2 = assoc err "filename" loc in + FStar_Compiler_Effect.op_Bar_Greater uu___2 + FStar_Interactive_JsonHelper.js_str in + let uu___2 = + let uu___3 = assoc err "line" loc in + FStar_Compiler_Effect.op_Bar_Greater uu___3 + FStar_Interactive_JsonHelper.js_int in + let uu___3 = + let uu___4 = assoc err "column" loc in + FStar_Compiler_Effect.op_Bar_Greater uu___4 + FStar_Interactive_JsonHelper.js_int in + (uu___1, uu___2, uu___3) in let uu___1 = match query with | "exit" -> FStar_Interactive_Ide_Types.Exit @@ -392,9 +406,17 @@ let (unpack_interactive_query : FStar_Interactive_Ide_Types.Push uu___2 | "full-buffer" -> let uu___2 = - let uu___3 = arg "code" in - FStar_Compiler_Effect.op_Bar_Greater uu___3 - FStar_Interactive_JsonHelper.js_str in + let uu___3 = + let uu___4 = arg "code" in + FStar_Compiler_Effect.op_Bar_Greater uu___4 + FStar_Interactive_JsonHelper.js_str in + let uu___4 = + let uu___5 = + let uu___6 = arg "kind" in + FStar_Compiler_Effect.op_Bar_Greater uu___6 + FStar_Interactive_JsonHelper.js_str in + uu___5 = "full" in + (uu___3, uu___4) in FStar_Interactive_Ide_Types.FullBuffer uu___2 | "autocomplete" -> let uu___2 = @@ -426,28 +448,14 @@ let (unpack_interactive_query : FStar_Interactive_JsonHelper.js_assoc) in FStar_Compiler_Effect.op_Bar_Greater uu___6 (FStar_Compiler_Util.map_option - (fun loc -> - let uu___7 = - let uu___8 = - assoc "[location]" "filename" loc in - FStar_Compiler_Effect.op_Bar_Greater uu___8 - FStar_Interactive_JsonHelper.js_str in - let uu___8 = - let uu___9 = assoc "[location]" "line" loc in - FStar_Compiler_Effect.op_Bar_Greater uu___9 - FStar_Interactive_JsonHelper.js_int in - let uu___9 = - let uu___10 = - assoc "[location]" "column" loc in - FStar_Compiler_Effect.op_Bar_Greater uu___10 - FStar_Interactive_JsonHelper.js_int in - (uu___7, uu___8, uu___9))) in + (read_position "[location]")) in let uu___6 = let uu___7 = arg "requested-info" in FStar_Compiler_Effect.op_Bar_Greater uu___7 (FStar_Interactive_JsonHelper.js_list FStar_Interactive_JsonHelper.js_str) in - (uu___3, uu___4, uu___5, uu___6) in + let uu___7 = try_arg "symbol-range" in + (uu___3, uu___4, uu___5, uu___6, uu___7) in FStar_Interactive_Ide_Types.Lookup uu___2 | "compute" -> let uu___2 = @@ -547,42 +555,53 @@ let json_of_opt : FStar_Compiler_Util.dflt FStar_Compiler_Util.JsonNull uu___ let (alist_of_symbol_lookup_result : FStar_Interactive_QueryHelper.sl_reponse -> - (Prims.string * FStar_Compiler_Util.json) Prims.list) + Prims.string -> + FStar_Compiler_Util.json FStar_Pervasives_Native.option -> + (Prims.string * FStar_Compiler_Util.json) Prims.list) = fun lr -> - let uu___ = - let uu___1 = - let uu___2 = - json_of_opt FStar_Compiler_Range.json_of_def_range - lr.FStar_Interactive_QueryHelper.slr_def_range in - ("defined-at", uu___2) in - let uu___2 = - let uu___3 = - let uu___4 = - json_of_opt (fun uu___5 -> FStar_Compiler_Util.JsonStr uu___5) - lr.FStar_Interactive_QueryHelper.slr_typ in - ("type", uu___4) in - let uu___4 = - let uu___5 = - let uu___6 = - json_of_opt (fun uu___7 -> FStar_Compiler_Util.JsonStr uu___7) - lr.FStar_Interactive_QueryHelper.slr_doc in - ("documentation", uu___6) in - let uu___6 = - let uu___7 = - let uu___8 = + fun symbol -> + fun symrange_opt -> + let uu___ = + let uu___1 = + let uu___2 = + json_of_opt FStar_Compiler_Range.json_of_def_range + lr.FStar_Interactive_QueryHelper.slr_def_range in + ("defined-at", uu___2) in + let uu___2 = + let uu___3 = + let uu___4 = json_of_opt - (fun uu___9 -> FStar_Compiler_Util.JsonStr uu___9) - lr.FStar_Interactive_QueryHelper.slr_def in - ("definition", uu___8) in - [uu___7] in - uu___5 :: uu___6 in - uu___3 :: uu___4 in - uu___1 :: uu___2 in - ("name", - (FStar_Compiler_Util.JsonStr - (lr.FStar_Interactive_QueryHelper.slr_name))) - :: uu___ + (fun uu___5 -> FStar_Compiler_Util.JsonStr uu___5) + lr.FStar_Interactive_QueryHelper.slr_typ in + ("type", uu___4) in + let uu___4 = + let uu___5 = + let uu___6 = + json_of_opt + (fun uu___7 -> FStar_Compiler_Util.JsonStr uu___7) + lr.FStar_Interactive_QueryHelper.slr_doc in + ("documentation", uu___6) in + let uu___6 = + let uu___7 = + let uu___8 = + json_of_opt + (fun uu___9 -> FStar_Compiler_Util.JsonStr uu___9) + lr.FStar_Interactive_QueryHelper.slr_def in + ("definition", uu___8) in + let uu___8 = + let uu___9 = + let uu___10 = json_of_opt (fun x -> x) symrange_opt in + ("symbol-range", uu___10) in + [uu___9; ("symbol", (FStar_Compiler_Util.JsonStr symbol))] in + uu___7 :: uu___8 in + uu___5 :: uu___6 in + uu___3 :: uu___4 in + uu___1 :: uu___2 in + ("name", + (FStar_Compiler_Util.JsonStr + (lr.FStar_Interactive_QueryHelper.slr_name))) + :: uu___ let (alist_of_protocol_info : (Prims.string * FStar_Compiler_Util.json) Prims.list) = let js_version = @@ -1519,27 +1538,31 @@ let (run_symbol_lookup : FStar_Interactive_QueryHelper.position FStar_Pervasives_Native.option -> Prims.string Prims.list -> - (Prims.string, - (Prims.string * (Prims.string * FStar_Compiler_Util.json) - Prims.list)) - FStar_Pervasives.either) + FStar_Compiler_Util.json FStar_Pervasives_Native.option -> + (Prims.string, + (Prims.string * (Prims.string * FStar_Compiler_Util.json) + Prims.list)) + FStar_Pervasives.either) = fun st -> fun symbol -> fun pos_opt -> fun requested_info -> - let uu___ = - FStar_Interactive_QueryHelper.symlookup - st.FStar_Interactive_ReplState.repl_env symbol pos_opt - requested_info in - match uu___ with - | FStar_Pervasives_Native.None -> - FStar_Pervasives.Inl "Symbol not found" - | FStar_Pervasives_Native.Some result -> - let uu___1 = - let uu___2 = alist_of_symbol_lookup_result result in - ("symbol", uu___2) in - FStar_Pervasives.Inr uu___1 + fun symbol_range_opt -> + let uu___ = + FStar_Interactive_QueryHelper.symlookup + st.FStar_Interactive_ReplState.repl_env symbol pos_opt + requested_info in + match uu___ with + | FStar_Pervasives_Native.None -> + FStar_Pervasives.Inl "Symbol not found" + | FStar_Pervasives_Native.Some result -> + let uu___1 = + let uu___2 = + alist_of_symbol_lookup_result result symbol + symbol_range_opt in + ("symbol", uu___2) in + FStar_Pervasives.Inr uu___1 let (run_option_lookup : Prims.string -> (Prims.string, @@ -1599,25 +1622,28 @@ let (run_code_lookup : FStar_Interactive_QueryHelper.position FStar_Pervasives_Native.option -> Prims.string Prims.list -> - (Prims.string, - (Prims.string * (Prims.string * FStar_Compiler_Util.json) - Prims.list)) - FStar_Pervasives.either) + FStar_Compiler_Util.json FStar_Pervasives_Native.option -> + (Prims.string, + (Prims.string * (Prims.string * FStar_Compiler_Util.json) + Prims.list)) + FStar_Pervasives.either) = fun st -> fun symbol -> fun pos_opt -> fun requested_info -> - let uu___ = run_symbol_lookup st symbol pos_opt requested_info in - match uu___ with - | FStar_Pervasives.Inr alist -> FStar_Pervasives.Inr alist - | FStar_Pervasives.Inl uu___1 -> - let uu___2 = run_module_lookup st symbol in - (match uu___2 with - | FStar_Pervasives.Inr alist -> FStar_Pervasives.Inr alist - | FStar_Pervasives.Inl err_msg -> - FStar_Pervasives.Inl - "No such symbol, module, or namespace.") + fun symrange_opt -> + let uu___ = + run_symbol_lookup st symbol pos_opt requested_info symrange_opt in + match uu___ with + | FStar_Pervasives.Inr alist -> FStar_Pervasives.Inr alist + | FStar_Pervasives.Inl uu___1 -> + let uu___2 = run_module_lookup st symbol in + (match uu___2 with + | FStar_Pervasives.Inr alist -> FStar_Pervasives.Inr alist + | FStar_Pervasives.Inl err_msg -> + FStar_Pervasives.Inl + "No such symbol, module, or namespace.") let (run_lookup' : FStar_Interactive_ReplState.repl_state -> Prims.string -> @@ -1625,25 +1651,27 @@ let (run_lookup' : FStar_Interactive_QueryHelper.position FStar_Pervasives_Native.option -> Prims.string Prims.list -> - (Prims.string, - (Prims.string * (Prims.string * FStar_Compiler_Util.json) - Prims.list)) - FStar_Pervasives.either) + FStar_Compiler_Util.json FStar_Pervasives_Native.option -> + (Prims.string, + (Prims.string * (Prims.string * FStar_Compiler_Util.json) + Prims.list)) + FStar_Pervasives.either) = fun st -> fun symbol -> fun context -> fun pos_opt -> fun requested_info -> - match context with - | FStar_Interactive_Ide_Types.LKSymbolOnly -> - run_symbol_lookup st symbol pos_opt requested_info - | FStar_Interactive_Ide_Types.LKModule -> - run_module_lookup st symbol - | FStar_Interactive_Ide_Types.LKOption -> - run_option_lookup symbol - | FStar_Interactive_Ide_Types.LKCode -> - run_code_lookup st symbol pos_opt requested_info + fun symrange -> + match context with + | FStar_Interactive_Ide_Types.LKSymbolOnly -> + run_symbol_lookup st symbol pos_opt requested_info symrange + | FStar_Interactive_Ide_Types.LKModule -> + run_module_lookup st symbol + | FStar_Interactive_Ide_Types.LKOption -> + run_option_lookup symbol + | FStar_Interactive_Ide_Types.LKCode -> + run_code_lookup st symbol pos_opt requested_info symrange let run_lookup : 'uuuuu . FStar_Interactive_ReplState.repl_state -> @@ -1652,27 +1680,30 @@ let run_lookup : FStar_Interactive_QueryHelper.position FStar_Pervasives_Native.option -> Prims.string Prims.list -> - ((FStar_Interactive_Ide_Types.query_status * - FStar_Compiler_Util.json) * - (FStar_Interactive_ReplState.repl_state, 'uuuuu) - FStar_Pervasives.either) + FStar_Compiler_Util.json FStar_Pervasives_Native.option -> + ((FStar_Interactive_Ide_Types.query_status * + FStar_Compiler_Util.json) * + (FStar_Interactive_ReplState.repl_state, 'uuuuu) + FStar_Pervasives.either) = fun st -> fun symbol -> fun context -> fun pos_opt -> fun requested_info -> - let uu___ = run_lookup' st symbol context pos_opt requested_info in - match uu___ with - | FStar_Pervasives.Inl err_msg -> - ((FStar_Interactive_Ide_Types.QueryNOK, - (FStar_Compiler_Util.JsonStr err_msg)), - (FStar_Pervasives.Inl st)) - | FStar_Pervasives.Inr (kind, info) -> - ((FStar_Interactive_Ide_Types.QueryOK, - (FStar_Compiler_Util.JsonAssoc - (("kind", (FStar_Compiler_Util.JsonStr kind)) :: info))), - (FStar_Pervasives.Inl st)) + fun symrange -> + let uu___ = + run_lookup' st symbol context pos_opt requested_info symrange in + match uu___ with + | FStar_Pervasives.Inl err_msg -> + ((FStar_Interactive_Ide_Types.QueryNOK, + (FStar_Compiler_Util.JsonStr err_msg)), + (FStar_Pervasives.Inl st)) + | FStar_Pervasives.Inr (kind, info) -> + ((FStar_Interactive_Ide_Types.QueryOK, + (FStar_Compiler_Util.JsonAssoc + (("kind", (FStar_Compiler_Util.JsonStr kind)) :: + info))), (FStar_Pervasives.Inl st)) let run_code_autocomplete : 'uuuuu . FStar_Interactive_ReplState.repl_state -> @@ -2338,18 +2369,18 @@ let rec (run_query : let uu___ = run_push st pquery in as_json_list uu___ | FStar_Interactive_Ide_Types.Pop -> let uu___ = run_pop st in as_json_list uu___ - | FStar_Interactive_Ide_Types.FullBuffer code -> + | FStar_Interactive_Ide_Types.FullBuffer (code, full) -> let queries = FStar_Interactive_Incremental.run_full_buffer st - q.FStar_Interactive_Ide_Types.qid code + q.FStar_Interactive_Ide_Types.qid code full write_full_buffer_fragment_progress in fold_query validate_and_run_query queries st [] | FStar_Interactive_Ide_Types.AutoComplete (search_term1, context) -> let uu___ = run_autocomplete st search_term1 context in as_json_list uu___ | FStar_Interactive_Ide_Types.Lookup - (symbol, context, pos_opt, rq_info) -> - let uu___ = run_lookup st symbol context pos_opt rq_info in + (symbol, context, pos_opt, rq_info, symrange) -> + let uu___ = run_lookup st symbol context pos_opt rq_info symrange in as_json_list uu___ | FStar_Interactive_Ide_Types.Compute (term, rules) -> let uu___ = run_compute st term rules in as_json_list uu___ diff --git a/ocaml/fstar-lib/generated/FStar_Interactive_Ide_Types.ml b/ocaml/fstar-lib/generated/FStar_Interactive_Ide_Types.ml index c28b680d470..e81383bb94a 100755 --- a/ocaml/fstar-lib/generated/FStar_Interactive_Ide_Types.ml +++ b/ocaml/fstar-lib/generated/FStar_Interactive_Ide_Types.ml @@ -136,6 +136,7 @@ let (__proj__Mkpush_query__item__push_code_or_decl : match projectee with | { push_kind; push_line; push_column; push_peek_only; push_code_or_decl;_} -> push_code_or_decl +type lookup_symbol_range = FStar_Compiler_Util.json type query' = | Exit | DescribeProtocol @@ -146,13 +147,14 @@ type query' = | VfsAdd of (Prims.string FStar_Pervasives_Native.option * Prims.string) | AutoComplete of (Prims.string * completion_context) | Lookup of (Prims.string * lookup_context * position - FStar_Pervasives_Native.option * Prims.string Prims.list) + FStar_Pervasives_Native.option * Prims.string Prims.list * + lookup_symbol_range FStar_Pervasives_Native.option) | Compute of (Prims.string * FStar_TypeChecker_Env.step Prims.list FStar_Pervasives_Native.option) | Search of Prims.string | GenericError of Prims.string | ProtocolViolation of Prims.string - | FullBuffer of Prims.string + | FullBuffer of (Prims.string * Prims.bool) and query = { qq: query' ; qid: Prims.string } @@ -190,7 +192,8 @@ let (uu___is_Lookup : query' -> Prims.bool) = let (__proj__Lookup__item___0 : query' -> (Prims.string * lookup_context * position FStar_Pervasives_Native.option - * Prims.string Prims.list)) + * Prims.string Prims.list * lookup_symbol_range + FStar_Pervasives_Native.option)) = fun projectee -> match projectee with | Lookup _0 -> _0 let (uu___is_Compute : query' -> Prims.bool) = fun projectee -> match projectee with | Compute _0 -> true | uu___ -> false @@ -216,7 +219,7 @@ let (__proj__ProtocolViolation__item___0 : query' -> Prims.string) = let (uu___is_FullBuffer : query' -> Prims.bool) = fun projectee -> match projectee with | FullBuffer _0 -> true | uu___ -> false -let (__proj__FullBuffer__item___0 : query' -> Prims.string) = +let (__proj__FullBuffer__item___0 : query' -> (Prims.string * Prims.bool)) = fun projectee -> match projectee with | FullBuffer _0 -> _0 let (__proj__Mkquery__item__qq : query -> query') = fun projectee -> match projectee with | { qq; qid;_} -> qq diff --git a/ocaml/fstar-lib/generated/FStar_Interactive_Incremental.ml b/ocaml/fstar-lib/generated/FStar_Interactive_Incremental.ml index 171a63d9b64..cfc4c8d8a4b 100755 --- a/ocaml/fstar-lib/generated/FStar_Interactive_Incremental.ml +++ b/ocaml/fstar-lib/generated/FStar_Interactive_Incremental.ml @@ -174,83 +174,89 @@ let (run_full_buffer : FStar_Interactive_ReplState.repl_state -> Prims.string -> Prims.string -> - ((FStar_Parser_AST.decl, FStar_Errors.issue Prims.list) - FStar_Pervasives.either -> unit) - -> FStar_Interactive_Ide_Types.query Prims.list) + Prims.bool -> + ((FStar_Parser_AST.decl, FStar_Errors.issue Prims.list) + FStar_Pervasives.either -> unit) + -> FStar_Interactive_Ide_Types.query Prims.list) = fun st -> fun qid1 -> fun code -> - fun write_full_buffer_fragment_progress -> - let parse_result = - let uu___ = - let uu___1 = - let uu___2 = - FStar_Compiler_Range.file_of_range - FStar_Interactive_Ide_Types.initial_range in - let uu___3 = - let uu___4 = - FStar_Compiler_Range.start_of_range - FStar_Interactive_Ide_Types.initial_range in - FStar_Compiler_Range.line_of_pos uu___4 in - let uu___4 = - let uu___5 = - FStar_Compiler_Range.start_of_range + fun full -> + fun write_full_buffer_fragment_progress -> + let parse_result = + let uu___ = + let uu___1 = + let uu___2 = + FStar_Compiler_Range.file_of_range FStar_Interactive_Ide_Types.initial_range in - FStar_Compiler_Range.col_of_pos uu___5 in - { - FStar_Parser_ParseIt.frag_fname = uu___2; - FStar_Parser_ParseIt.frag_text = code; - FStar_Parser_ParseIt.frag_line = uu___3; - FStar_Parser_ParseIt.frag_col = uu___4 - } in - FStar_Parser_ParseIt.Incremental uu___1 in - FStar_Parser_ParseIt.parse uu___ in - let log_syntax_issues err = - match err with - | FStar_Pervasives_Native.None -> () - | FStar_Pervasives_Native.Some (raw_error, msg, range) -> - let uu___ = FStar_Errors.lookup raw_error in - (match uu___ with - | (uu___1, uu___2, num) -> - let issue = - { - FStar_Errors.issue_msg = msg; - FStar_Errors.issue_level = FStar_Errors.EError; - FStar_Errors.issue_range = - (FStar_Pervasives_Native.Some range); - FStar_Errors.issue_number = - (FStar_Pervasives_Native.Some num); - FStar_Errors.issue_ctx = [] - } in - write_full_buffer_fragment_progress - (FStar_Pervasives.Inr [issue])) in - let qs = - match parse_result with - | FStar_Parser_ParseIt.IncrementalFragment - (decls, uu___, err_opt) -> - let queries = - let uu___1 = - let uu___2 = - FStar_Compiler_Effect.op_Bang - FStar_Interactive_PushHelper.repl_stack in - inspect_repl_stack uu___2 decls - write_full_buffer_fragment_progress in - run_qst uu___1 qid1 in - (log_syntax_issues err_opt; - (let uu___3 = FStar_Options.debug_any () in - if uu___3 - then + let uu___3 = let uu___4 = - let uu___5 = - FStar_Compiler_List.map - FStar_Interactive_Ide_Types.query_to_string queries in - FStar_String.concat "\n" uu___5 in - FStar_Compiler_Util.print1 "Generating queries\n%s\n" - uu___4 - else ()); - queries) - | FStar_Parser_ParseIt.ParseError err -> - (log_syntax_issues (FStar_Pervasives_Native.Some err); []) - | uu___ -> failwith "Unexpected parse result" in - qs \ No newline at end of file + FStar_Compiler_Range.start_of_range + FStar_Interactive_Ide_Types.initial_range in + FStar_Compiler_Range.line_of_pos uu___4 in + let uu___4 = + let uu___5 = + FStar_Compiler_Range.start_of_range + FStar_Interactive_Ide_Types.initial_range in + FStar_Compiler_Range.col_of_pos uu___5 in + { + FStar_Parser_ParseIt.frag_fname = uu___2; + FStar_Parser_ParseIt.frag_text = code; + FStar_Parser_ParseIt.frag_line = uu___3; + FStar_Parser_ParseIt.frag_col = uu___4 + } in + FStar_Parser_ParseIt.Incremental uu___1 in + FStar_Parser_ParseIt.parse uu___ in + let log_syntax_issues err = + match err with + | FStar_Pervasives_Native.None -> () + | FStar_Pervasives_Native.Some (raw_error, msg, range) -> + let uu___ = FStar_Errors.lookup raw_error in + (match uu___ with + | (uu___1, uu___2, num) -> + let issue = + { + FStar_Errors.issue_msg = msg; + FStar_Errors.issue_level = FStar_Errors.EError; + FStar_Errors.issue_range = + (FStar_Pervasives_Native.Some range); + FStar_Errors.issue_number = + (FStar_Pervasives_Native.Some num); + FStar_Errors.issue_ctx = [] + } in + write_full_buffer_fragment_progress + (FStar_Pervasives.Inr [issue])) in + let qs = + match parse_result with + | FStar_Parser_ParseIt.IncrementalFragment + (decls, uu___, err_opt) -> + let queries = + let uu___1 = + let uu___2 = + FStar_Compiler_Effect.op_Bang + FStar_Interactive_PushHelper.repl_stack in + inspect_repl_stack uu___2 decls + write_full_buffer_fragment_progress in + run_qst uu___1 qid1 in + (if full then log_syntax_issues err_opt else (); + (let uu___3 = FStar_Options.debug_any () in + if uu___3 + then + let uu___4 = + let uu___5 = + FStar_Compiler_List.map + FStar_Interactive_Ide_Types.query_to_string + queries in + FStar_String.concat "\n" uu___5 in + FStar_Compiler_Util.print1 "Generating queries\n%s\n" + uu___4 + else ()); + if full then queries else []) + | FStar_Parser_ParseIt.ParseError err -> + (if full + then log_syntax_issues (FStar_Pervasives_Native.Some err) + else (); + []) + | uu___ -> failwith "Unexpected parse result" in + qs \ No newline at end of file diff --git a/src/fstar/FStar.Interactive.Ide.Types.fst b/src/fstar/FStar.Interactive.Ide.Types.fst index 7ad07536aeb..cfaa1cd8602 100755 --- a/src/fstar/FStar.Interactive.Ide.Types.fst +++ b/src/fstar/FStar.Interactive.Ide.Types.fst @@ -120,6 +120,8 @@ type push_query = push_peek_only: bool; push_code_or_decl: either string FStar.Parser.AST.decl} +type lookup_symbol_range = json + type query' = | Exit | DescribeProtocol @@ -129,12 +131,12 @@ type query' = | Push of push_query | VfsAdd of option string (* fname *) * string (* contents *) | AutoComplete of string * completion_context -| Lookup of string * lookup_context * option position * list string +| Lookup of string * lookup_context * option position * list string * option lookup_symbol_range | Compute of string * option (list FStar.TypeChecker.Env.step) | Search of string | GenericError of string | ProtocolViolation of string -| FullBuffer of string +| FullBuffer of string & bool //if true, check the buffer, otherwise just find verified prefix and query = { qq: query'; qid: string } let push_query_to_string pq = diff --git a/src/fstar/FStar.Interactive.Ide.fst b/src/fstar/FStar.Interactive.Ide.fst index 619ee24cd02..84b5ad018e1 100644 --- a/src/fstar/FStar.Interactive.Ide.fst +++ b/src/fstar/FStar.Interactive.Ide.fst @@ -213,6 +213,11 @@ let unpack_interactive_query json = | Some JsonNull -> None | other -> other in + let read_position err loc = + assoc err "filename" loc |> js_str, + assoc err "line" loc |> js_int, + assoc err "column" loc |> js_int + in { qid = qid; qq = match query with | "exit" -> Exit @@ -225,18 +230,16 @@ let unpack_interactive_query json = push_line = arg "line" |> js_int; push_column = arg "column" |> js_int; push_peek_only = query = "peek" }) - | "full-buffer" -> FullBuffer (arg "code" |> js_str) + | "full-buffer" -> FullBuffer (arg "code" |> js_str, (arg "kind" |> js_str) = "full") | "autocomplete" -> AutoComplete (arg "partial-symbol" |> js_str, try_arg "context" |> js_optional_completion_context) | "lookup" -> Lookup (arg "symbol" |> js_str, try_arg "context" |> js_optional_lookup_context, try_arg "location" |> Util.map_option js_assoc - |> Util.map_option (fun loc -> - (assoc "[location]" "filename" loc |> js_str, - assoc "[location]" "line" loc |> js_int, - assoc "[location]" "column" loc |> js_int)), - arg "requested-info" |> js_list js_str) + |> Util.map_option (read_position "[location]"), + arg "requested-info" |> js_list js_str, + try_arg "symbol-range") | "compute" -> Compute (arg "term" |> js_str, try_arg "rules" |> Util.map_option (js_list js_reductionrule)) @@ -268,12 +271,14 @@ let read_interactive_query stream : query = let json_of_opt json_of_a opt_a = Util.dflt JsonNull (Util.map_option json_of_a opt_a) -let alist_of_symbol_lookup_result lr = +let alist_of_symbol_lookup_result lr symbol symrange_opt= [("name", JsonStr lr.slr_name); ("defined-at", json_of_opt json_of_def_range lr.slr_def_range); ("type", json_of_opt JsonStr lr.slr_typ); ("documentation", json_of_opt JsonStr lr.slr_doc); - ("definition", json_of_opt JsonStr lr.slr_def)] + ("definition", json_of_opt JsonStr lr.slr_def); + ("symbol-range", json_of_opt (fun x -> x) symrange_opt); + ("symbol", JsonStr symbol)] let alist_of_protocol_info = let js_version = JsonInt interactive_protocol_vernum in @@ -620,10 +625,11 @@ let run_push st query = else run_push_without_deps st query -let run_symbol_lookup st symbol pos_opt requested_info = +let run_symbol_lookup st symbol pos_opt requested_info (symbol_range_opt:option json) = match QH.symlookup st.repl_env symbol pos_opt requested_info with | None -> Inl "Symbol not found" - | Some result -> Inr ("symbol", alist_of_symbol_lookup_result result) + | Some result -> + Inr ("symbol", alist_of_symbol_lookup_result result symbol symbol_range_opt) let run_option_lookup opt_name = let _, trimmed_name = trim_option_name opt_name in @@ -641,22 +647,22 @@ let run_module_lookup st symbol = | Some (CTable.Namespace ns_info) -> Inr ("namespace", CTable.alist_of_ns_info ns_info) -let run_code_lookup st symbol pos_opt requested_info = - match run_symbol_lookup st symbol pos_opt requested_info with +let run_code_lookup st symbol pos_opt requested_info symrange_opt= + match run_symbol_lookup st symbol pos_opt requested_info symrange_opt with | Inr alist -> Inr alist | Inl _ -> match run_module_lookup st symbol with | Inr alist -> Inr alist | Inl err_msg -> Inl "No such symbol, module, or namespace." -let run_lookup' st symbol context pos_opt requested_info = +let run_lookup' st symbol context pos_opt requested_info symrange = match context with - | LKSymbolOnly -> run_symbol_lookup st symbol pos_opt requested_info + | LKSymbolOnly -> run_symbol_lookup st symbol pos_opt requested_info symrange | LKModule -> run_module_lookup st symbol | LKOption -> run_option_lookup symbol - | LKCode -> run_code_lookup st symbol pos_opt requested_info + | LKCode -> run_code_lookup st symbol pos_opt requested_info symrange -let run_lookup st symbol context pos_opt requested_info = - match run_lookup' st symbol context pos_opt requested_info with +let run_lookup st symbol context pos_opt requested_info symrange = + match run_lookup' st symbol context pos_opt requested_info symrange with | Inl err_msg -> ((QueryNOK, JsonStr err_msg), Inl st) | Inr (kind, info) -> @@ -932,15 +938,15 @@ let rec run_query st (q: query) : (query_status * list json) * either repl_state | VfsAdd (fname, contents) -> as_json_list (run_vfs_add st fname contents) | Push pquery -> as_json_list (run_push st pquery) | Pop -> as_json_list (run_pop st) - | FullBuffer code -> + | FullBuffer (code, full) -> let queries = - FStar.Interactive.Incremental.run_full_buffer st q.qid code write_full_buffer_fragment_progress + FStar.Interactive.Incremental.run_full_buffer st q.qid code full write_full_buffer_fragment_progress in fold_query validate_and_run_query queries st [] | AutoComplete (search_term, context) -> as_json_list (run_autocomplete st search_term context) - | Lookup (symbol, context, pos_opt, rq_info) -> - as_json_list (run_lookup st symbol context pos_opt rq_info) + | Lookup (symbol, context, pos_opt, rq_info, symrange) -> + as_json_list (run_lookup st symbol context pos_opt rq_info symrange) | Compute (term, rules) -> as_json_list (run_compute st term rules) | Search term -> diff --git a/src/fstar/FStar.Interactive.Incremental.fst b/src/fstar/FStar.Interactive.Incremental.fst index 130d9992b3c..fdeb26ac1f1 100755 --- a/src/fstar/FStar.Interactive.Incremental.fst +++ b/src/fstar/FStar.Interactive.Incremental.fst @@ -170,6 +170,7 @@ let inspect_repl_stack (s:repl_stack_t) let run_full_buffer (st:repl_state) (qid:string) (code:string) + (full:bool) (write_full_buffer_fragment_progress: either decl (list issue) -> unit) : list query = let parse_result = @@ -200,16 +201,16 @@ let run_full_buffer (st:repl_state) let queries = run_qst (inspect_repl_stack (!repl_stack) decls write_full_buffer_fragment_progress) qid in - log_syntax_issues err_opt; + if full then log_syntax_issues err_opt; if Options.debug_any() then ( BU.print1 "Generating queries\n%s\n" (String.concat "\n" (List.map query_to_string queries)) ); - queries + if full then queries else [] | ParseError err -> - log_syntax_issues (Some err); + if full then log_syntax_issues (Some err); [] | _ -> failwith "Unexpected parse result" diff --git a/src/fstar/FStar.Interactive.Incremental.fsti b/src/fstar/FStar.Interactive.Incremental.fsti index c2ee8263a38..be6230862c1 100755 --- a/src/fstar/FStar.Interactive.Incremental.fsti +++ b/src/fstar/FStar.Interactive.Incremental.fsti @@ -38,5 +38,6 @@ open FStar.Interactive.Ide.Types val run_full_buffer (st:repl_state) (qid:string) (code:string) + (full:bool) (write_full_buffer_fragment_progress: either decl (list issue) -> unit) : list query From a8739fb2f235e6647d8dd5b3f09d0b4669af5d91 Mon Sep 17 00:00:00 2001 From: Nikhil Swamy Date: Mon, 13 Mar 2023 10:18:56 -0700 Subject: [PATCH 09/48] reset gensym in incremental parsing; implement parser decl equality of constants; add a config file for the F* compiler for use in vscode --- ocaml/fstar-lib/FStar_Parser_ParseIt.ml | 1 + .../generated/FStar_Parser_AST_Comparison.ml | 18 +++++++++++++++++- src/FStarCompiler.fst.config.json | 3 +++ src/parser/FStar.Parser.AST.Comparison.fst | 15 +++++++++++++-- 4 files changed, 34 insertions(+), 3 deletions(-) create mode 100644 src/FStarCompiler.fst.config.json diff --git a/ocaml/fstar-lib/FStar_Parser_ParseIt.ml b/ocaml/fstar-lib/FStar_Parser_ParseIt.ml index e627e9fb297..2d5badc2da2 100644 --- a/ocaml/fstar-lib/FStar_Parser_ParseIt.ml +++ b/ocaml/fstar-lib/FStar_Parser_ParseIt.ml @@ -138,6 +138,7 @@ let parse fn = let _, _, r = err_of_parse_error () in let d = try + FStar_Ident.reset_gensym(); Inl (parse_one_decl lexer) with | FStar_Errors.Error(e, msg, r, _ctx) -> diff --git a/ocaml/fstar-lib/generated/FStar_Parser_AST_Comparison.ml b/ocaml/fstar-lib/generated/FStar_Parser_AST_Comparison.ml index d22e09c9edc..43ddb6bcd56 100755 --- a/ocaml/fstar-lib/generated/FStar_Parser_AST_Comparison.ml +++ b/ocaml/fstar-lib/generated/FStar_Parser_AST_Comparison.ml @@ -26,7 +26,23 @@ let eq_option : t21) -> f t11 t21 | uu___ -> false let (eq_sconst : FStar_Const.sconst -> FStar_Const.sconst -> Prims.bool) = - fun c1 -> fun c2 -> false + fun c1 -> + fun c2 -> + match (c1, c2) with + | (FStar_Const.Const_effect, FStar_Const.Const_effect) -> true + | (FStar_Const.Const_unit, FStar_Const.Const_unit) -> true + | (FStar_Const.Const_bool b1, FStar_Const.Const_bool b2) -> b1 = b2 + | (FStar_Const.Const_int (s1, sw1), FStar_Const.Const_int (s2, sw2)) -> + (s1 = s2) && (sw1 = sw2) + | (FStar_Const.Const_char c11, FStar_Const.Const_char c21) -> c11 = c21 + | (FStar_Const.Const_string (s1, uu___), FStar_Const.Const_string + (s2, uu___1)) -> s1 = s2 + | (FStar_Const.Const_real s1, FStar_Const.Const_real s2) -> s1 = s2 + | (FStar_Const.Const_range r1, FStar_Const.Const_range r2) -> r1 = r2 + | (FStar_Const.Const_reify, FStar_Const.Const_reify) -> true + | (FStar_Const.Const_reflect l1, FStar_Const.Const_reflect l2) -> + FStar_Ident.lid_equals l1 l2 + | uu___ -> false let rec (eq_term : FStar_Parser_AST.term -> FStar_Parser_AST.term -> Prims.bool) = fun t1 -> fun t2 -> eq_term' t1.FStar_Parser_AST.tm t2.FStar_Parser_AST.tm diff --git a/src/FStarCompiler.fst.config.json b/src/FStarCompiler.fst.config.json new file mode 100644 index 00000000000..5a062556de9 --- /dev/null +++ b/src/FStarCompiler.fst.config.json @@ -0,0 +1,3 @@ +{ "fstar_exe":"fstar.exe", + "options":["--MLish", "--lax", "--cache_dir", ".cache.boot", "--no_location_info", "--warn_error", "-271-272-241-319-274"], + "include_dirs":["../ulib", "basic", "basic/boot", "extraction", "fstar", "parser", "prettyprint", "prettyprint/boot", "reflection", "smtencoding", "syntax", "tactics", "tosyntax", "typechecker", "tests", "tests/boot"] } diff --git a/src/parser/FStar.Parser.AST.Comparison.fst b/src/parser/FStar.Parser.AST.Comparison.fst index 4e0e636fe84..21a2def23b3 100755 --- a/src/parser/FStar.Parser.AST.Comparison.fst +++ b/src/parser/FStar.Parser.AST.Comparison.fst @@ -37,7 +37,6 @@ let eq_list (f: 'a -> 'a -> bool) (t1 t2:list 'a) = List.length t1 = List.length t2 && List.forall2 f t1 t2 -// eq_opt: A equality on option a let eq_option (f: 'a -> 'a -> bool) (t1 t2:option 'a) : bool = match t1, t2 with @@ -45,7 +44,19 @@ let eq_option (f: 'a -> 'a -> bool) (t1 t2:option 'a) | Some t1, Some t2 -> f t1 t2 | _ -> false -let eq_sconst (c1 c2: sconst) : bool = false +let eq_sconst (c1 c2: sconst) : bool = + match c1, c2 with + | Const_effect, Const_effect -> true + | Const_unit, Const_unit -> true + | Const_bool b1, Const_bool b2 -> b1 = b2 + | Const_int (s1, sw1), Const_int (s2, sw2) -> s1=s2 && sw1=sw2 + | Const_char c1, Const_char c2 -> c1 = c2 + | Const_string (s1, _), Const_string (s2, _) -> s1 = s2 + | Const_real s1, Const_real s2 -> s1 = s2 + | Const_range r1, Const_range r2 -> r1 = r2 + | Const_reify, Const_reify -> true + | Const_reflect l1, Const_reflect l2 -> Ident.lid_equals l1 l2 + | _ -> false let rec eq_term (t1 t2:term) : bool From 6086ec46ba7f3f8f9085071a229efbe6e904b82d Mon Sep 17 00:00:00 2001 From: Nikhil Swamy Date: Mon, 13 Mar 2023 11:23:56 -0700 Subject: [PATCH 10/48] only echo back symbol-range in symbol lookups rather than including null fields (emacs compat) --- .../generated/FStar_Interactive_Ide.ml | 67 ++++++++++--------- src/fstar/FStar.Interactive.Ide.fst | 12 +++- 2 files changed, 45 insertions(+), 34 deletions(-) diff --git a/ocaml/fstar-lib/generated/FStar_Interactive_Ide.ml b/ocaml/fstar-lib/generated/FStar_Interactive_Ide.ml index db4c490b5de..0f8a0b016c9 100644 --- a/ocaml/fstar-lib/generated/FStar_Interactive_Ide.ml +++ b/ocaml/fstar-lib/generated/FStar_Interactive_Ide.ml @@ -565,43 +565,48 @@ let (alist_of_symbol_lookup_result : let uu___ = let uu___1 = let uu___2 = - json_of_opt FStar_Compiler_Range.json_of_def_range - lr.FStar_Interactive_QueryHelper.slr_def_range in - ("defined-at", uu___2) in - let uu___2 = + let uu___3 = + json_of_opt FStar_Compiler_Range.json_of_def_range + lr.FStar_Interactive_QueryHelper.slr_def_range in + ("defined-at", uu___3) in let uu___3 = let uu___4 = - json_of_opt - (fun uu___5 -> FStar_Compiler_Util.JsonStr uu___5) - lr.FStar_Interactive_QueryHelper.slr_typ in - ("type", uu___4) in - let uu___4 = + let uu___5 = + json_of_opt + (fun uu___6 -> FStar_Compiler_Util.JsonStr uu___6) + lr.FStar_Interactive_QueryHelper.slr_typ in + ("type", uu___5) in let uu___5 = let uu___6 = - json_of_opt - (fun uu___7 -> FStar_Compiler_Util.JsonStr uu___7) - lr.FStar_Interactive_QueryHelper.slr_doc in - ("documentation", uu___6) in - let uu___6 = + let uu___7 = + json_of_opt + (fun uu___8 -> FStar_Compiler_Util.JsonStr uu___8) + lr.FStar_Interactive_QueryHelper.slr_doc in + ("documentation", uu___7) in let uu___7 = let uu___8 = - json_of_opt - (fun uu___9 -> FStar_Compiler_Util.JsonStr uu___9) - lr.FStar_Interactive_QueryHelper.slr_def in - ("definition", uu___8) in - let uu___8 = - let uu___9 = - let uu___10 = json_of_opt (fun x -> x) symrange_opt in - ("symbol-range", uu___10) in - [uu___9; ("symbol", (FStar_Compiler_Util.JsonStr symbol))] in - uu___7 :: uu___8 in - uu___5 :: uu___6 in - uu___3 :: uu___4 in - uu___1 :: uu___2 in - ("name", - (FStar_Compiler_Util.JsonStr - (lr.FStar_Interactive_QueryHelper.slr_name))) - :: uu___ + let uu___9 = + json_of_opt + (fun uu___10 -> FStar_Compiler_Util.JsonStr uu___10) + lr.FStar_Interactive_QueryHelper.slr_def in + ("definition", uu___9) in + [uu___8] in + uu___6 :: uu___7 in + uu___4 :: uu___5 in + uu___2 :: uu___3 in + ("name", + (FStar_Compiler_Util.JsonStr + (lr.FStar_Interactive_QueryHelper.slr_name))) + :: uu___1 in + let uu___1 = + match symrange_opt with + | FStar_Pervasives_Native.None -> [] + | FStar_Pervasives_Native.Some symrange -> + let uu___2 = + let uu___3 = json_of_opt (fun x -> x) symrange_opt in + ("symbol-range", uu___3) in + [uu___2; ("symbol", (FStar_Compiler_Util.JsonStr symbol))] in + FStar_Compiler_List.op_At uu___ uu___1 let (alist_of_protocol_info : (Prims.string * FStar_Compiler_Util.json) Prims.list) = let js_version = diff --git a/src/fstar/FStar.Interactive.Ide.fst b/src/fstar/FStar.Interactive.Ide.fst index 84b5ad018e1..33534c0cf5b 100644 --- a/src/fstar/FStar.Interactive.Ide.fst +++ b/src/fstar/FStar.Interactive.Ide.fst @@ -276,9 +276,15 @@ let alist_of_symbol_lookup_result lr symbol symrange_opt= ("defined-at", json_of_opt json_of_def_range lr.slr_def_range); ("type", json_of_opt JsonStr lr.slr_typ); ("documentation", json_of_opt JsonStr lr.slr_doc); - ("definition", json_of_opt JsonStr lr.slr_def); - ("symbol-range", json_of_opt (fun x -> x) symrange_opt); - ("symbol", JsonStr symbol)] + ("definition", json_of_opt JsonStr lr.slr_def)] @ ( + // echo back the symbol-range and symbol, if symbol-range was provided + // (don't include it otherwise, for backwards compat with fstar-mode.el) + match symrange_opt with + | None -> [] + | Some symrange -> + [("symbol-range", json_of_opt (fun x -> x) symrange_opt); + ("symbol", JsonStr symbol)] + ) let alist_of_protocol_info = let js_version = JsonInt interactive_protocol_vernum in From cf1ef5a632a60ddbfd688ce7218454c5b47410c6 Mon Sep 17 00:00:00 2001 From: Nikhil Swamy Date: Mon, 13 Mar 2023 20:54:26 +0000 Subject: [PATCH 11/48] fix file permissions in generated snapshot --- ocaml/fstar-lib/generated/FStar_Interactive_Ide_Types.ml | 0 ocaml/fstar-lib/generated/FStar_Interactive_Incremental.ml | 0 ocaml/fstar-lib/generated/FStar_Interactive_ReplState.ml | 2 +- ocaml/fstar-lib/generated/FStar_Parser_AST_Comparison.ml | 0 4 files changed, 1 insertion(+), 1 deletion(-) mode change 100755 => 100644 ocaml/fstar-lib/generated/FStar_Interactive_Ide_Types.ml mode change 100755 => 100644 ocaml/fstar-lib/generated/FStar_Interactive_Incremental.ml mode change 100755 => 100644 ocaml/fstar-lib/generated/FStar_Interactive_ReplState.ml mode change 100755 => 100644 ocaml/fstar-lib/generated/FStar_Parser_AST_Comparison.ml diff --git a/ocaml/fstar-lib/generated/FStar_Interactive_Ide_Types.ml b/ocaml/fstar-lib/generated/FStar_Interactive_Ide_Types.ml old mode 100755 new mode 100644 diff --git a/ocaml/fstar-lib/generated/FStar_Interactive_Incremental.ml b/ocaml/fstar-lib/generated/FStar_Interactive_Incremental.ml old mode 100755 new mode 100644 diff --git a/ocaml/fstar-lib/generated/FStar_Interactive_ReplState.ml b/ocaml/fstar-lib/generated/FStar_Interactive_ReplState.ml old mode 100755 new mode 100644 index 1e318053bf6..0a505ac7e74 --- a/ocaml/fstar-lib/generated/FStar_Interactive_ReplState.ml +++ b/ocaml/fstar-lib/generated/FStar_Interactive_ReplState.ml @@ -180,5 +180,5 @@ let (repl_state_to_string : repl_state -> Prims.string) = uu___3 :: uu___4 in uu___1 :: uu___2 in FStar_Compiler_Util.format - "{\n\trepl_line=%s;\n\trepl_column=%s;\n\trepl_fname=%s;\n\trepl_cur_mod=%s;\n\t\\ \r\n repl_deps_stack={%s}\n}" + "{\n\trepl_line=%s;\n\trepl_column=%s;\n\trepl_fname=%s;\n\trepl_cur_mod=%s;\n\t\\ \n repl_deps_stack={%s}\n}" uu___ \ No newline at end of file diff --git a/ocaml/fstar-lib/generated/FStar_Parser_AST_Comparison.ml b/ocaml/fstar-lib/generated/FStar_Parser_AST_Comparison.ml old mode 100755 new mode 100644 From fcf8e512c147387c01475b06350d4906d3a0c4c9 Mon Sep 17 00:00:00 2001 From: Nikhil Swamy Date: Mon, 13 Mar 2023 20:19:56 -0700 Subject: [PATCH 12/48] a Callback query entry to enable instrumenting the repl with progress markers --- .../generated/FStar_Interactive_Ide.ml | 21 +++- .../generated/FStar_Interactive_Ide_Types.ml | 35 ++++-- .../FStar_Interactive_Incremental.ml | 107 ++++++++++++------ src/fstar/FStar.Interactive.Ide.Types.fst | 16 ++- src/fstar/FStar.Interactive.Ide.fst | 14 ++- src/fstar/FStar.Interactive.Incremental.fst | 28 +++-- src/fstar/FStar.Interactive.Incremental.fsti | 7 +- 7 files changed, 157 insertions(+), 71 deletions(-) diff --git a/ocaml/fstar-lib/generated/FStar_Interactive_Ide.ml b/ocaml/fstar-lib/generated/FStar_Interactive_Ide.ml index 0f8a0b016c9..914234c0225 100644 --- a/ocaml/fstar-lib/generated/FStar_Interactive_Ide.ml +++ b/ocaml/fstar-lib/generated/FStar_Interactive_Ide.ml @@ -1237,12 +1237,20 @@ let (rephrase_dependency_error : FStar_Errors.issue -> FStar_Errors.issue) = FStar_Errors.issue_ctx = (issue.FStar_Errors.issue_ctx) } let (write_full_buffer_fragment_progress : - (FStar_Parser_AST.decl, FStar_Errors.issue Prims.list) - FStar_Pervasives.either -> unit) - = + FStar_Interactive_Incremental.fragment_progress -> unit) = fun di -> match di with - | FStar_Pervasives.Inl d -> + | FStar_Interactive_Incremental.FragmentStarted d -> + let uu___ = + let uu___1 = + let uu___2 = + FStar_Compiler_Range.json_of_def_range + d.FStar_Parser_AST.drange in + ("ranges", uu___2) in + [uu___1] in + write_progress + (FStar_Pervasives_Native.Some "full-buffer-fragment-started") uu___ + | FStar_Interactive_Incremental.FragmentSuccess d -> let uu___ = let uu___1 = let uu___2 = @@ -1252,7 +1260,7 @@ let (write_full_buffer_fragment_progress : [uu___1] in write_progress (FStar_Pervasives_Native.Some "full-buffer-fragment-ok") uu___ - | FStar_Pervasives.Inr issues -> + | FStar_Interactive_Incremental.FragmentError issues -> let qid = let uu___ = FStar_Compiler_Effect.op_Bang repl_current_qid in match uu___ with @@ -1442,7 +1450,7 @@ let (run_push_without_deps : | FStar_Pervasives.Inr d when Prims.op_Negation has_error -> write_full_buffer_fragment_progress - (FStar_Pervasives.Inl d) + (FStar_Interactive_Incremental.FragmentSuccess d) | uu___4 -> ()); (let json_errors = let uu___4 = @@ -2391,6 +2399,7 @@ let rec (run_query : let uu___ = run_compute st term rules in as_json_list uu___ | FStar_Interactive_Ide_Types.Search term -> let uu___ = run_search st term in as_json_list uu___ + | FStar_Interactive_Ide_Types.Callback f -> f st and (validate_and_run_query : FStar_Interactive_ReplState.repl_state -> FStar_Interactive_Ide_Types.query -> run_query_result) diff --git a/ocaml/fstar-lib/generated/FStar_Interactive_Ide_Types.ml b/ocaml/fstar-lib/generated/FStar_Interactive_Ide_Types.ml index e81383bb94a..02b78d7e366 100644 --- a/ocaml/fstar-lib/generated/FStar_Interactive_Ide_Types.ml +++ b/ocaml/fstar-lib/generated/FStar_Interactive_Ide_Types.ml @@ -137,6 +137,22 @@ let (__proj__Mkpush_query__item__push_code_or_decl : | { push_kind; push_line; push_column; push_peek_only; push_code_or_decl;_} -> push_code_or_decl type lookup_symbol_range = FStar_Compiler_Util.json +type query_status = + | QueryOK + | QueryNOK + | QueryViolatesProtocol +let (uu___is_QueryOK : query_status -> Prims.bool) = + fun projectee -> match projectee with | QueryOK -> true | uu___ -> false +let (uu___is_QueryNOK : query_status -> Prims.bool) = + fun projectee -> match projectee with | QueryNOK -> true | uu___ -> false +let (uu___is_QueryViolatesProtocol : query_status -> Prims.bool) = + fun projectee -> + match projectee with | QueryViolatesProtocol -> true | uu___ -> false +type callback_t = + FStar_Interactive_ReplState.repl_state -> + ((query_status * FStar_Compiler_Util.json Prims.list) * + (FStar_Interactive_ReplState.repl_state, Prims.int) + FStar_Pervasives.either) type query' = | Exit | DescribeProtocol @@ -155,6 +171,7 @@ type query' = | GenericError of Prims.string | ProtocolViolation of Prims.string | FullBuffer of (Prims.string * Prims.bool) + | Callback of callback_t and query = { qq: query' ; qid: Prims.string } @@ -221,6 +238,11 @@ let (uu___is_FullBuffer : query' -> Prims.bool) = match projectee with | FullBuffer _0 -> true | uu___ -> false let (__proj__FullBuffer__item___0 : query' -> (Prims.string * Prims.bool)) = fun projectee -> match projectee with | FullBuffer _0 -> _0 +let (uu___is_Callback : query' -> Prims.bool) = + fun projectee -> + match projectee with | Callback _0 -> true | uu___ -> false +let (__proj__Callback__item___0 : query' -> callback_t) = + fun projectee -> match projectee with | Callback _0 -> _0 let (__proj__Mkquery__item__qq : query -> query') = fun projectee -> match projectee with | { qq; qid;_} -> qq let (__proj__Mkquery__item__qid : query -> Prims.string) = @@ -270,6 +292,7 @@ let (query_to_string : query -> Prims.string) = | GenericError uu___ -> "GenericError" | ProtocolViolation uu___ -> "ProtocolViolation" | FullBuffer uu___ -> "FullBuffer" + | Callback uu___ -> "Callback" let (query_needs_current_module : query' -> Prims.bool) = fun uu___ -> match uu___ with @@ -286,6 +309,7 @@ let (query_needs_current_module : query' -> Prims.bool) = | GenericError uu___1 -> false | ProtocolViolation uu___1 -> false | FullBuffer uu___1 -> false + | Callback uu___1 -> false | Push uu___1 -> true | AutoComplete uu___1 -> true | Lookup uu___1 -> true @@ -315,17 +339,6 @@ let (interactive_protocol_features : Prims.string Prims.list) = "interrupt"; "progress"; "full-buffer"] -type query_status = - | QueryOK - | QueryNOK - | QueryViolatesProtocol -let (uu___is_QueryOK : query_status -> Prims.bool) = - fun projectee -> match projectee with | QueryOK -> true | uu___ -> false -let (uu___is_QueryNOK : query_status -> Prims.bool) = - fun projectee -> match projectee with | QueryNOK -> true | uu___ -> false -let (uu___is_QueryViolatesProtocol : query_status -> Prims.bool) = - fun projectee -> - match projectee with | QueryViolatesProtocol -> true | uu___ -> false let (json_of_issue_level : FStar_Errors.issue_level -> FStar_Compiler_Util.json) = fun i -> diff --git a/ocaml/fstar-lib/generated/FStar_Interactive_Incremental.ml b/ocaml/fstar-lib/generated/FStar_Interactive_Incremental.ml index cfc4c8d8a4b..d9e9cf53bf4 100644 --- a/ocaml/fstar-lib/generated/FStar_Interactive_Incremental.ml +++ b/ocaml/fstar-lib/generated/FStar_Interactive_Incremental.ml @@ -1,4 +1,26 @@ open Prims +type fragment_progress = + | FragmentStarted of FStar_Parser_AST.decl + | FragmentSuccess of FStar_Parser_AST.decl + | FragmentError of FStar_Errors.issue Prims.list +let (uu___is_FragmentStarted : fragment_progress -> Prims.bool) = + fun projectee -> + match projectee with | FragmentStarted _0 -> true | uu___ -> false +let (__proj__FragmentStarted__item___0 : + fragment_progress -> FStar_Parser_AST.decl) = + fun projectee -> match projectee with | FragmentStarted _0 -> _0 +let (uu___is_FragmentSuccess : fragment_progress -> Prims.bool) = + fun projectee -> + match projectee with | FragmentSuccess _0 -> true | uu___ -> false +let (__proj__FragmentSuccess__item___0 : + fragment_progress -> FStar_Parser_AST.decl) = + fun projectee -> match projectee with | FragmentSuccess _0 -> _0 +let (uu___is_FragmentError : fragment_progress -> Prims.bool) = + fun projectee -> + match projectee with | FragmentError _0 -> true | uu___ -> false +let (__proj__FragmentError__item___0 : + fragment_progress -> FStar_Errors.issue Prims.list) = + fun projectee -> match projectee with | FragmentError _0 -> _0 type qid = (Prims.string * Prims.int) type 'a qst = qid -> ('a * qid) let return : 'a . 'a -> 'a qst = fun x -> fun q -> (x, q) @@ -51,31 +73,47 @@ let (as_query : } in return uu___1) let (push_decl : - FStar_Parser_AST.decl -> FStar_Interactive_Ide_Types.query qst) = - fun d -> - let pq = - let uu___ = + (fragment_progress -> unit) -> + FStar_Parser_AST.decl -> FStar_Interactive_Ide_Types.query Prims.list qst) + = + fun write_full_buffer_fragment_progress -> + fun d -> + let pq = + let uu___ = + let uu___1 = + FStar_Compiler_Range.start_of_range d.FStar_Parser_AST.drange in + FStar_Compiler_Range.line_of_pos uu___1 in let uu___1 = - FStar_Compiler_Range.start_of_range d.FStar_Parser_AST.drange in - FStar_Compiler_Range.line_of_pos uu___1 in - let uu___1 = - let uu___2 = - FStar_Compiler_Range.start_of_range d.FStar_Parser_AST.drange in - FStar_Compiler_Range.col_of_pos uu___2 in - { - FStar_Interactive_Ide_Types.push_kind = - FStar_Interactive_PushHelper.FullCheck; - FStar_Interactive_Ide_Types.push_line = uu___; - FStar_Interactive_Ide_Types.push_column = uu___1; - FStar_Interactive_Ide_Types.push_peek_only = false; - FStar_Interactive_Ide_Types.push_code_or_decl = - (FStar_Pervasives.Inr d) - } in - as_query (FStar_Interactive_Ide_Types.Push pq) + let uu___2 = + FStar_Compiler_Range.start_of_range d.FStar_Parser_AST.drange in + FStar_Compiler_Range.col_of_pos uu___2 in + { + FStar_Interactive_Ide_Types.push_kind = + FStar_Interactive_PushHelper.FullCheck; + FStar_Interactive_Ide_Types.push_line = uu___; + FStar_Interactive_Ide_Types.push_column = uu___1; + FStar_Interactive_Ide_Types.push_peek_only = false; + FStar_Interactive_Ide_Types.push_code_or_decl = + (FStar_Pervasives.Inr d) + } in + let progress st = + write_full_buffer_fragment_progress (FragmentStarted d); + ((FStar_Interactive_Ide_Types.QueryOK, []), + (FStar_Pervasives.Inl st)) in + let uu___ = as_query (FStar_Interactive_Ide_Types.Callback progress) in + op_let_Bang uu___ + (fun cb -> + let uu___1 = as_query (FStar_Interactive_Ide_Types.Push pq) in + op_let_Bang uu___1 (fun push -> return [cb; push])) let (push_decls : - FStar_Parser_AST.decl Prims.list -> - FStar_Interactive_Ide_Types.query Prims.list qst) - = fun ds -> map push_decl ds + (fragment_progress -> unit) -> + FStar_Parser_AST.decl Prims.list -> + FStar_Interactive_Ide_Types.query Prims.list qst) + = + fun write_full_buffer_fragment_progress -> + fun ds -> + let uu___ = map (push_decl write_full_buffer_fragment_progress) ds in + op_let_Bang uu___ (fun qs -> return (FStar_Compiler_List.flatten qs)) let (pop_entries : FStar_Interactive_ReplState.repl_stack_entry_t Prims.list -> FStar_Interactive_Ide_Types.query Prims.list qst) @@ -105,14 +143,14 @@ let (response_success : let (inspect_repl_stack : FStar_Interactive_ReplState.repl_stack_t -> FStar_Parser_AST.decl Prims.list -> - ((FStar_Parser_AST.decl, FStar_Errors.issue Prims.list) - FStar_Pervasives.either -> unit) - -> FStar_Interactive_Ide_Types.query Prims.list qst) + (fragment_progress -> unit) -> + FStar_Interactive_Ide_Types.query Prims.list qst) = fun s -> fun ds -> fun write_full_buffer_fragment_progress -> let entries = FStar_Compiler_List.rev s in + let push_decls1 = push_decls write_full_buffer_fragment_progress in let uu___ = FStar_Compiler_Util.prefix_until (fun uu___1 -> @@ -123,7 +161,7 @@ let (inspect_repl_stack : | uu___2 -> false) entries in match uu___ with | FStar_Pervasives_Native.None -> - let uu___1 = push_decls ds in + let uu___1 = push_decls1 ds in op_let_Bang uu___1 (fun ds1 -> return ds1) | FStar_Pervasives_Native.Some (prefix, first_push, rest) -> let entries1 = first_push :: rest in @@ -141,7 +179,7 @@ let (inspect_repl_stack : let uu___1 = pop_entries (e :: entries3) in op_let_Bang uu___1 (fun pops -> - let uu___2 = push_decls (d :: ds2) in + let uu___2 = push_decls1 (d :: ds2) in op_let_Bang uu___2 (fun pushes -> return @@ -152,19 +190,19 @@ let (inspect_repl_stack : if uu___1 then (write_full_buffer_fragment_progress - (FStar_Pervasives.Inl d); + (FragmentSuccess d); matching_prefix entries3 ds2) else (let uu___3 = pop_entries (e :: entries3) in op_let_Bang uu___3 (fun pops -> - let uu___4 = push_decls (d :: ds2) in + let uu___4 = push_decls1 (d :: ds2) in op_let_Bang uu___4 (fun pushes -> return (FStar_Compiler_List.op_At pops pushes))))) | ([], ds2) -> - let uu___1 = push_decls ds2 in + let uu___1 = push_decls1 ds2 in op_let_Bang uu___1 (fun pushes -> return pushes) | (es, []) -> let uu___1 = pop_entries es in @@ -175,9 +213,8 @@ let (run_full_buffer : Prims.string -> Prims.string -> Prims.bool -> - ((FStar_Parser_AST.decl, FStar_Errors.issue Prims.list) - FStar_Pervasives.either -> unit) - -> FStar_Interactive_Ide_Types.query Prims.list) + (fragment_progress -> unit) -> + FStar_Interactive_Ide_Types.query Prims.list) = fun st -> fun qid1 -> @@ -226,7 +263,7 @@ let (run_full_buffer : FStar_Errors.issue_ctx = [] } in write_full_buffer_fragment_progress - (FStar_Pervasives.Inr [issue])) in + (FragmentError [issue])) in let qs = match parse_result with | FStar_Parser_ParseIt.IncrementalFragment diff --git a/src/fstar/FStar.Interactive.Ide.Types.fst b/src/fstar/FStar.Interactive.Ide.Types.fst index cfaa1cd8602..a8ecc11ea02 100755 --- a/src/fstar/FStar.Interactive.Ide.Types.fst +++ b/src/fstar/FStar.Interactive.Ide.Types.fst @@ -41,7 +41,11 @@ module TcErr = FStar.TypeChecker.Err module TcEnv = FStar.TypeChecker.Env module CTable = FStar.Interactive.CompletionTable module QH = FStar.Interactive.QueryHelper +module RS = FStar.Interactive.ReplState +(***********************) +(* Global state setup *) +(***********************) let initial_range = Range.mk_range "" (Range.mk_pos 1 0) (Range.mk_pos 1 0) @@ -122,6 +126,10 @@ type push_query = type lookup_symbol_range = json +type query_status = | QueryOK | QueryNOK | QueryViolatesProtocol + +type callback_t = RS.repl_state -> (query_status * list json) * either RS.repl_state int + type query' = | Exit | DescribeProtocol @@ -137,6 +145,7 @@ type query' = | GenericError of string | ProtocolViolation of string | FullBuffer of string & bool //if true, check the buffer, otherwise just find verified prefix +| Callback of callback_t and query = { qq: query'; qid: string } let push_query_to_string pq = @@ -173,13 +182,13 @@ let query_to_string q = match q.qq with | GenericError _ -> "GenericError" | ProtocolViolation _ -> "ProtocolViolation" | FullBuffer _ -> "FullBuffer" - +| Callback _ -> "Callback" let query_needs_current_module = function | Exit | DescribeProtocol | DescribeRepl | Segment _ | Pop | Push { push_peek_only = false } | VfsAdd _ - | GenericError _ | ProtocolViolation _ | FullBuffer _ -> false + | GenericError _ | ProtocolViolation _ | FullBuffer _ | Callback _ -> false | Push _ | AutoComplete _ | Lookup _ | Compute _ | Search _ -> true let interactive_protocol_vernum = 2 @@ -192,9 +201,6 @@ let interactive_protocol_features = "peek"; "pop"; "push"; "search"; "segment"; "vfs-add"; "tactic-ranges"; "interrupt"; "progress"; "full-buffer"] -type query_status = | QueryOK | QueryNOK | QueryViolatesProtocol - - let json_of_issue_level i = JsonStr (match i with | ENotImplemented -> "not-implemented" diff --git a/src/fstar/FStar.Interactive.Ide.fst b/src/fstar/FStar.Interactive.Ide.fst index 33534c0cf5b..0e202292bdd 100644 --- a/src/fstar/FStar.Interactive.Ide.fst +++ b/src/fstar/FStar.Interactive.Ide.fst @@ -549,12 +549,16 @@ let rephrase_dependency_error issue = format1 "Error while computing or loading dependencies:\n%s" issue.issue_msg } -let write_full_buffer_fragment_progress (di:either FStar.Parser.AST.decl (list issue)) = +let write_full_buffer_fragment_progress (di:Incremental.fragment_progress) = + let open FStar.Interactive.Incremental in match di with - | Inl d -> + | FragmentStarted d -> + write_progress (Some "full-buffer-fragment-started") + ["ranges", json_of_def_range d.FStar.Parser.AST.drange] + | FragmentSuccess d -> write_progress (Some "full-buffer-fragment-ok") ["ranges", json_of_def_range d.FStar.Parser.AST.drange] - | Inr issues -> + | FragmentError issues -> let qid = match !repl_current_qid with | None -> "unknown" @@ -604,7 +608,7 @@ let run_push_without_deps st query let _ = match code_or_decl with | Inr d when not has_error -> - write_full_buffer_fragment_progress (Inl d) + write_full_buffer_fragment_progress (Incremental.FragmentSuccess d) | _ -> () in let json_errors = JsonList (errs |> List.map json_of_issue) in @@ -957,6 +961,8 @@ let rec run_query st (q: query) : (query_status * list json) * either repl_state as_json_list (run_compute st term rules) | Search term -> as_json_list (run_search st term) + | Callback f -> + f st and validate_and_run_query st query = let query = validate_query st query in repl_current_qid := Some query.qid; diff --git a/src/fstar/FStar.Interactive.Incremental.fst b/src/fstar/FStar.Interactive.Incremental.fst index fdeb26ac1f1..7ad32d77b40 100755 --- a/src/fstar/FStar.Interactive.Incremental.fst +++ b/src/fstar/FStar.Interactive.Incremental.fst @@ -87,8 +87,9 @@ let as_query (q:query') qid=qid_prefix ^ "." ^ string_of_int i } -let push_decl (d:decl) - : qst query +let push_decl (write_full_buffer_fragment_progress: fragment_progress -> unit) + (d:decl) + : qst (list query) = let open FStar.Compiler.Range in let pq = { push_kind = FullCheck; @@ -97,11 +98,19 @@ let push_decl (d:decl) push_peek_only = false; push_code_or_decl = Inr d } in - as_query (Push pq) + let progress st = + write_full_buffer_fragment_progress (FragmentStarted d); + (QueryOK, []), Inl st + in + let! cb = as_query (Callback progress) in + let! push = as_query (Push pq) in + return [cb; push] -let push_decls (ds:list decl) +let push_decls (write_full_buffer_fragment_progress : fragment_progress -> unit) + (ds:list decl) : qst (list query) - = map push_decl ds + = let! qs = map (push_decl write_full_buffer_fragment_progress) ds in + return (List.flatten qs) let pop_entries (e:list repl_stack_entry_t) : qst (list query) @@ -118,9 +127,10 @@ let response_success (d:decl) let inspect_repl_stack (s:repl_stack_t) (ds:list decl) - (write_full_buffer_fragment_progress: either decl (list issue) -> unit) + (write_full_buffer_fragment_progress: fragment_progress -> unit) : qst (list query) // & list json) = let entries = List.rev s in + let push_decls = push_decls write_full_buffer_fragment_progress in match BU.prefix_until (function (_, (PushFragment _, _)) -> true | _ -> false) entries @@ -149,7 +159,7 @@ let inspect_repl_stack (s:repl_stack_t) | PushFragment (Inr d') -> if eq_decl d d' then ( - write_full_buffer_fragment_progress (Inl d); + write_full_buffer_fragment_progress (FragmentSuccess d); matching_prefix entries ds ) else let! pops = pop_entries (e::entries) in @@ -171,7 +181,7 @@ let run_full_buffer (st:repl_state) (qid:string) (code:string) (full:bool) - (write_full_buffer_fragment_progress: either decl (list issue) -> unit) + (write_full_buffer_fragment_progress: fragment_progress -> unit) : list query = let parse_result = P.parse (Incremental { @@ -193,7 +203,7 @@ let run_full_buffer (st:repl_state) issue_number = Some num; issue_ctx = [] } in - write_full_buffer_fragment_progress (Inr [issue]) + write_full_buffer_fragment_progress (FragmentError [issue]) in let qs = match parse_result with diff --git a/src/fstar/FStar.Interactive.Incremental.fsti b/src/fstar/FStar.Interactive.Incremental.fsti index be6230862c1..3d0e4afb3b3 100755 --- a/src/fstar/FStar.Interactive.Incremental.fsti +++ b/src/fstar/FStar.Interactive.Incremental.fsti @@ -25,6 +25,11 @@ open FStar.Parser.AST open FStar.Errors open FStar.Interactive.Ide.Types +type fragment_progress = + | FragmentStarted of decl + | FragmentSuccess of decl + | FragmentError of list issue + (* Translates a full-buffer(qid, code) query by 1. Parsing the code into its declarations 2. Finding a prefix of the repl state that matches a prefix of the declarations @@ -39,5 +44,5 @@ val run_full_buffer (st:repl_state) (qid:string) (code:string) (full:bool) - (write_full_buffer_fragment_progress: either decl (list issue) -> unit) + (write_full_buffer_fragment_progress: fragment_progress -> unit) : list query From cc4ef0b64a65debf10ebad3319088c0079f51cf3 Mon Sep 17 00:00:00 2001 From: Nikhil Swamy Date: Mon, 13 Mar 2023 22:28:46 -0700 Subject: [PATCH 13/48] report failure messages on full-buffer fragments too --- .../generated/FStar_Interactive_Ide.ml | 22 +++++++++++++++---- .../FStar_Interactive_Incremental.ml | 7 ++++++ src/fstar/FStar.Interactive.Ide.fst | 9 ++++++-- src/fstar/FStar.Interactive.Incremental.fsti | 1 + 4 files changed, 33 insertions(+), 6 deletions(-) diff --git a/ocaml/fstar-lib/generated/FStar_Interactive_Ide.ml b/ocaml/fstar-lib/generated/FStar_Interactive_Ide.ml index 914234c0225..498855d396f 100644 --- a/ocaml/fstar-lib/generated/FStar_Interactive_Ide.ml +++ b/ocaml/fstar-lib/generated/FStar_Interactive_Ide.ml @@ -1260,6 +1260,16 @@ let (write_full_buffer_fragment_progress : [uu___1] in write_progress (FStar_Pervasives_Native.Some "full-buffer-fragment-ok") uu___ + | FStar_Interactive_Incremental.FragmentFailed d -> + let uu___ = + let uu___1 = + let uu___2 = + FStar_Compiler_Range.json_of_def_range + d.FStar_Parser_AST.drange in + ("ranges", uu___2) in + [uu___1] in + write_progress + (FStar_Pervasives_Native.Some "full-buffer-fragment-failed") uu___ | FStar_Interactive_Incremental.FragmentError issues -> let qid = let uu___ = FStar_Compiler_Effect.op_Bang repl_current_qid in @@ -1447,10 +1457,14 @@ let (run_push_without_deps : | FStar_Errors.ENotImplemented -> true | uu___3 -> false) errs in ((match code_or_decl with - | FStar_Pervasives.Inr d when Prims.op_Negation has_error - -> - write_full_buffer_fragment_progress - (FStar_Interactive_Incremental.FragmentSuccess d) + | FStar_Pervasives.Inr d -> + if Prims.op_Negation has_error + then + write_full_buffer_fragment_progress + (FStar_Interactive_Incremental.FragmentSuccess d) + else + write_full_buffer_fragment_progress + (FStar_Interactive_Incremental.FragmentFailed d) | uu___4 -> ()); (let json_errors = let uu___4 = diff --git a/ocaml/fstar-lib/generated/FStar_Interactive_Incremental.ml b/ocaml/fstar-lib/generated/FStar_Interactive_Incremental.ml index d9e9cf53bf4..4c860735515 100644 --- a/ocaml/fstar-lib/generated/FStar_Interactive_Incremental.ml +++ b/ocaml/fstar-lib/generated/FStar_Interactive_Incremental.ml @@ -2,6 +2,7 @@ open Prims type fragment_progress = | FragmentStarted of FStar_Parser_AST.decl | FragmentSuccess of FStar_Parser_AST.decl + | FragmentFailed of FStar_Parser_AST.decl | FragmentError of FStar_Errors.issue Prims.list let (uu___is_FragmentStarted : fragment_progress -> Prims.bool) = fun projectee -> @@ -15,6 +16,12 @@ let (uu___is_FragmentSuccess : fragment_progress -> Prims.bool) = let (__proj__FragmentSuccess__item___0 : fragment_progress -> FStar_Parser_AST.decl) = fun projectee -> match projectee with | FragmentSuccess _0 -> _0 +let (uu___is_FragmentFailed : fragment_progress -> Prims.bool) = + fun projectee -> + match projectee with | FragmentFailed _0 -> true | uu___ -> false +let (__proj__FragmentFailed__item___0 : + fragment_progress -> FStar_Parser_AST.decl) = + fun projectee -> match projectee with | FragmentFailed _0 -> _0 let (uu___is_FragmentError : fragment_progress -> Prims.bool) = fun projectee -> match projectee with | FragmentError _0 -> true | uu___ -> false diff --git a/src/fstar/FStar.Interactive.Ide.fst b/src/fstar/FStar.Interactive.Ide.fst index 0e202292bdd..be7e2592f70 100644 --- a/src/fstar/FStar.Interactive.Ide.fst +++ b/src/fstar/FStar.Interactive.Ide.fst @@ -558,6 +558,9 @@ let write_full_buffer_fragment_progress (di:Incremental.fragment_progress) = | FragmentSuccess d -> write_progress (Some "full-buffer-fragment-ok") ["ranges", json_of_def_range d.FStar.Parser.AST.drange] + | FragmentFailed d -> + write_progress (Some "full-buffer-fragment-failed") + ["ranges", json_of_def_range d.FStar.Parser.AST.drange] | FragmentError issues -> let qid = match !repl_current_qid with @@ -607,8 +610,10 @@ let run_push_without_deps st query in let _ = match code_or_decl with - | Inr d when not has_error -> - write_full_buffer_fragment_progress (Incremental.FragmentSuccess d) + | Inr d -> + if not has_error + then write_full_buffer_fragment_progress (Incremental.FragmentSuccess d) + else write_full_buffer_fragment_progress (Incremental.FragmentFailed d) | _ -> () in let json_errors = JsonList (errs |> List.map json_of_issue) in diff --git a/src/fstar/FStar.Interactive.Incremental.fsti b/src/fstar/FStar.Interactive.Incremental.fsti index 3d0e4afb3b3..9d4f9df7459 100755 --- a/src/fstar/FStar.Interactive.Incremental.fsti +++ b/src/fstar/FStar.Interactive.Incremental.fsti @@ -28,6 +28,7 @@ open FStar.Interactive.Ide.Types type fragment_progress = | FragmentStarted of decl | FragmentSuccess of decl + | FragmentFailed of decl | FragmentError of list issue (* Translates a full-buffer(qid, code) query by From 0e1c81f9993cce70a48b9dccf76872ae2e49618d Mon Sep 17 00:00:00 2001 From: Nikhil Swamy Date: Tue, 14 Mar 2023 09:03:16 -0700 Subject: [PATCH 14/48] a code formatting command in the IDE protocol --- .../generated/FStar_Interactive_Ide.ml | 41 ++++++ .../generated/FStar_Interactive_Ide_Types.ml | 10 +- .../FStar_Interactive_Incremental.ml | 124 ++++++++++++------ src/fstar/FStar.Interactive.Ide.Types.fst | 8 +- src/fstar/FStar.Interactive.Ide.fst | 13 ++ src/fstar/FStar.Interactive.Incremental.fst | 61 ++++++--- src/fstar/FStar.Interactive.Incremental.fsti | 4 + 7 files changed, 199 insertions(+), 62 deletions(-) diff --git a/ocaml/fstar-lib/generated/FStar_Interactive_Ide.ml b/ocaml/fstar-lib/generated/FStar_Interactive_Ide.ml index 498855d396f..4fa09c25fdc 100644 --- a/ocaml/fstar-lib/generated/FStar_Interactive_Ide.ml +++ b/ocaml/fstar-lib/generated/FStar_Interactive_Ide.ml @@ -490,6 +490,12 @@ let (unpack_interactive_query : FStar_Interactive_JsonHelper.js_str in (uu___3, uu___4) in FStar_Interactive_Ide_Types.VfsAdd uu___2 + | "format" -> + let uu___2 = + let uu___3 = arg "code" in + FStar_Compiler_Effect.op_Bar_Greater uu___3 + FStar_Interactive_JsonHelper.js_str in + FStar_Interactive_Ide_Types.Format uu___2 | uu___2 -> let uu___3 = FStar_Compiler_Util.format1 "Unknown query '%s'" query in @@ -2293,6 +2299,39 @@ let run_search : (FStar_Interactive_Ide_Types.QueryNOK, (FStar_Compiler_Util.JsonStr s)) in (results, (FStar_Pervasives.Inl st)) +let run_format_code : + 'uuuuu . + FStar_Interactive_ReplState.repl_state -> + Prims.string -> + ((FStar_Interactive_Ide_Types.query_status * + FStar_Compiler_Util.json) * + (FStar_Interactive_ReplState.repl_state, 'uuuuu) + FStar_Pervasives.either) + = + fun st -> + fun code -> + let code_or_err = FStar_Interactive_Incremental.format_code st code in + match code_or_err with + | FStar_Pervasives.Inl code1 -> + let result = + FStar_Compiler_Util.JsonAssoc + [("formatted-code", (FStar_Compiler_Util.JsonStr code1))] in + ((FStar_Interactive_Ide_Types.QueryOK, result), + (FStar_Pervasives.Inl st)) + | FStar_Pervasives.Inr issue -> + let result = + let uu___ = + let uu___1 = + let uu___2 = + let uu___3 = + FStar_Compiler_List.map + FStar_Interactive_Ide_Types.json_of_issue issue in + FStar_Compiler_Util.JsonList uu___3 in + ("formatted-code-issue", uu___2) in + [uu___1] in + FStar_Compiler_Util.JsonAssoc uu___ in + ((FStar_Interactive_Ide_Types.QueryNOK, result), + (FStar_Pervasives.Inl st)) let (as_json_list : ((FStar_Interactive_Ide_Types.query_status * FStar_Compiler_Util.json) * (FStar_Interactive_ReplState.repl_state, Prims.int) @@ -2414,6 +2453,8 @@ let rec (run_query : | FStar_Interactive_Ide_Types.Search term -> let uu___ = run_search st term in as_json_list uu___ | FStar_Interactive_Ide_Types.Callback f -> f st + | FStar_Interactive_Ide_Types.Format code -> + let uu___ = run_format_code st code in as_json_list uu___ and (validate_and_run_query : FStar_Interactive_ReplState.repl_state -> FStar_Interactive_Ide_Types.query -> run_query_result) diff --git a/ocaml/fstar-lib/generated/FStar_Interactive_Ide_Types.ml b/ocaml/fstar-lib/generated/FStar_Interactive_Ide_Types.ml index 02b78d7e366..cbaf6d19e93 100644 --- a/ocaml/fstar-lib/generated/FStar_Interactive_Ide_Types.ml +++ b/ocaml/fstar-lib/generated/FStar_Interactive_Ide_Types.ml @@ -172,6 +172,7 @@ type query' = | ProtocolViolation of Prims.string | FullBuffer of (Prims.string * Prims.bool) | Callback of callback_t + | Format of Prims.string and query = { qq: query' ; qid: Prims.string } @@ -243,6 +244,10 @@ let (uu___is_Callback : query' -> Prims.bool) = match projectee with | Callback _0 -> true | uu___ -> false let (__proj__Callback__item___0 : query' -> callback_t) = fun projectee -> match projectee with | Callback _0 -> _0 +let (uu___is_Format : query' -> Prims.bool) = + fun projectee -> match projectee with | Format _0 -> true | uu___ -> false +let (__proj__Format__item___0 : query' -> Prims.string) = + fun projectee -> match projectee with | Format _0 -> _0 let (__proj__Mkquery__item__qq : query -> query') = fun projectee -> match projectee with | { qq; qid;_} -> qq let (__proj__Mkquery__item__qid : query -> Prims.string) = @@ -293,6 +298,7 @@ let (query_to_string : query -> Prims.string) = | ProtocolViolation uu___ -> "ProtocolViolation" | FullBuffer uu___ -> "FullBuffer" | Callback uu___ -> "Callback" + | Format uu___ -> "Format" let (query_needs_current_module : query' -> Prims.bool) = fun uu___ -> match uu___ with @@ -310,6 +316,7 @@ let (query_needs_current_module : query' -> Prims.bool) = | ProtocolViolation uu___1 -> false | FullBuffer uu___1 -> false | Callback uu___1 -> false + | Format uu___1 -> false | Push uu___1 -> true | AutoComplete uu___1 -> true | Lookup uu___1 -> true @@ -338,7 +345,8 @@ let (interactive_protocol_features : Prims.string Prims.list) = "tactic-ranges"; "interrupt"; "progress"; - "full-buffer"] + "full-buffer"; + "format"] let (json_of_issue_level : FStar_Errors.issue_level -> FStar_Compiler_Util.json) = fun i -> diff --git a/ocaml/fstar-lib/generated/FStar_Interactive_Incremental.ml b/ocaml/fstar-lib/generated/FStar_Interactive_Incremental.ml index 4c860735515..6304808b6d6 100644 --- a/ocaml/fstar-lib/generated/FStar_Interactive_Incremental.ml +++ b/ocaml/fstar-lib/generated/FStar_Interactive_Incremental.ml @@ -215,6 +215,52 @@ let (inspect_repl_stack : let uu___1 = pop_entries es in op_let_Bang uu___1 (fun pops -> return pops) in matching_prefix entries1 ds +let (parse_code : Prims.string -> FStar_Parser_ParseIt.parse_result) = + fun code -> + let uu___ = + let uu___1 = + let uu___2 = + FStar_Compiler_Range.file_of_range + FStar_Interactive_Ide_Types.initial_range in + let uu___3 = + let uu___4 = + FStar_Compiler_Range.start_of_range + FStar_Interactive_Ide_Types.initial_range in + FStar_Compiler_Range.line_of_pos uu___4 in + let uu___4 = + let uu___5 = + FStar_Compiler_Range.start_of_range + FStar_Interactive_Ide_Types.initial_range in + FStar_Compiler_Range.col_of_pos uu___5 in + { + FStar_Parser_ParseIt.frag_fname = uu___2; + FStar_Parser_ParseIt.frag_text = code; + FStar_Parser_ParseIt.frag_line = uu___3; + FStar_Parser_ParseIt.frag_col = uu___4 + } in + FStar_Parser_ParseIt.Incremental uu___1 in + FStar_Parser_ParseIt.parse uu___ +let (syntax_issue : + (FStar_Errors_Codes.raw_error * Prims.string * FStar_Compiler_Range.range) + -> FStar_Errors.issue) + = + fun uu___ -> + match uu___ with + | (raw_error, msg, range) -> + let uu___1 = FStar_Errors.lookup raw_error in + (match uu___1 with + | (uu___2, uu___3, num) -> + let issue = + { + FStar_Errors.issue_msg = msg; + FStar_Errors.issue_level = FStar_Errors.EError; + FStar_Errors.issue_range = + (FStar_Pervasives_Native.Some range); + FStar_Errors.issue_number = + (FStar_Pervasives_Native.Some num); + FStar_Errors.issue_ctx = [] + } in + issue) let (run_full_buffer : FStar_Interactive_ReplState.repl_state -> Prims.string -> @@ -228,49 +274,13 @@ let (run_full_buffer : fun code -> fun full -> fun write_full_buffer_fragment_progress -> - let parse_result = - let uu___ = - let uu___1 = - let uu___2 = - FStar_Compiler_Range.file_of_range - FStar_Interactive_Ide_Types.initial_range in - let uu___3 = - let uu___4 = - FStar_Compiler_Range.start_of_range - FStar_Interactive_Ide_Types.initial_range in - FStar_Compiler_Range.line_of_pos uu___4 in - let uu___4 = - let uu___5 = - FStar_Compiler_Range.start_of_range - FStar_Interactive_Ide_Types.initial_range in - FStar_Compiler_Range.col_of_pos uu___5 in - { - FStar_Parser_ParseIt.frag_fname = uu___2; - FStar_Parser_ParseIt.frag_text = code; - FStar_Parser_ParseIt.frag_line = uu___3; - FStar_Parser_ParseIt.frag_col = uu___4 - } in - FStar_Parser_ParseIt.Incremental uu___1 in - FStar_Parser_ParseIt.parse uu___ in + let parse_result = parse_code code in let log_syntax_issues err = match err with | FStar_Pervasives_Native.None -> () - | FStar_Pervasives_Native.Some (raw_error, msg, range) -> - let uu___ = FStar_Errors.lookup raw_error in - (match uu___ with - | (uu___1, uu___2, num) -> - let issue = - { - FStar_Errors.issue_msg = msg; - FStar_Errors.issue_level = FStar_Errors.EError; - FStar_Errors.issue_range = - (FStar_Pervasives_Native.Some range); - FStar_Errors.issue_number = - (FStar_Pervasives_Native.Some num); - FStar_Errors.issue_ctx = [] - } in - write_full_buffer_fragment_progress - (FragmentError [issue])) in + | FStar_Pervasives_Native.Some err1 -> + let issue = syntax_issue err1 in + write_full_buffer_fragment_progress (FragmentError [issue]) in let qs = match parse_result with | FStar_Parser_ParseIt.IncrementalFragment @@ -303,4 +313,36 @@ let (run_full_buffer : else (); []) | uu___ -> failwith "Unexpected parse result" in - qs \ No newline at end of file + qs +let (format_code : + FStar_Interactive_ReplState.repl_state -> + Prims.string -> + (Prims.string, FStar_Errors.issue Prims.list) FStar_Pervasives.either) + = + fun st -> + fun code -> + let parse_result = parse_code code in + match parse_result with + | FStar_Parser_ParseIt.IncrementalFragment + (decls, uu___, FStar_Pervasives_Native.None) -> + let doc_to_string doc = + FStar_Pprint.pretty_string + (FStar_Compiler_Util.float_of_string "1.0") + (Prims.of_int (100)) doc in + let formatted_code = + let uu___1 = + FStar_Compiler_List.map + (fun d -> + let uu___2 = FStar_Parser_ToDocument.decl_to_document d in + doc_to_string uu___2) decls in + FStar_Compiler_Effect.op_Bar_Greater uu___1 + (FStar_String.concat "\n\n") in + FStar_Pervasives.Inl formatted_code + | FStar_Parser_ParseIt.IncrementalFragment + (uu___, uu___1, FStar_Pervasives_Native.Some err) -> + let uu___2 = let uu___3 = syntax_issue err in [uu___3] in + FStar_Pervasives.Inr uu___2 + | FStar_Parser_ParseIt.ParseError err -> + let uu___ = let uu___1 = syntax_issue err in [uu___1] in + FStar_Pervasives.Inr uu___ + | uu___ -> failwith "Unexpected parse result" \ No newline at end of file diff --git a/src/fstar/FStar.Interactive.Ide.Types.fst b/src/fstar/FStar.Interactive.Ide.Types.fst index a8ecc11ea02..a1fd7594e42 100755 --- a/src/fstar/FStar.Interactive.Ide.Types.fst +++ b/src/fstar/FStar.Interactive.Ide.Types.fst @@ -146,6 +146,7 @@ type query' = | ProtocolViolation of string | FullBuffer of string & bool //if true, check the buffer, otherwise just find verified prefix | Callback of callback_t +| Format of string and query = { qq: query'; qid: string } let push_query_to_string pq = @@ -183,12 +184,12 @@ let query_to_string q = match q.qq with | ProtocolViolation _ -> "ProtocolViolation" | FullBuffer _ -> "FullBuffer" | Callback _ -> "Callback" - +| Format _ -> "Format" let query_needs_current_module = function | Exit | DescribeProtocol | DescribeRepl | Segment _ | Pop | Push { push_peek_only = false } | VfsAdd _ - | GenericError _ | ProtocolViolation _ | FullBuffer _ | Callback _ -> false + | GenericError _ | ProtocolViolation _ | FullBuffer _ | Callback _ | Format _ -> false | Push _ | AutoComplete _ | Lookup _ | Compute _ | Search _ -> true let interactive_protocol_vernum = 2 @@ -199,7 +200,8 @@ let interactive_protocol_features = "describe-protocol"; "describe-repl"; "exit"; "lookup"; "lookup/context"; "lookup/documentation"; "lookup/definition"; "peek"; "pop"; "push"; "search"; "segment"; - "vfs-add"; "tactic-ranges"; "interrupt"; "progress"; "full-buffer"] + "vfs-add"; "tactic-ranges"; "interrupt"; "progress"; + "full-buffer"; "format"] let json_of_issue_level i = JsonStr (match i with diff --git a/src/fstar/FStar.Interactive.Ide.fst b/src/fstar/FStar.Interactive.Ide.fst index be7e2592f70..0f81cc59a2d 100644 --- a/src/fstar/FStar.Interactive.Ide.fst +++ b/src/fstar/FStar.Interactive.Ide.fst @@ -246,6 +246,7 @@ let unpack_interactive_query json = | "search" -> Search (arg "terms" |> js_str) | "vfs-add" -> VfsAdd (try_arg "filename" |> Util.map_option js_str, arg "contents" |> js_str) + | "format" -> Format (arg "code" |> js_str) | _ -> ProtocolViolation (Util.format1 "Unknown query '%s'" query) } with | InvalidQuery msg -> { qid = qid; qq = ProtocolViolation msg } @@ -909,6 +910,16 @@ let run_search st search_str = with InvalidSearch s -> (QueryNOK, JsonStr s) in (results, Inl st) +let run_format_code st code = + let code_or_err = FStar.Interactive.Incremental.format_code st code in + match code_or_err with + | Inl code -> + let result = JsonAssoc ["formatted-code", JsonStr code] in + (QueryOK, result), Inl st + | Inr issue -> + let result = JsonAssoc ["formatted-code-issue", JsonList (List.map json_of_issue issue)] in + (QueryNOK, result), Inl st + let as_json_list (q: (query_status & json) & either repl_state int) : (query_status & list json) & either repl_state int = let (q, j), s = q in @@ -968,6 +979,8 @@ let rec run_query st (q: query) : (query_status * list json) * either repl_state as_json_list (run_search st term) | Callback f -> f st + | Format code -> + as_json_list (run_format_code st code) and validate_and_run_query st query = let query = validate_query st query in repl_current_qid := Some query.qid; diff --git a/src/fstar/FStar.Interactive.Incremental.fst b/src/fstar/FStar.Interactive.Incremental.fst index 7ad32d77b40..4c41a89cdda 100755 --- a/src/fstar/FStar.Interactive.Incremental.fst +++ b/src/fstar/FStar.Interactive.Incremental.fst @@ -177,32 +177,38 @@ let inspect_repl_stack (s:repl_stack_t) in matching_prefix entries ds +let parse_code (code:string) = + P.parse (Incremental { + frag_fname = Range.file_of_range initial_range; + frag_text = code; + frag_line = Range.line_of_pos (Range.start_of_range initial_range); + frag_col = Range.col_of_pos (Range.start_of_range initial_range); + }) + + +let syntax_issue (raw_error, msg, range) = + let _, _, num = FStar.Errors.lookup raw_error in + let issue = { + issue_msg = msg; + issue_level = EError; + issue_range = Some range; + issue_number = Some num; + issue_ctx = [] + } in + issue + let run_full_buffer (st:repl_state) (qid:string) (code:string) (full:bool) (write_full_buffer_fragment_progress: fragment_progress -> unit) : list query - = let parse_result = - P.parse (Incremental { - frag_fname = Range.file_of_range initial_range; - frag_text = code; - frag_line = Range.line_of_pos (Range.start_of_range initial_range); - frag_col = Range.col_of_pos (Range.start_of_range initial_range); - }) - in + = let parse_result = parse_code code in let log_syntax_issues err = match err with | None -> () - | Some (raw_error, msg, range) -> - let _, _, num = FStar.Errors.lookup raw_error in - let issue = { - issue_msg = msg; - issue_level = EError; - issue_range = Some range; - issue_number = Some num; - issue_ctx = [] - } in + | Some err -> + let issue = syntax_issue err in write_full_buffer_fragment_progress (FragmentError [issue]) in let qs = @@ -226,3 +232,24 @@ let run_full_buffer (st:repl_state) failwith "Unexpected parse result" in qs + +let format_code (st:repl_state) (code:string) + = let parse_result = parse_code code in + match parse_result with + | IncrementalFragment (decls, _, None) -> + let doc_to_string doc = + FStar.Pprint.pretty_string (float_of_string "1.0") 100 doc + in + let formatted_code = + List.map + (fun d -> doc_to_string (FStar.Parser.ToDocument.decl_to_document d)) + decls + |> String.concat "\n\n" + in + Inl formatted_code + | IncrementalFragment (_, _, Some err) -> + Inr [syntax_issue err] + | ParseError err -> + Inr [syntax_issue err] + | _ -> + failwith "Unexpected parse result" \ No newline at end of file diff --git a/src/fstar/FStar.Interactive.Incremental.fsti b/src/fstar/FStar.Interactive.Incremental.fsti index 9d4f9df7459..1064578f313 100755 --- a/src/fstar/FStar.Interactive.Incremental.fsti +++ b/src/fstar/FStar.Interactive.Incremental.fsti @@ -47,3 +47,7 @@ val run_full_buffer (st:repl_state) (full:bool) (write_full_buffer_fragment_progress: fragment_progress -> unit) : list query + +val format_code (st:repl_state) + (code:string) + : either string (list issue) \ No newline at end of file From a8afe5933ecffc7213d994e50345298998a46e47 Mon Sep 17 00:00:00 2001 From: Nikhil Swamy Date: Sat, 18 Mar 2023 11:47:35 -0700 Subject: [PATCH 15/48] adding a command to reload-deps; preparing for handling a Cancel request --- ocaml/fstar-lib/FStar_Compiler_Util.ml | 9 + .../generated/FStar_Interactive_Ide.ml | 620 +++++++++++------- .../generated/FStar_Interactive_Ide_Types.ml | 410 +++++++++--- .../FStar_Interactive_Incremental.ml | 139 +++- .../generated/FStar_Interactive_Lsp.ml | 55 +- .../generated/FStar_Interactive_PushHelper.ml | 420 ++++++------ .../FStar_Interactive_QueryHelper.ml | 8 +- .../generated/FStar_Interactive_ReplState.ml | 185 +----- src/basic/boot/FStar.Compiler.Util.fsti | 1 + src/fstar/FStar.Interactive.Ide.Types.fst | 222 +++++-- src/fstar/FStar.Interactive.Ide.fst | 61 +- src/fstar/FStar.Interactive.Incremental.fst | 54 +- src/fstar/FStar.Interactive.Incremental.fsti | 2 +- src/fstar/FStar.Interactive.Lsp.fst | 14 +- src/fstar/FStar.Interactive.PushHelper.fst | 2 +- src/fstar/FStar.Interactive.PushHelper.fsti | 5 +- src/fstar/FStar.Interactive.QueryHelper.fsti | 2 +- src/fstar/FStar.Interactive.ReplState.fst | 30 +- 18 files changed, 1343 insertions(+), 896 deletions(-) diff --git a/ocaml/fstar-lib/FStar_Compiler_Util.ml b/ocaml/fstar-lib/FStar_Compiler_Util.ml index f6ebb2537bd..2294217ee0c 100644 --- a/ocaml/fstar-lib/FStar_Compiler_Util.ml +++ b/ocaml/fstar-lib/FStar_Compiler_Util.ml @@ -296,6 +296,15 @@ let nread (s:stream_reader) (n:Z.t) = with _ -> None +let poll_stdin (f:float) = + try + let ready_fds, _, _ = Unix.select [Unix.stdin] [] [] f in + match ready_fds with + | [] -> false + | _ -> true + with + | _ -> false + type string_builder = BatBuffer.t let new_string_builder () = BatBuffer.create 256 let clear_string_builder b = BatBuffer.clear b diff --git a/ocaml/fstar-lib/generated/FStar_Interactive_Ide.ml b/ocaml/fstar-lib/generated/FStar_Interactive_Ide.ml index 4fa09c25fdc..8a5f0754ad8 100644 --- a/ocaml/fstar-lib/generated/FStar_Interactive_Ide.ml +++ b/ocaml/fstar-lib/generated/FStar_Interactive_Ide.ml @@ -61,7 +61,7 @@ let (repl_current_qid : Prims.string FStar_Pervasives_Native.option FStar_Compiler_Effect.ref) = FStar_Compiler_Util.mk_ref FStar_Pervasives_Native.None let (nothing_left_to_pop : - FStar_Interactive_ReplState.repl_state -> Prims.bool) = + FStar_Interactive_Ide_Types.repl_state -> Prims.bool) = fun st -> let uu___ = let uu___1 = @@ -69,13 +69,13 @@ let (nothing_left_to_pop : FStar_Compiler_List.length uu___1 in uu___ = (FStar_Compiler_List.length - st.FStar_Interactive_ReplState.repl_deps_stack) + st.FStar_Interactive_Ide_Types.repl_deps_stack) let (run_repl_transaction : - FStar_Interactive_ReplState.repl_state -> - FStar_Interactive_PushHelper.push_kind -> + FStar_Interactive_Ide_Types.repl_state -> + FStar_Interactive_Ide_Types.push_kind -> Prims.bool -> - FStar_Interactive_ReplState.repl_task -> - (Prims.bool * FStar_Interactive_ReplState.repl_state)) + FStar_Interactive_Ide_Types.repl_task -> + (Prims.bool * FStar_Interactive_Ide_Types.repl_state)) = fun st -> fun push_kind -> @@ -86,7 +86,7 @@ let (run_repl_transaction : push_kind task st in let uu___ = FStar_Interactive_PushHelper.track_name_changes - st1.FStar_Interactive_ReplState.repl_env in + st1.FStar_Interactive_Ide_Types.repl_env in match uu___ with | (env, finish_name_tracking) -> let check_success uu___1 = @@ -99,7 +99,7 @@ let (run_repl_transaction : (fun env1 -> let uu___3 = FStar_Interactive_PushHelper.run_repl_task - st1.FStar_Interactive_ReplState.repl_curmod env1 + st1.FStar_Interactive_Ide_Types.repl_curmod env1 task in FStar_Compiler_Effect.op_Less_Bar (fun uu___4 -> FStar_Pervasives_Native.Some uu___4) @@ -108,7 +108,7 @@ let (run_repl_transaction : | FStar_Pervasives_Native.Some (curmod, env1) when check_success () -> (curmod, env1, true) | uu___3 -> - ((st1.FStar_Interactive_ReplState.repl_curmod), env, + ((st1.FStar_Interactive_Ide_Types.repl_curmod), env, false) in (match uu___1 with | (curmod, env1, success) -> @@ -120,21 +120,24 @@ let (run_repl_transaction : then let st3 = { - FStar_Interactive_ReplState.repl_line = - (st1.FStar_Interactive_ReplState.repl_line); - FStar_Interactive_ReplState.repl_column = - (st1.FStar_Interactive_ReplState.repl_column); - FStar_Interactive_ReplState.repl_fname = - (st1.FStar_Interactive_ReplState.repl_fname); - FStar_Interactive_ReplState.repl_deps_stack = - (st1.FStar_Interactive_ReplState.repl_deps_stack); - FStar_Interactive_ReplState.repl_curmod = + FStar_Interactive_Ide_Types.repl_line = + (st1.FStar_Interactive_Ide_Types.repl_line); + FStar_Interactive_Ide_Types.repl_column = + (st1.FStar_Interactive_Ide_Types.repl_column); + FStar_Interactive_Ide_Types.repl_fname = + (st1.FStar_Interactive_Ide_Types.repl_fname); + FStar_Interactive_Ide_Types.repl_deps_stack = + (st1.FStar_Interactive_Ide_Types.repl_deps_stack); + FStar_Interactive_Ide_Types.repl_curmod = curmod; - FStar_Interactive_ReplState.repl_env = env2; - FStar_Interactive_ReplState.repl_stdin = - (st1.FStar_Interactive_ReplState.repl_stdin); - FStar_Interactive_ReplState.repl_names = - (st1.FStar_Interactive_ReplState.repl_names) + FStar_Interactive_Ide_Types.repl_env = env2; + FStar_Interactive_Ide_Types.repl_stdin = + (st1.FStar_Interactive_Ide_Types.repl_stdin); + FStar_Interactive_Ide_Types.repl_names = + (st1.FStar_Interactive_Ide_Types.repl_names); + FStar_Interactive_Ide_Types.repl_buffered_input_queries + = + (st1.FStar_Interactive_Ide_Types.repl_buffered_input_queries) } in FStar_Interactive_PushHelper.commit_name_tracking st3 name_events @@ -143,11 +146,11 @@ let (run_repl_transaction : "run_repl_transaction" st1 in (success, st2))) let (run_repl_ld_transactions : - FStar_Interactive_ReplState.repl_state -> - FStar_Interactive_ReplState.repl_task Prims.list -> - (FStar_Interactive_ReplState.repl_task -> unit) -> - (FStar_Interactive_ReplState.repl_state, - FStar_Interactive_ReplState.repl_state) FStar_Pervasives.either) + FStar_Interactive_Ide_Types.repl_state -> + FStar_Interactive_Ide_Types.repl_task Prims.list -> + (FStar_Interactive_Ide_Types.repl_task -> unit) -> + (FStar_Interactive_Ide_Types.repl_state, + FStar_Interactive_Ide_Types.repl_state) FStar_Pervasives.either) = fun st -> fun tasks -> @@ -156,7 +159,7 @@ let (run_repl_ld_transactions : let uu___ = FStar_Options.debug_any () in if uu___ then - let uu___1 = FStar_Interactive_ReplState.string_of_repl_task task in + let uu___1 = FStar_Interactive_Ide_Types.string_of_repl_task task in FStar_Compiler_Util.print2 "%s %s" verb uu___1 else () in let rec revert_many st1 uu___ = @@ -181,27 +184,29 @@ let (run_repl_ld_transactions : "run_repl_ls_transactions" st1 in let dep_graph = FStar_TypeChecker_Env.dep_graph - st1.FStar_Interactive_ReplState.repl_env in + st1.FStar_Interactive_Ide_Types.repl_env in let st'1 = let uu___3 = FStar_TypeChecker_Env.set_dep_graph - st'.FStar_Interactive_ReplState.repl_env dep_graph in + st'.FStar_Interactive_Ide_Types.repl_env dep_graph in { - FStar_Interactive_ReplState.repl_line = - (st'.FStar_Interactive_ReplState.repl_line); - FStar_Interactive_ReplState.repl_column = - (st'.FStar_Interactive_ReplState.repl_column); - FStar_Interactive_ReplState.repl_fname = - (st'.FStar_Interactive_ReplState.repl_fname); - FStar_Interactive_ReplState.repl_deps_stack = - (st'.FStar_Interactive_ReplState.repl_deps_stack); - FStar_Interactive_ReplState.repl_curmod = - (st'.FStar_Interactive_ReplState.repl_curmod); - FStar_Interactive_ReplState.repl_env = uu___3; - FStar_Interactive_ReplState.repl_stdin = - (st'.FStar_Interactive_ReplState.repl_stdin); - FStar_Interactive_ReplState.repl_names = - (st'.FStar_Interactive_ReplState.repl_names) + FStar_Interactive_Ide_Types.repl_line = + (st'.FStar_Interactive_Ide_Types.repl_line); + FStar_Interactive_Ide_Types.repl_column = + (st'.FStar_Interactive_Ide_Types.repl_column); + FStar_Interactive_Ide_Types.repl_fname = + (st'.FStar_Interactive_Ide_Types.repl_fname); + FStar_Interactive_Ide_Types.repl_deps_stack = + (st'.FStar_Interactive_Ide_Types.repl_deps_stack); + FStar_Interactive_Ide_Types.repl_curmod = + (st'.FStar_Interactive_Ide_Types.repl_curmod); + FStar_Interactive_Ide_Types.repl_env = uu___3; + FStar_Interactive_Ide_Types.repl_stdin = + (st'.FStar_Interactive_Ide_Types.repl_stdin); + FStar_Interactive_Ide_Types.repl_names = + (st'.FStar_Interactive_Ide_Types.repl_names); + FStar_Interactive_Ide_Types.repl_buffered_input_queries = + (st'.FStar_Interactive_Ide_Types.repl_buffered_input_queries) } in revert_many st'1 entries)) in let rec aux st1 tasks1 previous = @@ -218,8 +223,8 @@ let (run_repl_ld_transactions : let push_kind = let uu___3 = FStar_Options.lax () in if uu___3 - then FStar_Interactive_PushHelper.LaxCheck - else FStar_Interactive_PushHelper.FullCheck in + then FStar_Interactive_Ide_Types.LaxCheck + else FStar_Interactive_Ide_Types.FullCheck in let uu___3 = run_repl_transaction st1 push_kind false timestamped_task in match uu___3 with @@ -231,22 +236,25 @@ let (run_repl_ld_transactions : FStar_Compiler_Effect.op_Bang FStar_Interactive_PushHelper.repl_stack in { - FStar_Interactive_ReplState.repl_line = - (st2.FStar_Interactive_ReplState.repl_line); - FStar_Interactive_ReplState.repl_column = - (st2.FStar_Interactive_ReplState.repl_column); - FStar_Interactive_ReplState.repl_fname = - (st2.FStar_Interactive_ReplState.repl_fname); - FStar_Interactive_ReplState.repl_deps_stack = + FStar_Interactive_Ide_Types.repl_line = + (st2.FStar_Interactive_Ide_Types.repl_line); + FStar_Interactive_Ide_Types.repl_column = + (st2.FStar_Interactive_Ide_Types.repl_column); + FStar_Interactive_Ide_Types.repl_fname = + (st2.FStar_Interactive_Ide_Types.repl_fname); + FStar_Interactive_Ide_Types.repl_deps_stack = uu___5; - FStar_Interactive_ReplState.repl_curmod = - (st2.FStar_Interactive_ReplState.repl_curmod); - FStar_Interactive_ReplState.repl_env = - (st2.FStar_Interactive_ReplState.repl_env); - FStar_Interactive_ReplState.repl_stdin = - (st2.FStar_Interactive_ReplState.repl_stdin); - FStar_Interactive_ReplState.repl_names = - (st2.FStar_Interactive_ReplState.repl_names) + FStar_Interactive_Ide_Types.repl_curmod = + (st2.FStar_Interactive_Ide_Types.repl_curmod); + FStar_Interactive_Ide_Types.repl_env = + (st2.FStar_Interactive_Ide_Types.repl_env); + FStar_Interactive_Ide_Types.repl_stdin = + (st2.FStar_Interactive_Ide_Types.repl_stdin); + FStar_Interactive_Ide_Types.repl_names = + (st2.FStar_Interactive_Ide_Types.repl_names); + FStar_Interactive_Ide_Types.repl_buffered_input_queries + = + (st2.FStar_Interactive_Ide_Types.repl_buffered_input_queries) } in aux uu___4 tasks2 [] else FStar_Pervasives.Inr st2)) @@ -260,7 +268,7 @@ let (run_repl_ld_transactions : let uu___ = revert_many st1 previous1 in aux uu___ tasks2 [] in aux st tasks (FStar_Compiler_List.rev - st.FStar_Interactive_ReplState.repl_deps_stack) + st.FStar_Interactive_Ide_Types.repl_deps_stack) let (wrap_js_failure : Prims.string -> Prims.string -> @@ -333,6 +341,15 @@ let (unpack_interactive_query : FStar_Compiler_Effect.op_Bar_Greater uu___4 FStar_Interactive_JsonHelper.js_int in (uu___1, uu___2, uu___3) in + let parse_full_buffer_kind uu___1 = + match uu___1 with + | "full" -> FStar_Interactive_Ide_Types.Full + | "cache" -> FStar_Interactive_Ide_Types.Cache + | "reload-deps" -> FStar_Interactive_Ide_Types.ReloadDeps + | uu___2 -> + FStar_Compiler_Effect.raise + (FStar_Interactive_JsonHelper.InvalidQuery + "Invalid full-buffer kind") in let uu___1 = match query with | "exit" -> FStar_Interactive_Ide_Types.Exit @@ -415,7 +432,7 @@ let (unpack_interactive_query : let uu___6 = arg "kind" in FStar_Compiler_Effect.op_Bar_Greater uu___6 FStar_Interactive_JsonHelper.js_str in - uu___5 = "full" in + parse_full_buffer_kind uu___5 in (uu___3, uu___4) in FStar_Interactive_Ide_Types.FullBuffer uu___2 | "autocomplete" -> @@ -542,14 +559,111 @@ let (parse_interactive_query : } | FStar_Pervasives_Native.Some request -> deserialize_interactive_query request +let (buffer_input_queries : + FStar_Interactive_Ide_Types.repl_state -> + FStar_Interactive_Ide_Types.repl_state) + = + fun st -> + let rec aux qs st1 = + let done1 qs1 st2 = + { + FStar_Interactive_Ide_Types.repl_line = + (st2.FStar_Interactive_Ide_Types.repl_line); + FStar_Interactive_Ide_Types.repl_column = + (st2.FStar_Interactive_Ide_Types.repl_column); + FStar_Interactive_Ide_Types.repl_fname = + (st2.FStar_Interactive_Ide_Types.repl_fname); + FStar_Interactive_Ide_Types.repl_deps_stack = + (st2.FStar_Interactive_Ide_Types.repl_deps_stack); + FStar_Interactive_Ide_Types.repl_curmod = + (st2.FStar_Interactive_Ide_Types.repl_curmod); + FStar_Interactive_Ide_Types.repl_env = + (st2.FStar_Interactive_Ide_Types.repl_env); + FStar_Interactive_Ide_Types.repl_stdin = + (st2.FStar_Interactive_Ide_Types.repl_stdin); + FStar_Interactive_Ide_Types.repl_names = + (st2.FStar_Interactive_Ide_Types.repl_names); + FStar_Interactive_Ide_Types.repl_buffered_input_queries = + (FStar_Compiler_List.op_At + st2.FStar_Interactive_Ide_Types.repl_buffered_input_queries + (FStar_Compiler_List.rev qs1)) + } in + let uu___ = + let uu___1 = + FStar_Compiler_Util.poll_stdin + (FStar_Compiler_Util.float_of_string "0.0") in + Prims.op_Negation uu___1 in + if uu___ + then done1 qs st1 + else + (let uu___2 = + FStar_Compiler_Util.read_line + st1.FStar_Interactive_Ide_Types.repl_stdin in + match uu___2 with + | FStar_Pervasives_Native.None -> done1 qs st1 + | FStar_Pervasives_Native.Some line -> + let q = parse_interactive_query line in + (match q.FStar_Interactive_Ide_Types.qq with + | FStar_Interactive_Ide_Types.Cancel uu___3 -> + { + FStar_Interactive_Ide_Types.repl_line = + (st1.FStar_Interactive_Ide_Types.repl_line); + FStar_Interactive_Ide_Types.repl_column = + (st1.FStar_Interactive_Ide_Types.repl_column); + FStar_Interactive_Ide_Types.repl_fname = + (st1.FStar_Interactive_Ide_Types.repl_fname); + FStar_Interactive_Ide_Types.repl_deps_stack = + (st1.FStar_Interactive_Ide_Types.repl_deps_stack); + FStar_Interactive_Ide_Types.repl_curmod = + (st1.FStar_Interactive_Ide_Types.repl_curmod); + FStar_Interactive_Ide_Types.repl_env = + (st1.FStar_Interactive_Ide_Types.repl_env); + FStar_Interactive_Ide_Types.repl_stdin = + (st1.FStar_Interactive_Ide_Types.repl_stdin); + FStar_Interactive_Ide_Types.repl_names = + (st1.FStar_Interactive_Ide_Types.repl_names); + FStar_Interactive_Ide_Types.repl_buffered_input_queries = + [q] + } + | uu___3 -> aux (q :: qs) st1)) in + aux [] st let (read_interactive_query : - FStar_Compiler_Util.stream_reader -> FStar_Interactive_Ide_Types.query) = - fun stream -> - let uu___ = FStar_Compiler_Util.read_line stream in - match uu___ with - | FStar_Pervasives_Native.None -> - FStar_Compiler_Effect.exit Prims.int_zero - | FStar_Pervasives_Native.Some line -> parse_interactive_query line + FStar_Interactive_Ide_Types.repl_state -> + (FStar_Interactive_Ide_Types.query * + FStar_Interactive_Ide_Types.repl_state)) + = + fun st -> + match st.FStar_Interactive_Ide_Types.repl_buffered_input_queries with + | [] -> + let uu___ = + FStar_Compiler_Util.read_line + st.FStar_Interactive_Ide_Types.repl_stdin in + (match uu___ with + | FStar_Pervasives_Native.None -> + FStar_Compiler_Effect.exit Prims.int_zero + | FStar_Pervasives_Native.Some line -> + let uu___1 = parse_interactive_query line in (uu___1, st)) + | q::qs -> + (q, + { + FStar_Interactive_Ide_Types.repl_line = + (st.FStar_Interactive_Ide_Types.repl_line); + FStar_Interactive_Ide_Types.repl_column = + (st.FStar_Interactive_Ide_Types.repl_column); + FStar_Interactive_Ide_Types.repl_fname = + (st.FStar_Interactive_Ide_Types.repl_fname); + FStar_Interactive_Ide_Types.repl_deps_stack = + (st.FStar_Interactive_Ide_Types.repl_deps_stack); + FStar_Interactive_Ide_Types.repl_curmod = + (st.FStar_Interactive_Ide_Types.repl_curmod); + FStar_Interactive_Ide_Types.repl_env = + (st.FStar_Interactive_Ide_Types.repl_env); + FStar_Interactive_Ide_Types.repl_stdin = + (st.FStar_Interactive_Ide_Types.repl_stdin); + FStar_Interactive_Ide_Types.repl_names = + (st.FStar_Interactive_Ide_Types.repl_names); + FStar_Interactive_Ide_Types.repl_buffered_input_queries = qs + }) let json_of_opt : 'uuuuu . ('uuuuu -> FStar_Compiler_Util.json) -> @@ -945,19 +1059,19 @@ let (trim_option_name : Prims.string -> (Prims.string * Prims.string)) = (opt_prefix, uu___) else ("", opt_name) let (json_of_repl_state : - FStar_Interactive_ReplState.repl_state -> FStar_Compiler_Util.json) = + FStar_Interactive_Ide_Types.repl_state -> FStar_Compiler_Util.json) = fun st -> let filenames uu___ = match uu___ with | (uu___1, (task, uu___2)) -> (match task with - | FStar_Interactive_ReplState.LDInterleaved (intf, impl) -> - [intf.FStar_Interactive_ReplState.tf_fname; - impl.FStar_Interactive_ReplState.tf_fname] - | FStar_Interactive_ReplState.LDSingle intf_or_impl -> - [intf_or_impl.FStar_Interactive_ReplState.tf_fname] - | FStar_Interactive_ReplState.LDInterfaceOfCurrentFile intf -> - [intf.FStar_Interactive_ReplState.tf_fname] + | FStar_Interactive_Ide_Types.LDInterleaved (intf, impl) -> + [intf.FStar_Interactive_Ide_Types.tf_fname; + impl.FStar_Interactive_Ide_Types.tf_fname] + | FStar_Interactive_Ide_Types.LDSingle intf_or_impl -> + [intf_or_impl.FStar_Interactive_Ide_Types.tf_fname] + | FStar_Interactive_Ide_Types.LDInterfaceOfCurrentFile intf -> + [intf.FStar_Interactive_Ide_Types.tf_fname] | uu___3 -> []) in let uu___ = let uu___1 = @@ -965,7 +1079,7 @@ let (json_of_repl_state : let uu___3 = let uu___4 = FStar_Compiler_List.concatMap filenames - st.FStar_Interactive_ReplState.repl_deps_stack in + st.FStar_Interactive_Ide_Types.repl_deps_stack in FStar_Compiler_List.map (fun uu___5 -> FStar_Compiler_Util.JsonStr uu___5) uu___4 in FStar_Compiler_Util.JsonList uu___3 in @@ -1002,9 +1116,9 @@ let run_describe_protocol : (FStar_Pervasives.Inl st)) let run_describe_repl : 'uuuuu . - FStar_Interactive_ReplState.repl_state -> + FStar_Interactive_Ide_Types.repl_state -> ((FStar_Interactive_Ide_Types.query_status * FStar_Compiler_Util.json) - * (FStar_Interactive_ReplState.repl_state, 'uuuuu) + * (FStar_Interactive_Ide_Types.repl_state, 'uuuuu) FStar_Pervasives.either) = fun st -> @@ -1041,11 +1155,11 @@ let (collect_errors : unit -> FStar_Errors.issue Prims.list) = let errors = FStar_Errors.report_all () in FStar_Errors.clear (); errors let run_segment : 'uuuuu . - FStar_Interactive_ReplState.repl_state -> + FStar_Interactive_Ide_Types.repl_state -> Prims.string -> ((FStar_Interactive_Ide_Types.query_status * FStar_Compiler_Util.json) * - (FStar_Interactive_ReplState.repl_state, 'uuuuu) + (FStar_Interactive_Ide_Types.repl_state, 'uuuuu) FStar_Pervasives.either) = fun st -> @@ -1067,7 +1181,7 @@ let run_segment : | FStar_Parser_Driver.Modul (FStar_Parser_AST.Interface (uu___2, decls, uu___3)) -> decls in let uu___ = - with_captured_errors st.FStar_Interactive_ReplState.repl_env + with_captured_errors st.FStar_Interactive_Ide_Types.repl_env FStar_Compiler_Util.sigint_ignore (fun uu___1 -> let uu___2 = collect_decls () in @@ -1102,28 +1216,28 @@ let run_segment : (FStar_Pervasives.Inl st)) let run_vfs_add : 'uuuuu . - FStar_Interactive_ReplState.repl_state -> + FStar_Interactive_Ide_Types.repl_state -> Prims.string FStar_Pervasives_Native.option -> Prims.string -> ((FStar_Interactive_Ide_Types.query_status * FStar_Compiler_Util.json) * - (FStar_Interactive_ReplState.repl_state, 'uuuuu) + (FStar_Interactive_Ide_Types.repl_state, 'uuuuu) FStar_Pervasives.either) = fun st -> fun opt_fname -> fun contents -> let fname = - FStar_Compiler_Util.dflt st.FStar_Interactive_ReplState.repl_fname + FStar_Compiler_Util.dflt st.FStar_Interactive_Ide_Types.repl_fname opt_fname in FStar_Parser_ParseIt.add_vfs_entry fname contents; ((FStar_Interactive_Ide_Types.QueryOK, FStar_Compiler_Util.JsonNull), (FStar_Pervasives.Inl st)) let run_pop : 'uuuuu . - FStar_Interactive_ReplState.repl_state -> + FStar_Interactive_Ide_Types.repl_state -> ((FStar_Interactive_Ide_Types.query_status * FStar_Compiler_Util.json) - * (FStar_Interactive_ReplState.repl_state, 'uuuuu) + * (FStar_Interactive_Ide_Types.repl_state, 'uuuuu) FStar_Pervasives.either) = fun st -> @@ -1159,41 +1273,41 @@ let (write_error : json_of_message "error" (FStar_Compiler_Util.JsonAssoc contents) in FStar_Interactive_JsonHelper.write_json uu___ let (write_repl_ld_task_progress : - FStar_Interactive_ReplState.repl_task -> unit) = + FStar_Interactive_Ide_Types.repl_task -> unit) = fun task -> match task with - | FStar_Interactive_ReplState.LDInterleaved (uu___, tf) -> + | FStar_Interactive_Ide_Types.LDInterleaved (uu___, tf) -> let modname = FStar_Parser_Dep.module_name_of_file - tf.FStar_Interactive_ReplState.tf_fname in + tf.FStar_Interactive_Ide_Types.tf_fname in write_progress (FStar_Pervasives_Native.Some "loading-dependency") [("modname", (FStar_Compiler_Util.JsonStr modname))] - | FStar_Interactive_ReplState.LDSingle tf -> + | FStar_Interactive_Ide_Types.LDSingle tf -> let modname = FStar_Parser_Dep.module_name_of_file - tf.FStar_Interactive_ReplState.tf_fname in + tf.FStar_Interactive_Ide_Types.tf_fname in write_progress (FStar_Pervasives_Native.Some "loading-dependency") [("modname", (FStar_Compiler_Util.JsonStr modname))] - | FStar_Interactive_ReplState.LDInterfaceOfCurrentFile tf -> + | FStar_Interactive_Ide_Types.LDInterfaceOfCurrentFile tf -> let modname = FStar_Parser_Dep.module_name_of_file - tf.FStar_Interactive_ReplState.tf_fname in + tf.FStar_Interactive_Ide_Types.tf_fname in write_progress (FStar_Pervasives_Native.Some "loading-dependency") [("modname", (FStar_Compiler_Util.JsonStr modname))] | uu___ -> () let (load_deps : - FStar_Interactive_ReplState.repl_state -> - ((FStar_Interactive_ReplState.repl_state * Prims.string Prims.list), - FStar_Interactive_ReplState.repl_state) FStar_Pervasives.either) + FStar_Interactive_Ide_Types.repl_state -> + ((FStar_Interactive_Ide_Types.repl_state * Prims.string Prims.list), + FStar_Interactive_Ide_Types.repl_state) FStar_Pervasives.either) = fun st -> let uu___ = - with_captured_errors st.FStar_Interactive_ReplState.repl_env + with_captured_errors st.FStar_Interactive_Ide_Types.repl_env FStar_Compiler_Util.sigint_ignore (fun _env -> let uu___1 = FStar_Interactive_PushHelper.deps_and_repl_ld_tasks_of_our_file - st.FStar_Interactive_ReplState.repl_fname in + st.FStar_Interactive_Ide_Types.repl_fname in FStar_Compiler_Effect.op_Less_Bar (fun uu___2 -> FStar_Pervasives_Native.Some uu___2) uu___1) in match uu___ with @@ -1202,23 +1316,25 @@ let (load_deps : let st1 = let uu___1 = FStar_TypeChecker_Env.set_dep_graph - st.FStar_Interactive_ReplState.repl_env dep_graph in + st.FStar_Interactive_Ide_Types.repl_env dep_graph in { - FStar_Interactive_ReplState.repl_line = - (st.FStar_Interactive_ReplState.repl_line); - FStar_Interactive_ReplState.repl_column = - (st.FStar_Interactive_ReplState.repl_column); - FStar_Interactive_ReplState.repl_fname = - (st.FStar_Interactive_ReplState.repl_fname); - FStar_Interactive_ReplState.repl_deps_stack = - (st.FStar_Interactive_ReplState.repl_deps_stack); - FStar_Interactive_ReplState.repl_curmod = - (st.FStar_Interactive_ReplState.repl_curmod); - FStar_Interactive_ReplState.repl_env = uu___1; - FStar_Interactive_ReplState.repl_stdin = - (st.FStar_Interactive_ReplState.repl_stdin); - FStar_Interactive_ReplState.repl_names = - (st.FStar_Interactive_ReplState.repl_names) + FStar_Interactive_Ide_Types.repl_line = + (st.FStar_Interactive_Ide_Types.repl_line); + FStar_Interactive_Ide_Types.repl_column = + (st.FStar_Interactive_Ide_Types.repl_column); + FStar_Interactive_Ide_Types.repl_fname = + (st.FStar_Interactive_Ide_Types.repl_fname); + FStar_Interactive_Ide_Types.repl_deps_stack = + (st.FStar_Interactive_Ide_Types.repl_deps_stack); + FStar_Interactive_Ide_Types.repl_curmod = + (st.FStar_Interactive_Ide_Types.repl_curmod); + FStar_Interactive_Ide_Types.repl_env = uu___1; + FStar_Interactive_Ide_Types.repl_stdin = + (st.FStar_Interactive_Ide_Types.repl_stdin); + FStar_Interactive_Ide_Types.repl_names = + (st.FStar_Interactive_Ide_Types.repl_names); + FStar_Interactive_Ide_Types.repl_buffered_input_queries = + (st.FStar_Interactive_Ide_Types.repl_buffered_input_queries) } in let uu___1 = run_repl_ld_transactions st1 tasks write_repl_ld_task_progress in @@ -1291,28 +1407,28 @@ let (write_full_buffer_fragment_progress : json_of_response qid FStar_Interactive_Ide_Types.QueryNOK uu___1 in FStar_Interactive_JsonHelper.write_json uu___ let (run_push_without_deps : - FStar_Interactive_ReplState.repl_state -> + FStar_Interactive_Ide_Types.repl_state -> FStar_Interactive_Ide_Types.push_query -> ((FStar_Interactive_Ide_Types.query_status * FStar_Compiler_Util.json) - * (FStar_Interactive_ReplState.repl_state, Prims.int) + * (FStar_Interactive_Ide_Types.repl_state, Prims.int) FStar_Pervasives.either)) = fun st -> fun query -> let set_nosynth_flag st1 flag = { - FStar_Interactive_ReplState.repl_line = - (st1.FStar_Interactive_ReplState.repl_line); - FStar_Interactive_ReplState.repl_column = - (st1.FStar_Interactive_ReplState.repl_column); - FStar_Interactive_ReplState.repl_fname = - (st1.FStar_Interactive_ReplState.repl_fname); - FStar_Interactive_ReplState.repl_deps_stack = - (st1.FStar_Interactive_ReplState.repl_deps_stack); - FStar_Interactive_ReplState.repl_curmod = - (st1.FStar_Interactive_ReplState.repl_curmod); - FStar_Interactive_ReplState.repl_env = - (let uu___ = st1.FStar_Interactive_ReplState.repl_env in + FStar_Interactive_Ide_Types.repl_line = + (st1.FStar_Interactive_Ide_Types.repl_line); + FStar_Interactive_Ide_Types.repl_column = + (st1.FStar_Interactive_Ide_Types.repl_column); + FStar_Interactive_Ide_Types.repl_fname = + (st1.FStar_Interactive_Ide_Types.repl_fname); + FStar_Interactive_Ide_Types.repl_deps_stack = + (st1.FStar_Interactive_Ide_Types.repl_deps_stack); + FStar_Interactive_Ide_Types.repl_curmod = + (st1.FStar_Interactive_Ide_Types.repl_curmod); + FStar_Interactive_Ide_Types.repl_env = + (let uu___ = st1.FStar_Interactive_Ide_Types.repl_env in { FStar_TypeChecker_Env.solver = (uu___.FStar_TypeChecker_Env.solver); @@ -1412,10 +1528,12 @@ let (run_push_without_deps : FStar_TypeChecker_Env.core_check = (uu___.FStar_TypeChecker_Env.core_check) }); - FStar_Interactive_ReplState.repl_stdin = - (st1.FStar_Interactive_ReplState.repl_stdin); - FStar_Interactive_ReplState.repl_names = - (st1.FStar_Interactive_ReplState.repl_names) + FStar_Interactive_Ide_Types.repl_stdin = + (st1.FStar_Interactive_Ide_Types.repl_stdin); + FStar_Interactive_Ide_Types.repl_names = + (st1.FStar_Interactive_Ide_Types.repl_names); + FStar_Interactive_Ide_Types.repl_buffered_input_queries = + (st1.FStar_Interactive_Ide_Types.repl_buffered_input_queries) } in let uu___ = query in match uu___ with @@ -1428,10 +1546,10 @@ let (run_push_without_deps : if uu___2 then FStar_TypeChecker_Env.toggle_id_info - st.FStar_Interactive_ReplState.repl_env false + st.FStar_Interactive_Ide_Types.repl_env false else FStar_TypeChecker_Env.toggle_id_info - st.FStar_Interactive_ReplState.repl_env true); + st.FStar_Interactive_Ide_Types.repl_env true); (let frag = match code_or_decl with | FStar_Pervasives.Inl text -> @@ -1446,7 +1564,7 @@ let (run_push_without_deps : let st1 = set_nosynth_flag st peek_only in let uu___2 = run_repl_transaction st1 push_kind peek_only - (FStar_Interactive_ReplState.PushFragment frag) in + (FStar_Interactive_Ide_Types.PushFragment frag) in match uu___2 with | (success, st2) -> let st3 = set_nosynth_flag st2 false in @@ -1482,28 +1600,31 @@ let (run_push_without_deps : if success then { - FStar_Interactive_ReplState.repl_line = line; - FStar_Interactive_ReplState.repl_column = column; - FStar_Interactive_ReplState.repl_fname = - (st3.FStar_Interactive_ReplState.repl_fname); - FStar_Interactive_ReplState.repl_deps_stack = - (st3.FStar_Interactive_ReplState.repl_deps_stack); - FStar_Interactive_ReplState.repl_curmod = - (st3.FStar_Interactive_ReplState.repl_curmod); - FStar_Interactive_ReplState.repl_env = - (st3.FStar_Interactive_ReplState.repl_env); - FStar_Interactive_ReplState.repl_stdin = - (st3.FStar_Interactive_ReplState.repl_stdin); - FStar_Interactive_ReplState.repl_names = - (st3.FStar_Interactive_ReplState.repl_names) + FStar_Interactive_Ide_Types.repl_line = line; + FStar_Interactive_Ide_Types.repl_column = column; + FStar_Interactive_Ide_Types.repl_fname = + (st3.FStar_Interactive_Ide_Types.repl_fname); + FStar_Interactive_Ide_Types.repl_deps_stack = + (st3.FStar_Interactive_Ide_Types.repl_deps_stack); + FStar_Interactive_Ide_Types.repl_curmod = + (st3.FStar_Interactive_Ide_Types.repl_curmod); + FStar_Interactive_Ide_Types.repl_env = + (st3.FStar_Interactive_Ide_Types.repl_env); + FStar_Interactive_Ide_Types.repl_stdin = + (st3.FStar_Interactive_Ide_Types.repl_stdin); + FStar_Interactive_Ide_Types.repl_names = + (st3.FStar_Interactive_Ide_Types.repl_names); + FStar_Interactive_Ide_Types.repl_buffered_input_queries + = + (st3.FStar_Interactive_Ide_Types.repl_buffered_input_queries) } else st3 in ((status, json_errors), (FStar_Pervasives.Inl st4)))))) let (run_push_with_deps : - FStar_Interactive_ReplState.repl_state -> + FStar_Interactive_Ide_Types.repl_state -> FStar_Interactive_Ide_Types.push_query -> ((FStar_Interactive_Ide_Types.query_status * FStar_Compiler_Util.json) - * (FStar_Interactive_ReplState.repl_state, Prims.int) + * (FStar_Interactive_Ide_Types.repl_state, Prims.int) FStar_Pervasives.either)) = fun st -> @@ -1513,7 +1634,7 @@ let (run_push_with_deps : then FStar_Compiler_Util.print_string "Reloading dependencies" else ()); FStar_TypeChecker_Env.toggle_id_info - st.FStar_Interactive_ReplState.repl_env false; + st.FStar_Interactive_Ide_Types.repl_env false; (let uu___2 = load_deps st in match uu___2 with | FStar_Pervasives.Inr st1 -> @@ -1532,31 +1653,33 @@ let (run_push_with_deps : FStar_Compiler_Effect.op_Bar_Greater uu___4 (fun uu___5 -> ())); (let names = FStar_Interactive_PushHelper.add_module_completions - st1.FStar_Interactive_ReplState.repl_fname deps - st1.FStar_Interactive_ReplState.repl_names in + st1.FStar_Interactive_Ide_Types.repl_fname deps + st1.FStar_Interactive_Ide_Types.repl_names in run_push_without_deps { - FStar_Interactive_ReplState.repl_line = - (st1.FStar_Interactive_ReplState.repl_line); - FStar_Interactive_ReplState.repl_column = - (st1.FStar_Interactive_ReplState.repl_column); - FStar_Interactive_ReplState.repl_fname = - (st1.FStar_Interactive_ReplState.repl_fname); - FStar_Interactive_ReplState.repl_deps_stack = - (st1.FStar_Interactive_ReplState.repl_deps_stack); - FStar_Interactive_ReplState.repl_curmod = - (st1.FStar_Interactive_ReplState.repl_curmod); - FStar_Interactive_ReplState.repl_env = - (st1.FStar_Interactive_ReplState.repl_env); - FStar_Interactive_ReplState.repl_stdin = - (st1.FStar_Interactive_ReplState.repl_stdin); - FStar_Interactive_ReplState.repl_names = names + FStar_Interactive_Ide_Types.repl_line = + (st1.FStar_Interactive_Ide_Types.repl_line); + FStar_Interactive_Ide_Types.repl_column = + (st1.FStar_Interactive_Ide_Types.repl_column); + FStar_Interactive_Ide_Types.repl_fname = + (st1.FStar_Interactive_Ide_Types.repl_fname); + FStar_Interactive_Ide_Types.repl_deps_stack = + (st1.FStar_Interactive_Ide_Types.repl_deps_stack); + FStar_Interactive_Ide_Types.repl_curmod = + (st1.FStar_Interactive_Ide_Types.repl_curmod); + FStar_Interactive_Ide_Types.repl_env = + (st1.FStar_Interactive_Ide_Types.repl_env); + FStar_Interactive_Ide_Types.repl_stdin = + (st1.FStar_Interactive_Ide_Types.repl_stdin); + FStar_Interactive_Ide_Types.repl_names = names; + FStar_Interactive_Ide_Types.repl_buffered_input_queries = + (st1.FStar_Interactive_Ide_Types.repl_buffered_input_queries) } query))) let (run_push : - FStar_Interactive_ReplState.repl_state -> + FStar_Interactive_Ide_Types.repl_state -> FStar_Interactive_Ide_Types.push_query -> ((FStar_Interactive_Ide_Types.query_status * FStar_Compiler_Util.json) - * (FStar_Interactive_ReplState.repl_state, Prims.int) + * (FStar_Interactive_Ide_Types.repl_state, Prims.int) FStar_Pervasives.either)) = fun st -> @@ -1566,7 +1689,7 @@ let (run_push : then run_push_with_deps st query else run_push_without_deps st query let (run_symbol_lookup : - FStar_Interactive_ReplState.repl_state -> + FStar_Interactive_Ide_Types.repl_state -> Prims.string -> FStar_Interactive_QueryHelper.position FStar_Pervasives_Native.option -> @@ -1584,7 +1707,7 @@ let (run_symbol_lookup : fun symbol_range_opt -> let uu___ = FStar_Interactive_QueryHelper.symlookup - st.FStar_Interactive_ReplState.repl_env symbol pos_opt + st.FStar_Interactive_Ide_Types.repl_env symbol pos_opt requested_info in match uu___ with | FStar_Pervasives_Native.None -> @@ -1620,7 +1743,7 @@ let (run_option_lookup : ("option", uu___4) in FStar_Pervasives.Inr uu___3) let (run_module_lookup : - FStar_Interactive_ReplState.repl_state -> + FStar_Interactive_Ide_Types.repl_state -> Prims.string -> (Prims.string, (Prims.string * (Prims.string * FStar_Compiler_Util.json) Prims.list)) @@ -1631,7 +1754,7 @@ let (run_module_lookup : let query = FStar_Compiler_Util.split symbol "." in let uu___ = FStar_Interactive_CompletionTable.find_module_or_ns - st.FStar_Interactive_ReplState.repl_names query in + st.FStar_Interactive_Ide_Types.repl_names query in match uu___ with | FStar_Pervasives_Native.None -> FStar_Pervasives.Inl "No such module or namespace" @@ -1650,7 +1773,7 @@ let (run_module_lookup : ("namespace", uu___2) in FStar_Pervasives.Inr uu___1 let (run_code_lookup : - FStar_Interactive_ReplState.repl_state -> + FStar_Interactive_Ide_Types.repl_state -> Prims.string -> FStar_Interactive_QueryHelper.position FStar_Pervasives_Native.option -> @@ -1678,7 +1801,7 @@ let (run_code_lookup : FStar_Pervasives.Inl "No such symbol, module, or namespace.") let (run_lookup' : - FStar_Interactive_ReplState.repl_state -> + FStar_Interactive_Ide_Types.repl_state -> Prims.string -> FStar_Interactive_Ide_Types.lookup_context -> FStar_Interactive_QueryHelper.position FStar_Pervasives_Native.option @@ -1707,7 +1830,7 @@ let (run_lookup' : run_code_lookup st symbol pos_opt requested_info symrange let run_lookup : 'uuuuu . - FStar_Interactive_ReplState.repl_state -> + FStar_Interactive_Ide_Types.repl_state -> Prims.string -> FStar_Interactive_Ide_Types.lookup_context -> FStar_Interactive_QueryHelper.position @@ -1716,7 +1839,7 @@ let run_lookup : FStar_Compiler_Util.json FStar_Pervasives_Native.option -> ((FStar_Interactive_Ide_Types.query_status * FStar_Compiler_Util.json) * - (FStar_Interactive_ReplState.repl_state, 'uuuuu) + (FStar_Interactive_Ide_Types.repl_state, 'uuuuu) FStar_Pervasives.either) = fun st -> @@ -1739,11 +1862,11 @@ let run_lookup : info))), (FStar_Pervasives.Inl st)) let run_code_autocomplete : 'uuuuu . - FStar_Interactive_ReplState.repl_state -> + FStar_Interactive_Ide_Types.repl_state -> Prims.string -> ((FStar_Interactive_Ide_Types.query_status * FStar_Compiler_Util.json) * - (FStar_Interactive_ReplState.repl_state, 'uuuuu) + (FStar_Interactive_Ide_Types.repl_state, 'uuuuu) FStar_Pervasives.either) = fun st -> @@ -1756,13 +1879,13 @@ let run_code_autocomplete : (FStar_Compiler_Util.JsonList js)), (FStar_Pervasives.Inl st)) let run_module_autocomplete : 'uuuuu 'uuuuu1 'uuuuu2 . - FStar_Interactive_ReplState.repl_state -> + FStar_Interactive_Ide_Types.repl_state -> Prims.string -> 'uuuuu -> 'uuuuu1 -> ((FStar_Interactive_Ide_Types.query_status * FStar_Compiler_Util.json) * - (FStar_Interactive_ReplState.repl_state, 'uuuuu2) + (FStar_Interactive_Ide_Types.repl_state, 'uuuuu2) FStar_Pervasives.either) = fun st -> @@ -1772,7 +1895,7 @@ let run_module_autocomplete : let needle = FStar_Compiler_Util.split search_term "." in let mods_and_nss = FStar_Interactive_CompletionTable.autocomplete_mod_or_ns - st.FStar_Interactive_ReplState.repl_names needle + st.FStar_Interactive_Ide_Types.repl_names needle (fun uu___ -> FStar_Pervasives_Native.Some uu___) in let json = FStar_Compiler_List.map @@ -1851,12 +1974,12 @@ let run_option_autocomplete : (FStar_Pervasives.Inl st)) let run_autocomplete : 'uuuuu . - FStar_Interactive_ReplState.repl_state -> + FStar_Interactive_Ide_Types.repl_state -> Prims.string -> FStar_Interactive_Ide_Types.completion_context -> ((FStar_Interactive_Ide_Types.query_status * FStar_Compiler_Util.json) * - (FStar_Interactive_ReplState.repl_state, 'uuuuu) + (FStar_Interactive_Ide_Types.repl_state, 'uuuuu) FStar_Pervasives.either) = fun st -> @@ -1872,10 +1995,10 @@ let run_autocomplete : run_module_autocomplete st search_term modules namespaces let run_and_rewind : 'uuuuu 'uuuuu1 . - FStar_Interactive_ReplState.repl_state -> + FStar_Interactive_Ide_Types.repl_state -> 'uuuuu -> - (FStar_Interactive_ReplState.repl_state -> 'uuuuu) -> - ('uuuuu * (FStar_Interactive_ReplState.repl_state, 'uuuuu1) + (FStar_Interactive_Ide_Types.repl_state -> 'uuuuu) -> + ('uuuuu * (FStar_Interactive_Ide_Types.repl_state, 'uuuuu1) FStar_Pervasives.either) = fun st -> @@ -1883,8 +2006,8 @@ let run_and_rewind : fun task -> let st1 = FStar_Interactive_PushHelper.push_repl "run_and_rewind" - FStar_Interactive_PushHelper.FullCheck - FStar_Interactive_ReplState.Noop st in + FStar_Interactive_Ide_Types.FullCheck + FStar_Interactive_Ide_Types.Noop st in let results = try (fun uu___ -> @@ -1907,7 +2030,7 @@ let run_and_rewind : | FStar_Pervasives.Inr e -> FStar_Compiler_Effect.raise e let run_with_parsed_and_tc_term : 'uuuuu 'uuuuu1 'uuuuu2 . - FStar_Interactive_ReplState.repl_state -> + FStar_Interactive_Ide_Types.repl_state -> Prims.string -> 'uuuuu -> 'uuuuu1 -> @@ -1918,7 +2041,7 @@ let run_with_parsed_and_tc_term : -> ((FStar_Interactive_Ide_Types.query_status * FStar_Compiler_Util.json) * - (FStar_Interactive_ReplState.repl_state, 'uuuuu2) + (FStar_Interactive_Ide_Types.repl_state, 'uuuuu2) FStar_Pervasives.either) = fun st -> @@ -1976,7 +2099,7 @@ let run_with_parsed_and_tc_term : (FStar_Interactive_Ide_Types.QueryNOK, (FStar_Compiler_Util.JsonStr "Computation interrupted")) (fun st1 -> - let tcenv = st1.FStar_Interactive_ReplState.repl_env in + let tcenv = st1.FStar_Interactive_Ide_Types.repl_env in let frag = dummy_let_fragment term in let uu___ = parse frag in match uu___ with @@ -2022,13 +2145,13 @@ let run_with_parsed_and_tc_term : FStar_Compiler_Effect.raise uu___3))) let run_compute : 'uuuuu . - FStar_Interactive_ReplState.repl_state -> + FStar_Interactive_Ide_Types.repl_state -> Prims.string -> FStar_TypeChecker_Env.step Prims.list FStar_Pervasives_Native.option -> ((FStar_Interactive_Ide_Types.query_status * FStar_Compiler_Util.json) * - (FStar_Interactive_ReplState.repl_state, 'uuuuu) + (FStar_Interactive_Ide_Types.repl_state, 'uuuuu) FStar_Pervasives.either) = fun st -> @@ -2181,16 +2304,16 @@ let (__proj__InvalidSearch__item__uu___ : Prims.exn -> Prims.string) = fun projectee -> match projectee with | InvalidSearch uu___ -> uu___ let run_search : 'uuuuu . - FStar_Interactive_ReplState.repl_state -> + FStar_Interactive_Ide_Types.repl_state -> Prims.string -> ((FStar_Interactive_Ide_Types.query_status * FStar_Compiler_Util.json) * - (FStar_Interactive_ReplState.repl_state, 'uuuuu) + (FStar_Interactive_Ide_Types.repl_state, 'uuuuu) FStar_Pervasives.either) = fun st -> fun search_str -> - let tcenv = st.FStar_Interactive_ReplState.repl_env in + let tcenv = st.FStar_Interactive_Ide_Types.repl_env in let empty_fv_set = FStar_Syntax_Syntax.new_fv_set () in let st_matches candidate term = let found = @@ -2301,11 +2424,11 @@ let run_search : (results, (FStar_Pervasives.Inl st)) let run_format_code : 'uuuuu . - FStar_Interactive_ReplState.repl_state -> + FStar_Interactive_Ide_Types.repl_state -> Prims.string -> ((FStar_Interactive_Ide_Types.query_status * FStar_Compiler_Util.json) * - (FStar_Interactive_ReplState.repl_state, 'uuuuu) + (FStar_Interactive_Ide_Types.repl_state, 'uuuuu) FStar_Pervasives.either) = fun st -> @@ -2334,23 +2457,23 @@ let run_format_code : (FStar_Pervasives.Inl st)) let (as_json_list : ((FStar_Interactive_Ide_Types.query_status * FStar_Compiler_Util.json) * - (FStar_Interactive_ReplState.repl_state, Prims.int) + (FStar_Interactive_Ide_Types.repl_state, Prims.int) FStar_Pervasives.either) -> ((FStar_Interactive_Ide_Types.query_status * FStar_Compiler_Util.json - Prims.list) * (FStar_Interactive_ReplState.repl_state, Prims.int) + Prims.list) * (FStar_Interactive_Ide_Types.repl_state, Prims.int) FStar_Pervasives.either)) = fun q -> let uu___ = q in match uu___ with | ((q1, j), s) -> ((q1, [j]), s) type run_query_result = ((FStar_Interactive_Ide_Types.query_status * FStar_Compiler_Util.json - Prims.list) * (FStar_Interactive_ReplState.repl_state, Prims.int) + Prims.list) * (FStar_Interactive_Ide_Types.repl_state, Prims.int) FStar_Pervasives.either) let rec (fold_query : - (FStar_Interactive_ReplState.repl_state -> + (FStar_Interactive_Ide_Types.repl_state -> FStar_Interactive_Ide_Types.query -> run_query_result) -> FStar_Interactive_Ide_Types.query Prims.list -> - FStar_Interactive_ReplState.repl_state -> + FStar_Interactive_Ide_Types.repl_state -> FStar_Compiler_Util.json Prims.list -> run_query_result) = fun f -> @@ -2369,10 +2492,11 @@ let rec (fold_query : (match (status, st') with | (FStar_Interactive_Ide_Types.QueryOK, FStar_Pervasives.Inl st1) -> - fold_query f l1 st1 responses1 + let st2 = buffer_input_queries st1 in + fold_query f l1 st2 responses1 | uu___1 -> ((status, responses1), st'))) let (validate_query : - FStar_Interactive_ReplState.repl_state -> + FStar_Interactive_Ide_Types.repl_state -> FStar_Interactive_Ide_Types.query -> FStar_Interactive_Ide_Types.query) = fun st -> @@ -2381,7 +2505,7 @@ let (validate_query : | FStar_Interactive_Ide_Types.Push { FStar_Interactive_Ide_Types.push_kind = - FStar_Interactive_PushHelper.SyntaxCheck; + FStar_Interactive_Ide_Types.SyntaxCheck; FStar_Interactive_Ide_Types.push_line = uu___; FStar_Interactive_Ide_Types.push_column = uu___1; FStar_Interactive_Ide_Types.push_peek_only = false; @@ -2395,7 +2519,7 @@ let (validate_query : (q.FStar_Interactive_Ide_Types.qid) } | uu___ -> - (match st.FStar_Interactive_ReplState.repl_curmod with + (match st.FStar_Interactive_Ide_Types.repl_curmod with | FStar_Pervasives_Native.None when FStar_Interactive_Ide_Types.query_needs_current_module q.FStar_Interactive_Ide_Types.qq @@ -2409,10 +2533,10 @@ let (validate_query : } | uu___1 -> q) let rec (run_query : - FStar_Interactive_ReplState.repl_state -> + FStar_Interactive_Ide_Types.repl_state -> FStar_Interactive_Ide_Types.query -> ((FStar_Interactive_Ide_Types.query_status * FStar_Compiler_Util.json - Prims.list) * (FStar_Interactive_ReplState.repl_state, Prims.int) + Prims.list) * (FStar_Interactive_Ide_Types.repl_state, Prims.int) FStar_Pervasives.either)) = fun st -> @@ -2456,7 +2580,7 @@ let rec (run_query : | FStar_Interactive_Ide_Types.Format code -> let uu___ = run_format_code st code in as_json_list uu___ and (validate_and_run_query : - FStar_Interactive_ReplState.repl_state -> + FStar_Interactive_Ide_Types.repl_state -> FStar_Interactive_Ide_Types.query -> run_query_result) = fun st -> @@ -2474,10 +2598,10 @@ and (validate_and_run_query : else ()); run_query st query1 let (js_repl_eval : - FStar_Interactive_ReplState.repl_state -> + FStar_Interactive_Ide_Types.repl_state -> FStar_Interactive_Ide_Types.query -> (FStar_Compiler_Util.json Prims.list * - (FStar_Interactive_ReplState.repl_state, Prims.int) + (FStar_Interactive_Ide_Types.repl_state, Prims.int) FStar_Pervasives.either)) = fun st -> @@ -2491,10 +2615,10 @@ let (js_repl_eval : responses in (js_responses, st_opt) let (js_repl_eval_js : - FStar_Interactive_ReplState.repl_state -> + FStar_Interactive_Ide_Types.repl_state -> FStar_Compiler_Util.json -> (FStar_Compiler_Util.json Prims.list * - (FStar_Interactive_ReplState.repl_state, Prims.int) + (FStar_Interactive_Ide_Types.repl_state, Prims.int) FStar_Pervasives.either)) = fun st -> @@ -2502,9 +2626,9 @@ let (js_repl_eval_js : let uu___ = deserialize_interactive_query query_js in js_repl_eval st uu___ let (js_repl_eval_str : - FStar_Interactive_ReplState.repl_state -> + FStar_Interactive_Ide_Types.repl_state -> Prims.string -> - (Prims.string Prims.list * (FStar_Interactive_ReplState.repl_state, + (Prims.string Prims.list * (FStar_Interactive_Ide_Types.repl_state, Prims.int) FStar_Pervasives.either)) = fun st -> @@ -2536,7 +2660,7 @@ let (js_repl_init_opts : unit -> unit) = failwith "repl_init: Too many file names given in --ide invocation" | uu___2 -> ())) -let rec (go : FStar_Interactive_ReplState.repl_state -> Prims.int) = +let rec (go : FStar_Interactive_Ide_Types.repl_state -> Prims.int) = fun st -> (let uu___1 = FStar_Options.debug_any () in if uu___1 @@ -2545,20 +2669,21 @@ let rec (go : FStar_Interactive_ReplState.repl_state -> Prims.int) = let uu___3 = FStar_Compiler_Effect.op_Bang FStar_Interactive_PushHelper.repl_stack in - FStar_Interactive_ReplState.string_of_repl_stack uu___3 in + FStar_Interactive_Ide_Types.string_of_repl_stack uu___3 in FStar_Compiler_Util.print1 "Repl stack is:\n%s\n" uu___2 else ()); - (let query = - read_interactive_query st.FStar_Interactive_ReplState.repl_stdin in - let uu___1 = validate_and_run_query st query in + (let uu___1 = read_interactive_query st in match uu___1 with - | ((status, responses), state_opt) -> - (FStar_Compiler_List.iter - (write_response query.FStar_Interactive_Ide_Types.qid status) - responses; - (match state_opt with - | FStar_Pervasives.Inl st' -> go st' - | FStar_Pervasives.Inr exitcode -> exitcode))) + | (query, st1) -> + let uu___2 = validate_and_run_query st1 query in + (match uu___2 with + | ((status, responses), state_opt) -> + (FStar_Compiler_List.iter + (write_response query.FStar_Interactive_Ide_Types.qid status) + responses; + (match state_opt with + | FStar_Pervasives.Inl st' -> go st' + | FStar_Pervasives.Inr exitcode -> exitcode)))) let (interactive_error_handler : FStar_Errors.error_handler) = let issues = FStar_Compiler_Util.mk_ref [] in let add_one e = @@ -2609,7 +2734,7 @@ let (install_ide_mode_hooks : (FStar_Compiler_Util.json -> unit) -> unit) = FStar_Compiler_Util.set_printer (interactive_printer printer); FStar_Errors.set_handler interactive_error_handler let (build_initial_repl_state : - Prims.string -> FStar_Interactive_ReplState.repl_state) = + Prims.string -> FStar_Interactive_Ide_Types.repl_state) = fun filename -> let env = FStar_Universal.init_env FStar_Parser_Dep.empty_deps in let env1 = @@ -2617,18 +2742,19 @@ let (build_initial_repl_state : FStar_Interactive_Ide_Types.initial_range in let uu___ = FStar_Compiler_Util.open_stdin () in { - FStar_Interactive_ReplState.repl_line = Prims.int_one; - FStar_Interactive_ReplState.repl_column = Prims.int_zero; - FStar_Interactive_ReplState.repl_fname = filename; - FStar_Interactive_ReplState.repl_deps_stack = []; - FStar_Interactive_ReplState.repl_curmod = FStar_Pervasives_Native.None; - FStar_Interactive_ReplState.repl_env = env1; - FStar_Interactive_ReplState.repl_stdin = uu___; - FStar_Interactive_ReplState.repl_names = - FStar_Interactive_CompletionTable.empty + FStar_Interactive_Ide_Types.repl_line = Prims.int_one; + FStar_Interactive_Ide_Types.repl_column = Prims.int_zero; + FStar_Interactive_Ide_Types.repl_fname = filename; + FStar_Interactive_Ide_Types.repl_deps_stack = []; + FStar_Interactive_Ide_Types.repl_curmod = FStar_Pervasives_Native.None; + FStar_Interactive_Ide_Types.repl_env = env1; + FStar_Interactive_Ide_Types.repl_stdin = uu___; + FStar_Interactive_Ide_Types.repl_names = + FStar_Interactive_CompletionTable.empty; + FStar_Interactive_Ide_Types.repl_buffered_input_queries = [] } let interactive_mode' : - 'uuuuu . FStar_Interactive_ReplState.repl_state -> 'uuuuu = + 'uuuuu . FStar_Interactive_Ide_Types.repl_state -> 'uuuuu = fun init_st -> write_hello (); (let exit_code = diff --git a/ocaml/fstar-lib/generated/FStar_Interactive_Ide_Types.ml b/ocaml/fstar-lib/generated/FStar_Interactive_Ide_Types.ml index cbaf6d19e93..88ecda335bb 100644 --- a/ocaml/fstar-lib/generated/FStar_Interactive_Ide_Types.ml +++ b/ocaml/fstar-lib/generated/FStar_Interactive_Ide_Types.ml @@ -3,28 +3,6 @@ let (initial_range : FStar_Compiler_Range.range) = let uu___ = FStar_Compiler_Range.mk_pos Prims.int_one Prims.int_zero in let uu___1 = FStar_Compiler_Range.mk_pos Prims.int_one Prims.int_zero in FStar_Compiler_Range.mk_range "" uu___ uu___1 -let (js_pushkind : - FStar_Compiler_Util.json -> FStar_Interactive_PushHelper.push_kind) = - fun s -> - let uu___ = FStar_Interactive_JsonHelper.js_str s in - match uu___ with - | "syntax" -> FStar_Interactive_PushHelper.SyntaxCheck - | "lax" -> FStar_Interactive_PushHelper.LaxCheck - | "full" -> FStar_Interactive_PushHelper.FullCheck - | uu___1 -> FStar_Interactive_JsonHelper.js_fail "push_kind" s -let (js_reductionrule : - FStar_Compiler_Util.json -> FStar_TypeChecker_Env.step) = - fun s -> - let uu___ = FStar_Interactive_JsonHelper.js_str s in - match uu___ with - | "beta" -> FStar_TypeChecker_Env.Beta - | "delta" -> - FStar_TypeChecker_Env.UnfoldUntil FStar_Syntax_Syntax.delta_constant - | "iota" -> FStar_TypeChecker_Env.Iota - | "zeta" -> FStar_TypeChecker_Env.Zeta - | "reify" -> FStar_TypeChecker_Env.Reify - | "pure-subterms" -> FStar_TypeChecker_Env.PureSubtermsWithinComputations - | uu___1 -> FStar_Interactive_JsonHelper.js_fail "reduction rule" s type completion_context = | CKCode | CKOption of Prims.bool @@ -42,28 +20,6 @@ let (uu___is_CKModuleOrNamespace : completion_context -> Prims.bool) = let (__proj__CKModuleOrNamespace__item___0 : completion_context -> (Prims.bool * Prims.bool)) = fun projectee -> match projectee with | CKModuleOrNamespace _0 -> _0 -let (js_optional_completion_context : - FStar_Compiler_Util.json FStar_Pervasives_Native.option -> - completion_context) - = - fun k -> - match k with - | FStar_Pervasives_Native.None -> CKCode - | FStar_Pervasives_Native.Some k1 -> - let uu___ = FStar_Interactive_JsonHelper.js_str k1 in - (match uu___ with - | "symbol" -> CKCode - | "code" -> CKCode - | "set-options" -> CKOption false - | "reset-options" -> CKOption true - | "open" -> CKModuleOrNamespace (true, true) - | "let-open" -> CKModuleOrNamespace (true, true) - | "include" -> CKModuleOrNamespace (true, false) - | "module-alias" -> CKModuleOrNamespace (true, false) - | uu___1 -> - FStar_Interactive_JsonHelper.js_fail - "completion context (code, set-options, reset-options, open, let-open, include, module-alias)" - k1) type lookup_context = | LKSymbolOnly | LKModule @@ -78,63 +34,52 @@ let (uu___is_LKOption : lookup_context -> Prims.bool) = fun projectee -> match projectee with | LKOption -> true | uu___ -> false let (uu___is_LKCode : lookup_context -> Prims.bool) = fun projectee -> match projectee with | LKCode -> true | uu___ -> false -let (js_optional_lookup_context : - FStar_Compiler_Util.json FStar_Pervasives_Native.option -> lookup_context) - = - fun k -> - match k with - | FStar_Pervasives_Native.None -> LKSymbolOnly - | FStar_Pervasives_Native.Some k1 -> - let uu___ = FStar_Interactive_JsonHelper.js_str k1 in - (match uu___ with - | "symbol-only" -> LKSymbolOnly - | "code" -> LKCode - | "set-options" -> LKOption - | "reset-options" -> LKOption - | "open" -> LKModule - | "let-open" -> LKModule - | "include" -> LKModule - | "module-alias" -> LKModule - | uu___1 -> - FStar_Interactive_JsonHelper.js_fail - "lookup context (symbol-only, code, set-options, reset-options, open, let-open, include, module-alias)" - k1) type position = (Prims.string * Prims.int * Prims.int) +type push_kind = + | SyntaxCheck + | LaxCheck + | FullCheck +let (uu___is_SyntaxCheck : push_kind -> Prims.bool) = + fun projectee -> + match projectee with | SyntaxCheck -> true | uu___ -> false +let (uu___is_LaxCheck : push_kind -> Prims.bool) = + fun projectee -> match projectee with | LaxCheck -> true | uu___ -> false +let (uu___is_FullCheck : push_kind -> Prims.bool) = + fun projectee -> match projectee with | FullCheck -> true | uu___ -> false type push_query = { - push_kind: FStar_Interactive_PushHelper.push_kind ; + push_kind: push_kind ; push_line: Prims.int ; push_column: Prims.int ; push_peek_only: Prims.bool ; push_code_or_decl: (Prims.string, FStar_Parser_AST.decl) FStar_Pervasives.either } -let (__proj__Mkpush_query__item__push_kind : - push_query -> FStar_Interactive_PushHelper.push_kind) = +let (__proj__Mkpush_query__item__push_kind : push_query -> push_kind) = fun projectee -> match projectee with - | { push_kind; push_line; push_column; push_peek_only; - push_code_or_decl;_} -> push_kind + | { push_kind = push_kind1; push_line; push_column; push_peek_only; + push_code_or_decl;_} -> push_kind1 let (__proj__Mkpush_query__item__push_line : push_query -> Prims.int) = fun projectee -> match projectee with - | { push_kind; push_line; push_column; push_peek_only; + | { push_kind = push_kind1; push_line; push_column; push_peek_only; push_code_or_decl;_} -> push_line let (__proj__Mkpush_query__item__push_column : push_query -> Prims.int) = fun projectee -> match projectee with - | { push_kind; push_line; push_column; push_peek_only; + | { push_kind = push_kind1; push_line; push_column; push_peek_only; push_code_or_decl;_} -> push_column let (__proj__Mkpush_query__item__push_peek_only : push_query -> Prims.bool) = fun projectee -> match projectee with - | { push_kind; push_line; push_column; push_peek_only; + | { push_kind = push_kind1; push_line; push_column; push_peek_only; push_code_or_decl;_} -> push_peek_only let (__proj__Mkpush_query__item__push_code_or_decl : push_query -> (Prims.string, FStar_Parser_AST.decl) FStar_Pervasives.either) = fun projectee -> match projectee with - | { push_kind; push_line; push_column; push_peek_only; + | { push_kind = push_kind1; push_line; push_column; push_peek_only; push_code_or_decl;_} -> push_code_or_decl type lookup_symbol_range = FStar_Compiler_Util.json type query_status = @@ -148,11 +93,64 @@ let (uu___is_QueryNOK : query_status -> Prims.bool) = let (uu___is_QueryViolatesProtocol : query_status -> Prims.bool) = fun projectee -> match projectee with | QueryViolatesProtocol -> true | uu___ -> false -type callback_t = - FStar_Interactive_ReplState.repl_state -> - ((query_status * FStar_Compiler_Util.json Prims.list) * - (FStar_Interactive_ReplState.repl_state, Prims.int) +type repl_depth_t = (FStar_TypeChecker_Env.tcenv_depth_t * Prims.int) +type optmod_t = FStar_Syntax_Syntax.modul FStar_Pervasives_Native.option +type timed_fname = + { + tf_fname: Prims.string ; + tf_modtime: FStar_Compiler_Util.time } +let (__proj__Mktimed_fname__item__tf_fname : timed_fname -> Prims.string) = + fun projectee -> + match projectee with | { tf_fname; tf_modtime;_} -> tf_fname +let (__proj__Mktimed_fname__item__tf_modtime : + timed_fname -> FStar_Compiler_Util.time) = + fun projectee -> + match projectee with | { tf_fname; tf_modtime;_} -> tf_modtime +type repl_task = + | LDInterleaved of (timed_fname * timed_fname) + | LDSingle of timed_fname + | LDInterfaceOfCurrentFile of timed_fname + | PushFragment of (FStar_Parser_ParseIt.input_frag, FStar_Parser_AST.decl) + FStar_Pervasives.either + | Noop +let (uu___is_LDInterleaved : repl_task -> Prims.bool) = + fun projectee -> + match projectee with | LDInterleaved _0 -> true | uu___ -> false +let (__proj__LDInterleaved__item___0 : + repl_task -> (timed_fname * timed_fname)) = + fun projectee -> match projectee with | LDInterleaved _0 -> _0 +let (uu___is_LDSingle : repl_task -> Prims.bool) = + fun projectee -> + match projectee with | LDSingle _0 -> true | uu___ -> false +let (__proj__LDSingle__item___0 : repl_task -> timed_fname) = + fun projectee -> match projectee with | LDSingle _0 -> _0 +let (uu___is_LDInterfaceOfCurrentFile : repl_task -> Prims.bool) = + fun projectee -> + match projectee with + | LDInterfaceOfCurrentFile _0 -> true + | uu___ -> false +let (__proj__LDInterfaceOfCurrentFile__item___0 : repl_task -> timed_fname) = + fun projectee -> match projectee with | LDInterfaceOfCurrentFile _0 -> _0 +let (uu___is_PushFragment : repl_task -> Prims.bool) = + fun projectee -> + match projectee with | PushFragment _0 -> true | uu___ -> false +let (__proj__PushFragment__item___0 : + repl_task -> + (FStar_Parser_ParseIt.input_frag, FStar_Parser_AST.decl) FStar_Pervasives.either) + = fun projectee -> match projectee with | PushFragment _0 -> _0 +let (uu___is_Noop : repl_task -> Prims.bool) = + fun projectee -> match projectee with | Noop -> true | uu___ -> false +type full_buffer_request_kind = + | Full + | Cache + | ReloadDeps +let (uu___is_Full : full_buffer_request_kind -> Prims.bool) = + fun projectee -> match projectee with | Full -> true | uu___ -> false +let (uu___is_Cache : full_buffer_request_kind -> Prims.bool) = + fun projectee -> match projectee with | Cache -> true | uu___ -> false +let (uu___is_ReloadDeps : full_buffer_request_kind -> Prims.bool) = + fun projectee -> match projectee with | ReloadDeps -> true | uu___ -> false type query' = | Exit | DescribeProtocol @@ -170,12 +168,28 @@ type query' = | Search of Prims.string | GenericError of Prims.string | ProtocolViolation of Prims.string - | FullBuffer of (Prims.string * Prims.bool) - | Callback of callback_t + | FullBuffer of (Prims.string * full_buffer_request_kind) + | Callback of + (repl_state -> + ((query_status * FStar_Compiler_Util.json Prims.list) * (repl_state, + Prims.int) FStar_Pervasives.either)) + | Format of Prims.string + | Cancel of position FStar_Pervasives_Native.option and query = { qq: query' ; qid: Prims.string } +and repl_state = + { + repl_line: Prims.int ; + repl_column: Prims.int ; + repl_fname: Prims.string ; + repl_deps_stack: (repl_depth_t * (repl_task * repl_state)) Prims.list ; + repl_curmod: optmod_t ; + repl_env: FStar_TypeChecker_Env.env ; + repl_stdin: FStar_Compiler_Util.stream_reader ; + repl_names: FStar_Interactive_CompletionTable.table ; + repl_buffered_input_queries: query Prims.list } let (uu___is_Exit : query' -> Prims.bool) = fun projectee -> match projectee with | Exit -> true | uu___ -> false let (uu___is_DescribeProtocol : query' -> Prims.bool) = @@ -237,28 +251,182 @@ let (__proj__ProtocolViolation__item___0 : query' -> Prims.string) = let (uu___is_FullBuffer : query' -> Prims.bool) = fun projectee -> match projectee with | FullBuffer _0 -> true | uu___ -> false -let (__proj__FullBuffer__item___0 : query' -> (Prims.string * Prims.bool)) = +let (__proj__FullBuffer__item___0 : + query' -> (Prims.string * full_buffer_request_kind)) = fun projectee -> match projectee with | FullBuffer _0 -> _0 let (uu___is_Callback : query' -> Prims.bool) = fun projectee -> match projectee with | Callback _0 -> true | uu___ -> false -let (__proj__Callback__item___0 : query' -> callback_t) = - fun projectee -> match projectee with | Callback _0 -> _0 +let (__proj__Callback__item___0 : + query' -> + repl_state -> + ((query_status * FStar_Compiler_Util.json Prims.list) * (repl_state, + Prims.int) FStar_Pervasives.either)) + = fun projectee -> match projectee with | Callback _0 -> _0 let (uu___is_Format : query' -> Prims.bool) = fun projectee -> match projectee with | Format _0 -> true | uu___ -> false let (__proj__Format__item___0 : query' -> Prims.string) = fun projectee -> match projectee with | Format _0 -> _0 +let (uu___is_Cancel : query' -> Prims.bool) = + fun projectee -> match projectee with | Cancel _0 -> true | uu___ -> false +let (__proj__Cancel__item___0 : + query' -> position FStar_Pervasives_Native.option) = + fun projectee -> match projectee with | Cancel _0 -> _0 let (__proj__Mkquery__item__qq : query -> query') = fun projectee -> match projectee with | { qq; qid;_} -> qq let (__proj__Mkquery__item__qid : query -> Prims.string) = fun projectee -> match projectee with | { qq; qid;_} -> qid +let (__proj__Mkrepl_state__item__repl_line : repl_state -> Prims.int) = + fun projectee -> + match projectee with + | { repl_line; repl_column; repl_fname; repl_deps_stack; repl_curmod; + repl_env; repl_stdin; repl_names; repl_buffered_input_queries;_} -> + repl_line +let (__proj__Mkrepl_state__item__repl_column : repl_state -> Prims.int) = + fun projectee -> + match projectee with + | { repl_line; repl_column; repl_fname; repl_deps_stack; repl_curmod; + repl_env; repl_stdin; repl_names; repl_buffered_input_queries;_} -> + repl_column +let (__proj__Mkrepl_state__item__repl_fname : repl_state -> Prims.string) = + fun projectee -> + match projectee with + | { repl_line; repl_column; repl_fname; repl_deps_stack; repl_curmod; + repl_env; repl_stdin; repl_names; repl_buffered_input_queries;_} -> + repl_fname +let (__proj__Mkrepl_state__item__repl_deps_stack : + repl_state -> (repl_depth_t * (repl_task * repl_state)) Prims.list) = + fun projectee -> + match projectee with + | { repl_line; repl_column; repl_fname; repl_deps_stack; repl_curmod; + repl_env; repl_stdin; repl_names; repl_buffered_input_queries;_} -> + repl_deps_stack +let (__proj__Mkrepl_state__item__repl_curmod : repl_state -> optmod_t) = + fun projectee -> + match projectee with + | { repl_line; repl_column; repl_fname; repl_deps_stack; repl_curmod; + repl_env; repl_stdin; repl_names; repl_buffered_input_queries;_} -> + repl_curmod +let (__proj__Mkrepl_state__item__repl_env : + repl_state -> FStar_TypeChecker_Env.env) = + fun projectee -> + match projectee with + | { repl_line; repl_column; repl_fname; repl_deps_stack; repl_curmod; + repl_env; repl_stdin; repl_names; repl_buffered_input_queries;_} -> + repl_env +let (__proj__Mkrepl_state__item__repl_stdin : + repl_state -> FStar_Compiler_Util.stream_reader) = + fun projectee -> + match projectee with + | { repl_line; repl_column; repl_fname; repl_deps_stack; repl_curmod; + repl_env; repl_stdin; repl_names; repl_buffered_input_queries;_} -> + repl_stdin +let (__proj__Mkrepl_state__item__repl_names : + repl_state -> FStar_Interactive_CompletionTable.table) = + fun projectee -> + match projectee with + | { repl_line; repl_column; repl_fname; repl_deps_stack; repl_curmod; + repl_env; repl_stdin; repl_names; repl_buffered_input_queries;_} -> + repl_names +let (__proj__Mkrepl_state__item__repl_buffered_input_queries : + repl_state -> query Prims.list) = + fun projectee -> + match projectee with + | { repl_line; repl_column; repl_fname; repl_deps_stack; repl_curmod; + repl_env; repl_stdin; repl_names; repl_buffered_input_queries;_} -> + repl_buffered_input_queries +type callback_t = + repl_state -> + ((query_status * FStar_Compiler_Util.json Prims.list) * (repl_state, + Prims.int) FStar_Pervasives.either) +type repl_stack_entry_t = (repl_depth_t * (repl_task * repl_state)) +type repl_stack_t = (repl_depth_t * (repl_task * repl_state)) Prims.list +type grepl_state = + { + grepl_repls: repl_state FStar_Compiler_Util.psmap ; + grepl_stdin: FStar_Compiler_Util.stream_reader } +let (__proj__Mkgrepl_state__item__grepl_repls : + grepl_state -> repl_state FStar_Compiler_Util.psmap) = + fun projectee -> + match projectee with | { grepl_repls; grepl_stdin;_} -> grepl_repls +let (__proj__Mkgrepl_state__item__grepl_stdin : + grepl_state -> FStar_Compiler_Util.stream_reader) = + fun projectee -> + match projectee with | { grepl_repls; grepl_stdin;_} -> grepl_stdin +let (t0 : FStar_Compiler_Util.time) = FStar_Compiler_Util.now () +let (dummy_tf_of_fname : Prims.string -> timed_fname) = + fun fname -> { tf_fname = fname; tf_modtime = t0 } +let (string_of_timed_fname : timed_fname -> Prims.string) = + fun uu___ -> + match uu___ with + | { tf_fname = fname; tf_modtime = modtime;_} -> + if modtime = t0 + then FStar_Compiler_Util.format1 "{ %s }" fname + else + (let uu___2 = FStar_Compiler_Util.string_of_time modtime in + FStar_Compiler_Util.format2 "{ %s; %s }" fname uu___2) +let (string_of_repl_task : repl_task -> Prims.string) = + fun uu___ -> + match uu___ with + | LDInterleaved (intf, impl) -> + let uu___1 = string_of_timed_fname intf in + let uu___2 = string_of_timed_fname impl in + FStar_Compiler_Util.format2 "LDInterleaved (%s, %s)" uu___1 uu___2 + | LDSingle intf_or_impl -> + let uu___1 = string_of_timed_fname intf_or_impl in + FStar_Compiler_Util.format1 "LDSingle %s" uu___1 + | LDInterfaceOfCurrentFile intf -> + let uu___1 = string_of_timed_fname intf in + FStar_Compiler_Util.format1 "LDInterfaceOfCurrentFile %s" uu___1 + | PushFragment (FStar_Pervasives.Inl frag) -> + FStar_Compiler_Util.format1 "PushFragment { code = %s }" + frag.FStar_Parser_ParseIt.frag_text + | PushFragment (FStar_Pervasives.Inr d) -> + let uu___1 = FStar_Parser_AST.decl_to_string d in + FStar_Compiler_Util.format1 "PushFragment { decl = %s }" uu___1 + | Noop -> "Noop {}" +let (string_of_repl_stack_entry : repl_stack_entry_t -> Prims.string) = + fun uu___ -> + match uu___ with + | ((depth, i), (task, state)) -> + let uu___1 = + let uu___2 = FStar_Compiler_Util.string_of_int i in + let uu___3 = let uu___4 = string_of_repl_task task in [uu___4] in + uu___2 :: uu___3 in + FStar_Compiler_Util.format "{depth=%s; task=%s}" uu___1 +let (string_of_repl_stack : repl_stack_entry_t Prims.list -> Prims.string) = + fun s -> + let uu___ = FStar_Compiler_List.map string_of_repl_stack_entry s in + FStar_String.concat ";\n\t\t" uu___ +let (repl_state_to_string : repl_state -> Prims.string) = + fun r -> + let uu___ = + let uu___1 = FStar_Compiler_Util.string_of_int r.repl_line in + let uu___2 = + let uu___3 = FStar_Compiler_Util.string_of_int r.repl_column in + let uu___4 = + let uu___5 = + let uu___6 = + match r.repl_curmod with + | FStar_Pervasives_Native.None -> "None" + | FStar_Pervasives_Native.Some m -> + FStar_Ident.string_of_lid m.FStar_Syntax_Syntax.name in + let uu___7 = + let uu___8 = string_of_repl_stack r.repl_deps_stack in [uu___8] in + uu___6 :: uu___7 in + (r.repl_fname) :: uu___5 in + uu___3 :: uu___4 in + uu___1 :: uu___2 in + FStar_Compiler_Util.format + "{\n\trepl_line=%s;\n\trepl_column=%s;\n\trepl_fname=%s;\n\trepl_cur_mod=%s;\n\t\\ \r\n repl_deps_stack={%s}\n}" + uu___ let (push_query_to_string : push_query -> Prims.string) = fun pq -> let pk = match pq.push_kind with - | FStar_Interactive_PushHelper.SyntaxCheck -> "SyntaxCheck" - | FStar_Interactive_PushHelper.LaxCheck -> "LaxCheck" - | FStar_Interactive_PushHelper.FullCheck -> "FullCheck" in + | SyntaxCheck -> "SyntaxCheck" + | LaxCheck -> "LaxCheck" + | FullCheck -> "FullCheck" in let code_or_decl = match pq.push_code_or_decl with | FStar_Pervasives.Inl code -> code @@ -299,6 +467,7 @@ let (query_to_string : query -> Prims.string) = | FullBuffer uu___ -> "FullBuffer" | Callback uu___ -> "Callback" | Format uu___ -> "Format" + | Cancel uu___ -> "Cancel" let (query_needs_current_module : query' -> Prims.bool) = fun uu___ -> match uu___ with @@ -346,7 +515,8 @@ let (interactive_protocol_features : Prims.string Prims.list) = "interrupt"; "progress"; "full-buffer"; - "format"] + "format"; + "cancel"] let (json_of_issue_level : FStar_Errors.issue_level -> FStar_Compiler_Util.json) = fun i -> @@ -400,4 +570,68 @@ let (json_of_issue : FStar_Errors.issue -> FStar_Compiler_Util.json) = [("level", (json_of_issue_level issue.FStar_Errors.issue_level))] uu___1 in FStar_Compiler_Effect.op_Less_Bar - (fun uu___1 -> FStar_Compiler_Util.JsonAssoc uu___1) uu___ \ No newline at end of file + (fun uu___1 -> FStar_Compiler_Util.JsonAssoc uu___1) uu___ +let (js_pushkind : FStar_Compiler_Util.json -> push_kind) = + fun s -> + let uu___ = FStar_Interactive_JsonHelper.js_str s in + match uu___ with + | "syntax" -> SyntaxCheck + | "lax" -> LaxCheck + | "full" -> FullCheck + | uu___1 -> FStar_Interactive_JsonHelper.js_fail "push_kind" s +let (js_reductionrule : + FStar_Compiler_Util.json -> FStar_TypeChecker_Env.step) = + fun s -> + let uu___ = FStar_Interactive_JsonHelper.js_str s in + match uu___ with + | "beta" -> FStar_TypeChecker_Env.Beta + | "delta" -> + FStar_TypeChecker_Env.UnfoldUntil FStar_Syntax_Syntax.delta_constant + | "iota" -> FStar_TypeChecker_Env.Iota + | "zeta" -> FStar_TypeChecker_Env.Zeta + | "reify" -> FStar_TypeChecker_Env.Reify + | "pure-subterms" -> FStar_TypeChecker_Env.PureSubtermsWithinComputations + | uu___1 -> FStar_Interactive_JsonHelper.js_fail "reduction rule" s +let (js_optional_completion_context : + FStar_Compiler_Util.json FStar_Pervasives_Native.option -> + completion_context) + = + fun k -> + match k with + | FStar_Pervasives_Native.None -> CKCode + | FStar_Pervasives_Native.Some k1 -> + let uu___ = FStar_Interactive_JsonHelper.js_str k1 in + (match uu___ with + | "symbol" -> CKCode + | "code" -> CKCode + | "set-options" -> CKOption false + | "reset-options" -> CKOption true + | "open" -> CKModuleOrNamespace (true, true) + | "let-open" -> CKModuleOrNamespace (true, true) + | "include" -> CKModuleOrNamespace (true, false) + | "module-alias" -> CKModuleOrNamespace (true, false) + | uu___1 -> + FStar_Interactive_JsonHelper.js_fail + "completion context (code, set-options, reset-options, open, let-open, include, module-alias)" + k1) +let (js_optional_lookup_context : + FStar_Compiler_Util.json FStar_Pervasives_Native.option -> lookup_context) + = + fun k -> + match k with + | FStar_Pervasives_Native.None -> LKSymbolOnly + | FStar_Pervasives_Native.Some k1 -> + let uu___ = FStar_Interactive_JsonHelper.js_str k1 in + (match uu___ with + | "symbol-only" -> LKSymbolOnly + | "code" -> LKCode + | "set-options" -> LKOption + | "reset-options" -> LKOption + | "open" -> LKModule + | "let-open" -> LKModule + | "include" -> LKModule + | "module-alias" -> LKModule + | uu___1 -> + FStar_Interactive_JsonHelper.js_fail + "lookup context (symbol-only, code, set-options, reset-options, open, let-open, include, module-alias)" + k1) \ No newline at end of file diff --git a/ocaml/fstar-lib/generated/FStar_Interactive_Incremental.ml b/ocaml/fstar-lib/generated/FStar_Interactive_Incremental.ml index 6304808b6d6..7effb4eefed 100644 --- a/ocaml/fstar-lib/generated/FStar_Interactive_Incremental.ml +++ b/ocaml/fstar-lib/generated/FStar_Interactive_Incremental.ml @@ -96,7 +96,7 @@ let (push_decl : FStar_Compiler_Range.col_of_pos uu___2 in { FStar_Interactive_Ide_Types.push_kind = - FStar_Interactive_PushHelper.FullCheck; + FStar_Interactive_Ide_Types.FullCheck; FStar_Interactive_Ide_Types.push_line = uu___; FStar_Interactive_Ide_Types.push_column = uu___1; FStar_Interactive_Ide_Types.push_peek_only = false; @@ -122,7 +122,7 @@ let (push_decls : let uu___ = map (push_decl write_full_buffer_fragment_progress) ds in op_let_Bang uu___ (fun qs -> return (FStar_Compiler_List.flatten qs)) let (pop_entries : - FStar_Interactive_ReplState.repl_stack_entry_t Prims.list -> + FStar_Interactive_Ide_Types.repl_stack_entry_t Prims.list -> FStar_Interactive_Ide_Types.query Prims.list qst) = fun e -> map (fun uu___ -> as_query FStar_Interactive_Ide_Types.Pop) e let (response_success : @@ -147,8 +147,11 @@ let (response_success : ("query-id", (FStar_Compiler_Util.JsonStr q)); ("status", (FStar_Compiler_Util.JsonStr "success")); ("contents", contents)])) +let repl_task : + 'uuuuu 'uuuuu1 'uuuuu2 . ('uuuuu * ('uuuuu1 * 'uuuuu2)) -> 'uuuuu1 = + fun uu___ -> match uu___ with | (uu___1, (p, uu___2)) -> p let (inspect_repl_stack : - FStar_Interactive_ReplState.repl_stack_t -> + FStar_Interactive_Ide_Types.repl_stack_t -> FStar_Parser_AST.decl Prims.list -> (fragment_progress -> unit) -> FStar_Interactive_Ide_Types.query Prims.list qst) @@ -163,7 +166,7 @@ let (inspect_repl_stack : (fun uu___1 -> match uu___1 with | (uu___2, - (FStar_Interactive_ReplState.PushFragment uu___3, uu___4)) + (FStar_Interactive_Ide_Types.PushFragment uu___3, uu___4)) -> true | uu___2 -> false) entries in match uu___ with @@ -172,16 +175,16 @@ let (inspect_repl_stack : op_let_Bang uu___1 (fun ds1 -> return ds1) | FStar_Pervasives_Native.Some (prefix, first_push, rest) -> let entries1 = first_push :: rest in - let repl_task uu___1 = + let repl_task1 uu___1 = match uu___1 with | (uu___2, (p, uu___3)) -> p in let rec matching_prefix entries2 ds1 = match (entries2, ds1) with | ([], []) -> return [] | (e::entries3, d::ds2) -> - (match repl_task e with - | FStar_Interactive_ReplState.Noop -> + (match repl_task1 e with + | FStar_Interactive_Ide_Types.Noop -> matching_prefix entries3 (d :: ds2) - | FStar_Interactive_ReplState.PushFragment + | FStar_Interactive_Ide_Types.PushFragment (FStar_Pervasives.Inl frag) -> let uu___1 = pop_entries (e :: entries3) in op_let_Bang uu___1 @@ -191,7 +194,7 @@ let (inspect_repl_stack : (fun pushes -> return (FStar_Compiler_List.op_At pops pushes))) - | FStar_Interactive_ReplState.PushFragment + | FStar_Interactive_Ide_Types.PushFragment (FStar_Pervasives.Inr d') -> let uu___1 = FStar_Parser_AST_Comparison.eq_decl d d' in if uu___1 @@ -215,6 +218,30 @@ let (inspect_repl_stack : let uu___1 = pop_entries es in op_let_Bang uu___1 (fun pops -> return pops) in matching_prefix entries1 ds +let reload_deps : + 'uuuuu 'uuuuu1 . + ('uuuuu * (FStar_Interactive_Ide_Types.repl_task * 'uuuuu1)) Prims.list + -> FStar_Interactive_Ide_Types.query Prims.list qst + = + fun repl_stack -> + let pop_until_deps entries = + let uu___ = + FStar_Compiler_Util.prefix_until + (fun e -> + match repl_task e with + | FStar_Interactive_Ide_Types.PushFragment uu___1 -> false + | FStar_Interactive_Ide_Types.Noop -> false + | uu___1 -> true) entries in + match uu___ with + | FStar_Pervasives_Native.None -> return [] + | FStar_Pervasives_Native.Some (prefix, uu___1, uu___2) -> + let uu___3 = as_query FStar_Interactive_Ide_Types.Pop in + op_let_Bang uu___3 + (fun pop -> + let uu___4 = + FStar_Compiler_List.map (fun uu___5 -> pop) prefix in + return uu___4) in + pop_until_deps repl_stack let (parse_code : Prims.string -> FStar_Parser_ParseIt.parse_result) = fun code -> let uu___ = @@ -262,17 +289,17 @@ let (syntax_issue : } in issue) let (run_full_buffer : - FStar_Interactive_ReplState.repl_state -> + FStar_Interactive_Ide_Types.repl_state -> Prims.string -> Prims.string -> - Prims.bool -> + FStar_Interactive_Ide_Types.full_buffer_request_kind -> (fragment_progress -> unit) -> FStar_Interactive_Ide_Types.query Prims.list) = fun st -> fun qid1 -> fun code -> - fun full -> + fun request_type -> fun write_full_buffer_fragment_progress -> let parse_result = parse_code code in let log_syntax_issues err = @@ -285,37 +312,77 @@ let (run_full_buffer : match parse_result with | FStar_Parser_ParseIt.IncrementalFragment (decls, uu___, err_opt) -> - let queries = - let uu___1 = - let uu___2 = - FStar_Compiler_Effect.op_Bang - FStar_Interactive_PushHelper.repl_stack in - inspect_repl_stack uu___2 decls - write_full_buffer_fragment_progress in - run_qst uu___1 qid1 in - (if full then log_syntax_issues err_opt else (); - (let uu___3 = FStar_Options.debug_any () in - if uu___3 - then - let uu___4 = - let uu___5 = - FStar_Compiler_List.map - FStar_Interactive_Ide_Types.query_to_string - queries in - FStar_String.concat "\n" uu___5 in - FStar_Compiler_Util.print1 "Generating queries\n%s\n" - uu___4 - else ()); - if full then queries else []) + (match request_type with + | FStar_Interactive_Ide_Types.Full -> + let queries = + let uu___1 = + let uu___2 = + FStar_Compiler_Effect.op_Bang + FStar_Interactive_PushHelper.repl_stack in + inspect_repl_stack uu___2 decls + write_full_buffer_fragment_progress in + run_qst uu___1 qid1 in + (if request_type = FStar_Interactive_Ide_Types.Full + then log_syntax_issues err_opt + else (); + (let uu___3 = FStar_Options.debug_any () in + if uu___3 + then + let uu___4 = + let uu___5 = + FStar_Compiler_List.map + FStar_Interactive_Ide_Types.query_to_string + queries in + FStar_String.concat "\n" uu___5 in + FStar_Compiler_Util.print1 + "Generating queries\n%s\n" uu___4 + else ()); + if request_type = FStar_Interactive_Ide_Types.Full + then queries + else []) + | FStar_Interactive_Ide_Types.Cache -> + let queries = + let uu___1 = + let uu___2 = + FStar_Compiler_Effect.op_Bang + FStar_Interactive_PushHelper.repl_stack in + inspect_repl_stack uu___2 decls + write_full_buffer_fragment_progress in + run_qst uu___1 qid1 in + (if request_type = FStar_Interactive_Ide_Types.Full + then log_syntax_issues err_opt + else (); + (let uu___3 = FStar_Options.debug_any () in + if uu___3 + then + let uu___4 = + let uu___5 = + FStar_Compiler_List.map + FStar_Interactive_Ide_Types.query_to_string + queries in + FStar_String.concat "\n" uu___5 in + FStar_Compiler_Util.print1 + "Generating queries\n%s\n" uu___4 + else ()); + if request_type = FStar_Interactive_Ide_Types.Full + then queries + else []) + | FStar_Interactive_Ide_Types.ReloadDeps -> + let uu___1 = + let uu___2 = + FStar_Compiler_Effect.op_Bang + FStar_Interactive_PushHelper.repl_stack in + reload_deps uu___2 in + run_qst uu___1 qid1) | FStar_Parser_ParseIt.ParseError err -> - (if full + (if request_type = FStar_Interactive_Ide_Types.Full then log_syntax_issues (FStar_Pervasives_Native.Some err) else (); []) | uu___ -> failwith "Unexpected parse result" in qs let (format_code : - FStar_Interactive_ReplState.repl_state -> + FStar_Interactive_Ide_Types.repl_state -> Prims.string -> (Prims.string, FStar_Errors.issue Prims.list) FStar_Pervasives.either) = diff --git a/ocaml/fstar-lib/generated/FStar_Interactive_Lsp.ml b/ocaml/fstar-lib/generated/FStar_Interactive_Lsp.ml index ab383721160..beb920ff56f 100644 --- a/ocaml/fstar-lib/generated/FStar_Interactive_Lsp.ml +++ b/ocaml/fstar-lib/generated/FStar_Interactive_Lsp.ml @@ -210,7 +210,7 @@ let (parse_lsp_query : } | FStar_Pervasives_Native.Some request -> deserialize_lsp_query request let (repl_state_init : - Prims.string -> FStar_Interactive_ReplState.repl_state) = + Prims.string -> FStar_Interactive_Ide_Types.repl_state) = fun fname -> let intial_range = let uu___ = FStar_Compiler_Range.mk_pos Prims.int_one Prims.int_zero in @@ -220,23 +220,24 @@ let (repl_state_init : let env1 = FStar_TypeChecker_Env.set_range env intial_range in let uu___ = FStar_Compiler_Util.open_stdin () in { - FStar_Interactive_ReplState.repl_line = Prims.int_one; - FStar_Interactive_ReplState.repl_column = Prims.int_zero; - FStar_Interactive_ReplState.repl_fname = fname; - FStar_Interactive_ReplState.repl_deps_stack = []; - FStar_Interactive_ReplState.repl_curmod = FStar_Pervasives_Native.None; - FStar_Interactive_ReplState.repl_env = env1; - FStar_Interactive_ReplState.repl_stdin = uu___; - FStar_Interactive_ReplState.repl_names = - FStar_Interactive_CompletionTable.empty + FStar_Interactive_Ide_Types.repl_line = Prims.int_one; + FStar_Interactive_Ide_Types.repl_column = Prims.int_zero; + FStar_Interactive_Ide_Types.repl_fname = fname; + FStar_Interactive_Ide_Types.repl_deps_stack = []; + FStar_Interactive_Ide_Types.repl_curmod = FStar_Pervasives_Native.None; + FStar_Interactive_Ide_Types.repl_env = env1; + FStar_Interactive_Ide_Types.repl_stdin = uu___; + FStar_Interactive_Ide_Types.repl_names = + FStar_Interactive_CompletionTable.empty; + FStar_Interactive_Ide_Types.repl_buffered_input_queries = [] } type optresponse = FStar_Interactive_JsonHelper.assoct FStar_Pervasives_Native.option type either_gst_exit = - (FStar_Interactive_ReplState.grepl_state, Prims.int) + (FStar_Interactive_Ide_Types.grepl_state, Prims.int) FStar_Pervasives.either let (invoke_full_lax : - FStar_Interactive_ReplState.grepl_state -> + FStar_Interactive_Ide_Types.grepl_state -> Prims.string -> Prims.string -> Prims.bool -> (optresponse * either_gst_exit)) = @@ -253,7 +254,7 @@ let (invoke_full_lax : | (diag, st') -> let repls = FStar_Compiler_Util.psmap_add - gst.FStar_Interactive_ReplState.grepl_repls fname st' in + gst.FStar_Interactive_Ide_Types.grepl_repls fname st' in let diag1 = if FStar_Compiler_Util.is_some diag then diag @@ -264,13 +265,13 @@ let (invoke_full_lax : (diag1, (FStar_Pervasives.Inl { - FStar_Interactive_ReplState.grepl_repls = repls; - FStar_Interactive_ReplState.grepl_stdin = - (gst.FStar_Interactive_ReplState.grepl_stdin) + FStar_Interactive_Ide_Types.grepl_repls = repls; + FStar_Interactive_Ide_Types.grepl_stdin = + (gst.FStar_Interactive_Ide_Types.grepl_stdin) }))) in let uu___ = FStar_Compiler_Util.psmap_try_find - gst.FStar_Interactive_ReplState.grepl_repls fname in + gst.FStar_Interactive_Ide_Types.grepl_repls fname in match uu___ with | FStar_Pervasives_Native.Some uu___1 -> if force @@ -278,7 +279,7 @@ let (invoke_full_lax : else (FStar_Pervasives_Native.None, (FStar_Pervasives.Inl gst)) | FStar_Pervasives_Native.None -> aux () let (run_query : - FStar_Interactive_ReplState.grepl_state -> + FStar_Interactive_Ide_Types.grepl_state -> FStar_Interactive_JsonHelper.lquery -> (optresponse * either_gst_exit)) = fun gst -> @@ -334,7 +335,7 @@ let (run_query : | FStar_Interactive_JsonHelper.Completion (txpos, ctx) -> let uu___ = FStar_Compiler_Util.psmap_try_find - gst.FStar_Interactive_ReplState.grepl_repls + gst.FStar_Interactive_Ide_Types.grepl_repls txpos.FStar_Interactive_JsonHelper.path in (match uu___ with | FStar_Pervasives_Native.Some st -> @@ -349,13 +350,13 @@ let (run_query : | FStar_Interactive_JsonHelper.Hover txpos -> let uu___ = FStar_Compiler_Util.psmap_try_find - gst.FStar_Interactive_ReplState.grepl_repls + gst.FStar_Interactive_Ide_Types.grepl_repls txpos.FStar_Interactive_JsonHelper.path in (match uu___ with | FStar_Pervasives_Native.Some st -> let uu___1 = FStar_Interactive_QueryHelper.hoverlookup - st.FStar_Interactive_ReplState.repl_env txpos in + st.FStar_Interactive_Ide_Types.repl_env txpos in (uu___1, (FStar_Pervasives.Inl gst)) | FStar_Pervasives_Native.None -> (FStar_Interactive_JsonHelper.nullResponse, @@ -369,13 +370,13 @@ let (run_query : | FStar_Interactive_JsonHelper.Definition txpos -> let uu___ = FStar_Compiler_Util.psmap_try_find - gst.FStar_Interactive_ReplState.grepl_repls + gst.FStar_Interactive_Ide_Types.grepl_repls txpos.FStar_Interactive_JsonHelper.path in (match uu___ with | FStar_Pervasives_Native.Some st -> let uu___1 = FStar_Interactive_QueryHelper.deflookup - st.FStar_Interactive_ReplState.repl_env txpos in + st.FStar_Interactive_Ide_Types.repl_env txpos in (uu___1, (FStar_Pervasives.Inl gst)) | FStar_Pervasives_Native.None -> (FStar_Interactive_JsonHelper.nullResponse, @@ -495,9 +496,9 @@ let rec (read_lsp_query : (FStar_Compiler_Util.print_error "[E] Malformed Content Header\n"; read_lsp_query stream) | FStar_Interactive_JsonHelper.InputExhausted -> read_lsp_query stream -let rec (go : FStar_Interactive_ReplState.grepl_state -> Prims.int) = +let rec (go : FStar_Interactive_Ide_Types.grepl_state -> Prims.int) = fun gst -> - let query = read_lsp_query gst.FStar_Interactive_ReplState.grepl_stdin in + let query = read_lsp_query gst.FStar_Interactive_Ide_Types.grepl_stdin in let uu___ = run_query gst query.FStar_Interactive_JsonHelper.q in match uu___ with | (r, state_opt) -> @@ -523,8 +524,8 @@ let (start_server : unit -> unit) = let uu___3 = FStar_Compiler_Util.psmap_empty () in let uu___4 = FStar_Compiler_Util.open_stdin () in { - FStar_Interactive_ReplState.grepl_repls = uu___3; - FStar_Interactive_ReplState.grepl_stdin = uu___4 + FStar_Interactive_Ide_Types.grepl_repls = uu___3; + FStar_Interactive_Ide_Types.grepl_stdin = uu___4 } in go uu___2 in FStar_Compiler_Effect.exit uu___1 \ No newline at end of file diff --git a/ocaml/fstar-lib/generated/FStar_Interactive_PushHelper.ml b/ocaml/fstar-lib/generated/FStar_Interactive_PushHelper.ml index 48bf9cad5e9..8cc0e5e5b68 100644 --- a/ocaml/fstar-lib/generated/FStar_Interactive_PushHelper.ml +++ b/ocaml/fstar-lib/generated/FStar_Interactive_PushHelper.ml @@ -1,21 +1,10 @@ open Prims -type push_kind = - | SyntaxCheck - | LaxCheck - | FullCheck -let (uu___is_SyntaxCheck : push_kind -> Prims.bool) = - fun projectee -> - match projectee with | SyntaxCheck -> true | uu___ -> false -let (uu___is_LaxCheck : push_kind -> Prims.bool) = - fun projectee -> match projectee with | LaxCheck -> true | uu___ -> false -let (uu___is_FullCheck : push_kind -> Prims.bool) = - fun projectee -> match projectee with | FullCheck -> true | uu___ -> false type ctx_depth_t = (Prims.int * Prims.int * FStar_TypeChecker_Env.solver_depth_t * Prims.int) type deps_t = FStar_Parser_Dep.deps type either_replst = - (FStar_Interactive_ReplState.repl_state, - FStar_Interactive_ReplState.repl_state) FStar_Pervasives.either + (FStar_Interactive_Ide_Types.repl_state, + FStar_Interactive_Ide_Types.repl_state) FStar_Pervasives.either type name_tracking_event = | NTAlias of (FStar_Ident.lid * FStar_Ident.ident * FStar_Ident.lid) | NTOpen of (FStar_Ident.lid * FStar_Syntax_DsEnv.open_module_or_namespace) @@ -50,16 +39,20 @@ let (__proj__NTBinding__item___0 : FStar_Pervasives.either) = fun projectee -> match projectee with | NTBinding _0 -> _0 let (repl_stack : - FStar_Interactive_ReplState.repl_stack_t FStar_Compiler_Effect.ref) = + FStar_Interactive_Ide_Types.repl_stack_t FStar_Compiler_Effect.ref) = FStar_Compiler_Util.mk_ref [] let (set_check_kind : - FStar_TypeChecker_Env.env_t -> push_kind -> FStar_TypeChecker_Env.env_t) = + FStar_TypeChecker_Env.env_t -> + FStar_Interactive_Ide_Types.push_kind -> FStar_TypeChecker_Env.env_t) + = fun env -> fun check_kind -> - let uu___ = (check_kind = LaxCheck) || (FStar_Options.lax ()) in + let uu___ = + (check_kind = FStar_Interactive_Ide_Types.LaxCheck) || + (FStar_Options.lax ()) in let uu___1 = FStar_Syntax_DsEnv.set_syntax_only env.FStar_TypeChecker_Env.dsenv - (check_kind = SyntaxCheck) in + (check_kind = FStar_Interactive_Ide_Types.SyntaxCheck) in { FStar_TypeChecker_Env.solver = (env.FStar_TypeChecker_Env.solver); FStar_TypeChecker_Env.range = (env.FStar_TypeChecker_Env.range); @@ -144,16 +137,16 @@ let (set_check_kind : } let (repl_ld_tasks_of_deps : Prims.string Prims.list -> - FStar_Interactive_ReplState.repl_task Prims.list -> - FStar_Interactive_ReplState.repl_task Prims.list) + FStar_Interactive_Ide_Types.repl_task Prims.list -> + FStar_Interactive_Ide_Types.repl_task Prims.list) = fun deps -> fun final_tasks -> let wrap fname = let uu___ = FStar_Compiler_Util.now () in { - FStar_Interactive_ReplState.tf_fname = fname; - FStar_Interactive_ReplState.tf_modtime = uu___ + FStar_Interactive_Ide_Types.tf_fname = fname; + FStar_Interactive_Ide_Types.tf_modtime = uu___ } in let rec aux deps1 final_tasks1 = match deps1 with @@ -163,18 +156,18 @@ let (repl_ld_tasks_of_deps : let uu___1 = let uu___2 = wrap intf in let uu___3 = wrap impl in (uu___2, uu___3) in - FStar_Interactive_ReplState.LDInterleaved uu___1 in + FStar_Interactive_Ide_Types.LDInterleaved uu___1 in let uu___1 = aux deps' final_tasks1 in uu___ :: uu___1 | intf_or_impl::deps' -> let uu___ = let uu___1 = wrap intf_or_impl in - FStar_Interactive_ReplState.LDSingle uu___1 in + FStar_Interactive_Ide_Types.LDSingle uu___1 in let uu___1 = aux deps' final_tasks1 in uu___ :: uu___1 | [] -> final_tasks1 in aux deps final_tasks let (deps_and_repl_ld_tasks_of_our_file : Prims.string -> - (Prims.string Prims.list * FStar_Interactive_ReplState.repl_task + (Prims.string Prims.list * FStar_Interactive_Ide_Types.repl_task Prims.list * deps_t)) = fun filename -> @@ -222,10 +215,10 @@ let (deps_and_repl_ld_tasks_of_our_file : let uu___5 = let uu___6 = FStar_Compiler_Util.now () in { - FStar_Interactive_ReplState.tf_fname = intf; - FStar_Interactive_ReplState.tf_modtime = uu___6 + FStar_Interactive_Ide_Types.tf_fname = intf; + FStar_Interactive_Ide_Types.tf_modtime = uu___6 } in - FStar_Interactive_ReplState.LDInterfaceOfCurrentFile + FStar_Interactive_Ide_Types.LDInterfaceOfCurrentFile uu___5 in [uu___4])) | impl::[] -> [] @@ -245,7 +238,7 @@ let (deps_and_repl_ld_tasks_of_our_file : let (snapshot_env : FStar_TypeChecker_Env.env -> Prims.string -> - (FStar_Interactive_ReplState.repl_depth_t * + (FStar_Interactive_Ide_Types.repl_depth_t * FStar_TypeChecker_Env.env_t)) = fun env -> @@ -258,40 +251,42 @@ let (snapshot_env : | (opt_depth, ()) -> ((ctx_depth, opt_depth), env1)) let (push_repl : Prims.string -> - push_kind -> - FStar_Interactive_ReplState.repl_task -> - FStar_Interactive_ReplState.repl_state -> - FStar_Interactive_ReplState.repl_state) + FStar_Interactive_Ide_Types.push_kind -> + FStar_Interactive_Ide_Types.repl_task -> + FStar_Interactive_Ide_Types.repl_state -> + FStar_Interactive_Ide_Types.repl_state) = fun msg -> - fun push_kind1 -> + fun push_kind -> fun task -> fun st -> let uu___ = - snapshot_env st.FStar_Interactive_ReplState.repl_env msg in + snapshot_env st.FStar_Interactive_Ide_Types.repl_env msg in match uu___ with | (depth, env) -> ((let uu___2 = let uu___3 = FStar_Compiler_Effect.op_Bang repl_stack in (depth, (task, st)) :: uu___3 in FStar_Compiler_Effect.op_Colon_Equals repl_stack uu___2); - (let uu___2 = set_check_kind env push_kind1 in + (let uu___2 = set_check_kind env push_kind in { - FStar_Interactive_ReplState.repl_line = - (st.FStar_Interactive_ReplState.repl_line); - FStar_Interactive_ReplState.repl_column = - (st.FStar_Interactive_ReplState.repl_column); - FStar_Interactive_ReplState.repl_fname = - (st.FStar_Interactive_ReplState.repl_fname); - FStar_Interactive_ReplState.repl_deps_stack = - (st.FStar_Interactive_ReplState.repl_deps_stack); - FStar_Interactive_ReplState.repl_curmod = - (st.FStar_Interactive_ReplState.repl_curmod); - FStar_Interactive_ReplState.repl_env = uu___2; - FStar_Interactive_ReplState.repl_stdin = - (st.FStar_Interactive_ReplState.repl_stdin); - FStar_Interactive_ReplState.repl_names = - (st.FStar_Interactive_ReplState.repl_names) + FStar_Interactive_Ide_Types.repl_line = + (st.FStar_Interactive_Ide_Types.repl_line); + FStar_Interactive_Ide_Types.repl_column = + (st.FStar_Interactive_Ide_Types.repl_column); + FStar_Interactive_Ide_Types.repl_fname = + (st.FStar_Interactive_Ide_Types.repl_fname); + FStar_Interactive_Ide_Types.repl_deps_stack = + (st.FStar_Interactive_Ide_Types.repl_deps_stack); + FStar_Interactive_Ide_Types.repl_curmod = + (st.FStar_Interactive_Ide_Types.repl_curmod); + FStar_Interactive_Ide_Types.repl_env = uu___2; + FStar_Interactive_Ide_Types.repl_stdin = + (st.FStar_Interactive_Ide_Types.repl_stdin); + FStar_Interactive_Ide_Types.repl_names = + (st.FStar_Interactive_Ide_Types.repl_names); + FStar_Interactive_Ide_Types.repl_buffered_input_queries = + (st.FStar_Interactive_Ide_Types.repl_buffered_input_queries) })) let (rollback_env : FStar_TypeChecker_Env.solver_t -> @@ -311,8 +306,8 @@ let (rollback_env : env) let (pop_repl : Prims.string -> - FStar_Interactive_ReplState.repl_state -> - FStar_Interactive_ReplState.repl_state) + FStar_Interactive_Ide_Types.repl_state -> + FStar_Interactive_Ide_Types.repl_state) = fun msg -> fun st -> @@ -322,12 +317,12 @@ let (pop_repl : | (depth, (uu___1, st'))::stack_tl -> let env = rollback_env - (st.FStar_Interactive_ReplState.repl_env).FStar_TypeChecker_Env.solver + (st.FStar_Interactive_Ide_Types.repl_env).FStar_TypeChecker_Env.solver msg depth in (FStar_Compiler_Effect.op_Colon_Equals repl_stack stack_tl; (let uu___4 = FStar_Compiler_Util.physical_equality env - st'.FStar_Interactive_ReplState.repl_env in + st'.FStar_Interactive_Ide_Types.repl_env in FStar_Common.runtime_assert uu___4 "Inconsistent stack state"); st') let (tc_one : @@ -347,35 +342,35 @@ let (tc_one : FStar_Universal.tc_one_file_for_ide env intf_opt modf parse_data in match uu___ with | (uu___1, env1) -> env1 let (run_repl_task : - FStar_Interactive_ReplState.optmod_t -> + FStar_Interactive_Ide_Types.optmod_t -> FStar_TypeChecker_Env.env_t -> - FStar_Interactive_ReplState.repl_task -> - (FStar_Interactive_ReplState.optmod_t * FStar_TypeChecker_Env.env_t)) + FStar_Interactive_Ide_Types.repl_task -> + (FStar_Interactive_Ide_Types.optmod_t * FStar_TypeChecker_Env.env_t)) = fun curmod -> fun env -> fun task -> match task with - | FStar_Interactive_ReplState.LDInterleaved (intf, impl) -> + | FStar_Interactive_Ide_Types.LDInterleaved (intf, impl) -> let uu___ = tc_one env (FStar_Pervasives_Native.Some - (intf.FStar_Interactive_ReplState.tf_fname)) - impl.FStar_Interactive_ReplState.tf_fname in + (intf.FStar_Interactive_Ide_Types.tf_fname)) + impl.FStar_Interactive_Ide_Types.tf_fname in (curmod, uu___) - | FStar_Interactive_ReplState.LDSingle intf_or_impl -> + | FStar_Interactive_Ide_Types.LDSingle intf_or_impl -> let uu___ = tc_one env FStar_Pervasives_Native.None - intf_or_impl.FStar_Interactive_ReplState.tf_fname in + intf_or_impl.FStar_Interactive_Ide_Types.tf_fname in (curmod, uu___) - | FStar_Interactive_ReplState.LDInterfaceOfCurrentFile intf -> + | FStar_Interactive_Ide_Types.LDInterfaceOfCurrentFile intf -> let uu___ = FStar_Universal.load_interface_decls env - intf.FStar_Interactive_ReplState.tf_fname in + intf.FStar_Interactive_Ide_Types.tf_fname in (curmod, uu___) - | FStar_Interactive_ReplState.PushFragment frag -> + | FStar_Interactive_Ide_Types.PushFragment frag -> FStar_Universal.tc_one_fragment curmod env frag - | FStar_Interactive_ReplState.Noop -> (curmod, env) + | FStar_Interactive_Ide_Types.Noop -> (curmod, env) let (query_of_ids : FStar_Ident.ident Prims.list -> FStar_Interactive_CompletionTable.query) = fun ids -> FStar_Compiler_List.map FStar_Ident.string_of_id ids @@ -464,30 +459,32 @@ let (commit_name_tracking' : let updater = update_names_from_event cur_mod_str in FStar_Compiler_List.fold_left updater names name_events let (commit_name_tracking : - FStar_Interactive_ReplState.repl_state -> - name_tracking_event Prims.list -> FStar_Interactive_ReplState.repl_state) + FStar_Interactive_Ide_Types.repl_state -> + name_tracking_event Prims.list -> FStar_Interactive_Ide_Types.repl_state) = fun st -> fun name_events -> let names = - commit_name_tracking' st.FStar_Interactive_ReplState.repl_curmod - st.FStar_Interactive_ReplState.repl_names name_events in + commit_name_tracking' st.FStar_Interactive_Ide_Types.repl_curmod + st.FStar_Interactive_Ide_Types.repl_names name_events in { - FStar_Interactive_ReplState.repl_line = - (st.FStar_Interactive_ReplState.repl_line); - FStar_Interactive_ReplState.repl_column = - (st.FStar_Interactive_ReplState.repl_column); - FStar_Interactive_ReplState.repl_fname = - (st.FStar_Interactive_ReplState.repl_fname); - FStar_Interactive_ReplState.repl_deps_stack = - (st.FStar_Interactive_ReplState.repl_deps_stack); - FStar_Interactive_ReplState.repl_curmod = - (st.FStar_Interactive_ReplState.repl_curmod); - FStar_Interactive_ReplState.repl_env = - (st.FStar_Interactive_ReplState.repl_env); - FStar_Interactive_ReplState.repl_stdin = - (st.FStar_Interactive_ReplState.repl_stdin); - FStar_Interactive_ReplState.repl_names = names + FStar_Interactive_Ide_Types.repl_line = + (st.FStar_Interactive_Ide_Types.repl_line); + FStar_Interactive_Ide_Types.repl_column = + (st.FStar_Interactive_Ide_Types.repl_column); + FStar_Interactive_Ide_Types.repl_fname = + (st.FStar_Interactive_Ide_Types.repl_fname); + FStar_Interactive_Ide_Types.repl_deps_stack = + (st.FStar_Interactive_Ide_Types.repl_deps_stack); + FStar_Interactive_Ide_Types.repl_curmod = + (st.FStar_Interactive_Ide_Types.repl_curmod); + FStar_Interactive_Ide_Types.repl_env = + (st.FStar_Interactive_Ide_Types.repl_env); + FStar_Interactive_Ide_Types.repl_stdin = + (st.FStar_Interactive_Ide_Types.repl_stdin); + FStar_Interactive_Ide_Types.repl_names = names; + FStar_Interactive_Ide_Types.repl_buffered_input_queries = + (st.FStar_Interactive_Ide_Types.repl_buffered_input_queries) } let (fresh_name_tracking_hooks : unit -> @@ -565,48 +562,51 @@ let (track_name_changes : FStar_Compiler_List.rev uu___5 in (uu___3, uu___4))))) let (repl_tx : - FStar_Interactive_ReplState.repl_state -> - push_kind -> - FStar_Interactive_ReplState.repl_task -> + FStar_Interactive_Ide_Types.repl_state -> + FStar_Interactive_Ide_Types.push_kind -> + FStar_Interactive_Ide_Types.repl_task -> (FStar_Interactive_JsonHelper.assoct FStar_Pervasives_Native.option * - FStar_Interactive_ReplState.repl_state)) + FStar_Interactive_Ide_Types.repl_state)) = fun st -> - fun push_kind1 -> + fun push_kind -> fun task -> try (fun uu___ -> match () with | () -> - let st1 = push_repl "repl_tx" push_kind1 task st in + let st1 = push_repl "repl_tx" push_kind task st in let uu___1 = track_name_changes - st1.FStar_Interactive_ReplState.repl_env in + st1.FStar_Interactive_Ide_Types.repl_env in (match uu___1 with | (env, finish_name_tracking) -> let uu___2 = run_repl_task - st1.FStar_Interactive_ReplState.repl_curmod env + st1.FStar_Interactive_Ide_Types.repl_curmod env task in (match uu___2 with | (curmod, env1) -> let st2 = { - FStar_Interactive_ReplState.repl_line = - (st1.FStar_Interactive_ReplState.repl_line); - FStar_Interactive_ReplState.repl_column = - (st1.FStar_Interactive_ReplState.repl_column); - FStar_Interactive_ReplState.repl_fname = - (st1.FStar_Interactive_ReplState.repl_fname); - FStar_Interactive_ReplState.repl_deps_stack = - (st1.FStar_Interactive_ReplState.repl_deps_stack); - FStar_Interactive_ReplState.repl_curmod = + FStar_Interactive_Ide_Types.repl_line = + (st1.FStar_Interactive_Ide_Types.repl_line); + FStar_Interactive_Ide_Types.repl_column = + (st1.FStar_Interactive_Ide_Types.repl_column); + FStar_Interactive_Ide_Types.repl_fname = + (st1.FStar_Interactive_Ide_Types.repl_fname); + FStar_Interactive_Ide_Types.repl_deps_stack = + (st1.FStar_Interactive_Ide_Types.repl_deps_stack); + FStar_Interactive_Ide_Types.repl_curmod = curmod; - FStar_Interactive_ReplState.repl_env = env1; - FStar_Interactive_ReplState.repl_stdin = - (st1.FStar_Interactive_ReplState.repl_stdin); - FStar_Interactive_ReplState.repl_names = - (st1.FStar_Interactive_ReplState.repl_names) + FStar_Interactive_Ide_Types.repl_env = env1; + FStar_Interactive_Ide_Types.repl_stdin = + (st1.FStar_Interactive_Ide_Types.repl_stdin); + FStar_Interactive_Ide_Types.repl_names = + (st1.FStar_Interactive_Ide_Types.repl_names); + FStar_Interactive_Ide_Types.repl_buffered_input_queries + = + (st1.FStar_Interactive_Ide_Types.repl_buffered_input_queries) } in let uu___3 = finish_name_tracking env1 in (match uu___3 with @@ -619,7 +619,7 @@ let (repl_tx : let uu___1 = let uu___2 = FStar_Interactive_JsonHelper.js_diag - st.FStar_Interactive_ReplState.repl_fname msg + st.FStar_Interactive_Ide_Types.repl_fname msg FStar_Pervasives_Native.None in FStar_Pervasives_Native.Some uu___2 in (uu___1, st) @@ -630,7 +630,7 @@ let (repl_tx : let uu___1 = let uu___2 = FStar_Interactive_JsonHelper.js_diag - st.FStar_Interactive_ReplState.repl_fname msg + st.FStar_Interactive_Ide_Types.repl_fname msg (FStar_Pervasives_Native.Some r) in FStar_Pervasives_Native.Some uu___2 in (uu___1, st) @@ -638,43 +638,43 @@ let (repl_tx : let uu___1 = let uu___2 = FStar_Interactive_JsonHelper.js_diag - st.FStar_Interactive_ReplState.repl_fname msg + st.FStar_Interactive_Ide_Types.repl_fname msg FStar_Pervasives_Native.None in FStar_Pervasives_Native.Some uu___2 in (uu___1, st) | FStar_Errors.Stop -> (FStar_Compiler_Util.print_error "[E] Stop"; (FStar_Pervasives_Native.None, st)) -let (tf_of_fname : Prims.string -> FStar_Interactive_ReplState.timed_fname) = +let (tf_of_fname : Prims.string -> FStar_Interactive_Ide_Types.timed_fname) = fun fname -> let uu___ = FStar_Parser_ParseIt.get_file_last_modification_time fname in { - FStar_Interactive_ReplState.tf_fname = fname; - FStar_Interactive_ReplState.tf_modtime = uu___ + FStar_Interactive_Ide_Types.tf_fname = fname; + FStar_Interactive_Ide_Types.tf_modtime = uu___ } let (update_task_timestamps : - FStar_Interactive_ReplState.repl_task -> - FStar_Interactive_ReplState.repl_task) + FStar_Interactive_Ide_Types.repl_task -> + FStar_Interactive_Ide_Types.repl_task) = fun uu___ -> match uu___ with - | FStar_Interactive_ReplState.LDInterleaved (intf, impl) -> + | FStar_Interactive_Ide_Types.LDInterleaved (intf, impl) -> let uu___1 = - let uu___2 = tf_of_fname intf.FStar_Interactive_ReplState.tf_fname in - let uu___3 = tf_of_fname impl.FStar_Interactive_ReplState.tf_fname in + let uu___2 = tf_of_fname intf.FStar_Interactive_Ide_Types.tf_fname in + let uu___3 = tf_of_fname impl.FStar_Interactive_Ide_Types.tf_fname in (uu___2, uu___3) in - FStar_Interactive_ReplState.LDInterleaved uu___1 - | FStar_Interactive_ReplState.LDSingle intf_or_impl -> + FStar_Interactive_Ide_Types.LDInterleaved uu___1 + | FStar_Interactive_Ide_Types.LDSingle intf_or_impl -> let uu___1 = - tf_of_fname intf_or_impl.FStar_Interactive_ReplState.tf_fname in - FStar_Interactive_ReplState.LDSingle uu___1 - | FStar_Interactive_ReplState.LDInterfaceOfCurrentFile intf -> - let uu___1 = tf_of_fname intf.FStar_Interactive_ReplState.tf_fname in - FStar_Interactive_ReplState.LDInterfaceOfCurrentFile uu___1 + tf_of_fname intf_or_impl.FStar_Interactive_Ide_Types.tf_fname in + FStar_Interactive_Ide_Types.LDSingle uu___1 + | FStar_Interactive_Ide_Types.LDInterfaceOfCurrentFile intf -> + let uu___1 = tf_of_fname intf.FStar_Interactive_Ide_Types.tf_fname in + FStar_Interactive_Ide_Types.LDInterfaceOfCurrentFile uu___1 | other -> other let (repl_ldtx : - FStar_Interactive_ReplState.repl_state -> - FStar_Interactive_ReplState.repl_task Prims.list -> either_replst) + FStar_Interactive_Ide_Types.repl_state -> + FStar_Interactive_Ide_Types.repl_task Prims.list -> either_replst) = fun st -> fun tasks -> @@ -685,27 +685,29 @@ let (repl_ldtx : let st' = pop_repl "repl_ldtx" st1 in let dep_graph = FStar_TypeChecker_Env.dep_graph - st1.FStar_Interactive_ReplState.repl_env in + st1.FStar_Interactive_Ide_Types.repl_env in let st'1 = let uu___1 = FStar_TypeChecker_Env.set_dep_graph - st'.FStar_Interactive_ReplState.repl_env dep_graph in + st'.FStar_Interactive_Ide_Types.repl_env dep_graph in { - FStar_Interactive_ReplState.repl_line = - (st'.FStar_Interactive_ReplState.repl_line); - FStar_Interactive_ReplState.repl_column = - (st'.FStar_Interactive_ReplState.repl_column); - FStar_Interactive_ReplState.repl_fname = - (st'.FStar_Interactive_ReplState.repl_fname); - FStar_Interactive_ReplState.repl_deps_stack = - (st'.FStar_Interactive_ReplState.repl_deps_stack); - FStar_Interactive_ReplState.repl_curmod = - (st'.FStar_Interactive_ReplState.repl_curmod); - FStar_Interactive_ReplState.repl_env = uu___1; - FStar_Interactive_ReplState.repl_stdin = - (st'.FStar_Interactive_ReplState.repl_stdin); - FStar_Interactive_ReplState.repl_names = - (st'.FStar_Interactive_ReplState.repl_names) + FStar_Interactive_Ide_Types.repl_line = + (st'.FStar_Interactive_Ide_Types.repl_line); + FStar_Interactive_Ide_Types.repl_column = + (st'.FStar_Interactive_Ide_Types.repl_column); + FStar_Interactive_Ide_Types.repl_fname = + (st'.FStar_Interactive_Ide_Types.repl_fname); + FStar_Interactive_Ide_Types.repl_deps_stack = + (st'.FStar_Interactive_Ide_Types.repl_deps_stack); + FStar_Interactive_Ide_Types.repl_curmod = + (st'.FStar_Interactive_Ide_Types.repl_curmod); + FStar_Interactive_Ide_Types.repl_env = uu___1; + FStar_Interactive_Ide_Types.repl_stdin = + (st'.FStar_Interactive_Ide_Types.repl_stdin); + FStar_Interactive_Ide_Types.repl_names = + (st'.FStar_Interactive_Ide_Types.repl_names); + FStar_Interactive_Ide_Types.repl_buffered_input_queries = + (st'.FStar_Interactive_Ide_Types.repl_buffered_input_queries) } in revert_many st'1 entries in let rec aux st1 tasks1 previous = @@ -713,7 +715,9 @@ let (repl_ldtx : | ([], []) -> FStar_Pervasives.Inl st1 | (task::tasks2, []) -> let timestamped_task = update_task_timestamps task in - let uu___ = repl_tx st1 LaxCheck timestamped_task in + let uu___ = + repl_tx st1 FStar_Interactive_Ide_Types.LaxCheck + timestamped_task in (match uu___ with | (diag, st2) -> if Prims.op_Negation (FStar_Compiler_Util.is_some diag) @@ -721,21 +725,24 @@ let (repl_ldtx : let uu___1 = let uu___2 = FStar_Compiler_Effect.op_Bang repl_stack in { - FStar_Interactive_ReplState.repl_line = - (st2.FStar_Interactive_ReplState.repl_line); - FStar_Interactive_ReplState.repl_column = - (st2.FStar_Interactive_ReplState.repl_column); - FStar_Interactive_ReplState.repl_fname = - (st2.FStar_Interactive_ReplState.repl_fname); - FStar_Interactive_ReplState.repl_deps_stack = uu___2; - FStar_Interactive_ReplState.repl_curmod = - (st2.FStar_Interactive_ReplState.repl_curmod); - FStar_Interactive_ReplState.repl_env = - (st2.FStar_Interactive_ReplState.repl_env); - FStar_Interactive_ReplState.repl_stdin = - (st2.FStar_Interactive_ReplState.repl_stdin); - FStar_Interactive_ReplState.repl_names = - (st2.FStar_Interactive_ReplState.repl_names) + FStar_Interactive_Ide_Types.repl_line = + (st2.FStar_Interactive_Ide_Types.repl_line); + FStar_Interactive_Ide_Types.repl_column = + (st2.FStar_Interactive_Ide_Types.repl_column); + FStar_Interactive_Ide_Types.repl_fname = + (st2.FStar_Interactive_Ide_Types.repl_fname); + FStar_Interactive_Ide_Types.repl_deps_stack = uu___2; + FStar_Interactive_Ide_Types.repl_curmod = + (st2.FStar_Interactive_Ide_Types.repl_curmod); + FStar_Interactive_Ide_Types.repl_env = + (st2.FStar_Interactive_Ide_Types.repl_env); + FStar_Interactive_Ide_Types.repl_stdin = + (st2.FStar_Interactive_Ide_Types.repl_stdin); + FStar_Interactive_Ide_Types.repl_names = + (st2.FStar_Interactive_Ide_Types.repl_names); + FStar_Interactive_Ide_Types.repl_buffered_input_queries + = + (st2.FStar_Interactive_Ide_Types.repl_buffered_input_queries) } in aux uu___1 tasks2 [] else FStar_Pervasives.Inr st2) @@ -748,11 +755,11 @@ let (repl_ldtx : let uu___ = revert_many st1 previous1 in aux uu___ tasks2 [] in aux st tasks (FStar_Compiler_List.rev - st.FStar_Interactive_ReplState.repl_deps_stack) + st.FStar_Interactive_Ide_Types.repl_deps_stack) let (ld_deps : - FStar_Interactive_ReplState.repl_state -> - ((FStar_Interactive_ReplState.repl_state * Prims.string Prims.list), - FStar_Interactive_ReplState.repl_state) FStar_Pervasives.either) + FStar_Interactive_Ide_Types.repl_state -> + ((FStar_Interactive_Ide_Types.repl_state * Prims.string Prims.list), + FStar_Interactive_Ide_Types.repl_state) FStar_Pervasives.either) = fun st -> try @@ -761,29 +768,32 @@ let (ld_deps : | () -> let uu___1 = deps_and_repl_ld_tasks_of_our_file - st.FStar_Interactive_ReplState.repl_fname in + st.FStar_Interactive_Ide_Types.repl_fname in (match uu___1 with | (deps, tasks, dep_graph) -> let st1 = let uu___2 = FStar_TypeChecker_Env.set_dep_graph - st.FStar_Interactive_ReplState.repl_env dep_graph in + st.FStar_Interactive_Ide_Types.repl_env dep_graph in { - FStar_Interactive_ReplState.repl_line = - (st.FStar_Interactive_ReplState.repl_line); - FStar_Interactive_ReplState.repl_column = - (st.FStar_Interactive_ReplState.repl_column); - FStar_Interactive_ReplState.repl_fname = - (st.FStar_Interactive_ReplState.repl_fname); - FStar_Interactive_ReplState.repl_deps_stack = - (st.FStar_Interactive_ReplState.repl_deps_stack); - FStar_Interactive_ReplState.repl_curmod = - (st.FStar_Interactive_ReplState.repl_curmod); - FStar_Interactive_ReplState.repl_env = uu___2; - FStar_Interactive_ReplState.repl_stdin = - (st.FStar_Interactive_ReplState.repl_stdin); - FStar_Interactive_ReplState.repl_names = - (st.FStar_Interactive_ReplState.repl_names) + FStar_Interactive_Ide_Types.repl_line = + (st.FStar_Interactive_Ide_Types.repl_line); + FStar_Interactive_Ide_Types.repl_column = + (st.FStar_Interactive_Ide_Types.repl_column); + FStar_Interactive_Ide_Types.repl_fname = + (st.FStar_Interactive_Ide_Types.repl_fname); + FStar_Interactive_Ide_Types.repl_deps_stack = + (st.FStar_Interactive_Ide_Types.repl_deps_stack); + FStar_Interactive_Ide_Types.repl_curmod = + (st.FStar_Interactive_Ide_Types.repl_curmod); + FStar_Interactive_Ide_Types.repl_env = uu___2; + FStar_Interactive_Ide_Types.repl_stdin = + (st.FStar_Interactive_Ide_Types.repl_stdin); + FStar_Interactive_Ide_Types.repl_names = + (st.FStar_Interactive_Ide_Types.repl_names); + FStar_Interactive_Ide_Types.repl_buffered_input_queries + = + (st.FStar_Interactive_Ide_Types.repl_buffered_input_queries) } in let uu___2 = repl_ldtx st1 tasks in (match uu___2 with @@ -849,18 +859,18 @@ let (add_module_completions : (FStar_Compiler_List.rev mods) let (full_lax : Prims.string -> - FStar_Interactive_ReplState.repl_state -> + FStar_Interactive_Ide_Types.repl_state -> (FStar_Interactive_JsonHelper.assoct FStar_Pervasives_Native.option * - FStar_Interactive_ReplState.repl_state)) + FStar_Interactive_Ide_Types.repl_state)) = fun text -> fun st -> FStar_TypeChecker_Env.toggle_id_info - st.FStar_Interactive_ReplState.repl_env true; + st.FStar_Interactive_Ide_Types.repl_env true; (let frag = { FStar_Parser_ParseIt.frag_fname = - (st.FStar_Interactive_ReplState.repl_fname); + (st.FStar_Interactive_Ide_Types.repl_fname); FStar_Parser_ParseIt.frag_text = text; FStar_Parser_ParseIt.frag_line = Prims.int_one; FStar_Parser_ParseIt.frag_col = Prims.int_zero @@ -870,26 +880,28 @@ let (full_lax : | FStar_Pervasives.Inl (st1, deps) -> let names = add_module_completions - st1.FStar_Interactive_ReplState.repl_fname deps - st1.FStar_Interactive_ReplState.repl_names in + st1.FStar_Interactive_Ide_Types.repl_fname deps + st1.FStar_Interactive_Ide_Types.repl_names in repl_tx { - FStar_Interactive_ReplState.repl_line = - (st1.FStar_Interactive_ReplState.repl_line); - FStar_Interactive_ReplState.repl_column = - (st1.FStar_Interactive_ReplState.repl_column); - FStar_Interactive_ReplState.repl_fname = - (st1.FStar_Interactive_ReplState.repl_fname); - FStar_Interactive_ReplState.repl_deps_stack = - (st1.FStar_Interactive_ReplState.repl_deps_stack); - FStar_Interactive_ReplState.repl_curmod = - (st1.FStar_Interactive_ReplState.repl_curmod); - FStar_Interactive_ReplState.repl_env = - (st1.FStar_Interactive_ReplState.repl_env); - FStar_Interactive_ReplState.repl_stdin = - (st1.FStar_Interactive_ReplState.repl_stdin); - FStar_Interactive_ReplState.repl_names = names - } LaxCheck - (FStar_Interactive_ReplState.PushFragment + FStar_Interactive_Ide_Types.repl_line = + (st1.FStar_Interactive_Ide_Types.repl_line); + FStar_Interactive_Ide_Types.repl_column = + (st1.FStar_Interactive_Ide_Types.repl_column); + FStar_Interactive_Ide_Types.repl_fname = + (st1.FStar_Interactive_Ide_Types.repl_fname); + FStar_Interactive_Ide_Types.repl_deps_stack = + (st1.FStar_Interactive_Ide_Types.repl_deps_stack); + FStar_Interactive_Ide_Types.repl_curmod = + (st1.FStar_Interactive_Ide_Types.repl_curmod); + FStar_Interactive_Ide_Types.repl_env = + (st1.FStar_Interactive_Ide_Types.repl_env); + FStar_Interactive_Ide_Types.repl_stdin = + (st1.FStar_Interactive_Ide_Types.repl_stdin); + FStar_Interactive_Ide_Types.repl_names = names; + FStar_Interactive_Ide_Types.repl_buffered_input_queries = + (st1.FStar_Interactive_Ide_Types.repl_buffered_input_queries) + } FStar_Interactive_Ide_Types.LaxCheck + (FStar_Interactive_Ide_Types.PushFragment (FStar_Pervasives.Inl frag)) | FStar_Pervasives.Inr st1 -> (FStar_Pervasives_Native.None, st1)) \ No newline at end of file diff --git a/ocaml/fstar-lib/generated/FStar_Interactive_QueryHelper.ml b/ocaml/fstar-lib/generated/FStar_Interactive_QueryHelper.ml index c69641c4a6d..ffe40f25332 100644 --- a/ocaml/fstar-lib/generated/FStar_Interactive_QueryHelper.ml +++ b/ocaml/fstar-lib/generated/FStar_Interactive_QueryHelper.ml @@ -171,7 +171,7 @@ let mod_filter : (pth, uu___2) in FStar_Pervasives_Native.Some uu___1 let (ck_completion : - FStar_Interactive_ReplState.repl_state -> + FStar_Interactive_Ide_Types.repl_state -> Prims.string -> FStar_Interactive_CompletionTable.completion_result Prims.list) = @@ -180,10 +180,10 @@ let (ck_completion : let needle = FStar_Compiler_Util.split search_term "." in let mods_and_nss = FStar_Interactive_CompletionTable.autocomplete_mod_or_ns - st.FStar_Interactive_ReplState.repl_names needle mod_filter in + st.FStar_Interactive_Ide_Types.repl_names needle mod_filter in let lids = FStar_Interactive_CompletionTable.autocomplete_lid - st.FStar_Interactive_ReplState.repl_names needle in + st.FStar_Interactive_Ide_Types.repl_names needle in FStar_Compiler_List.op_At lids mods_and_nss let (deflookup : FStar_TypeChecker_Env.env -> @@ -235,7 +235,7 @@ let (hoverlookup : ("value", (FStar_Compiler_Util.JsonStr hovertxt))]))]) | uu___1 -> FStar_Interactive_JsonHelper.nullResponse let (complookup : - FStar_Interactive_ReplState.repl_state -> + FStar_Interactive_Ide_Types.repl_state -> FStar_Interactive_JsonHelper.txdoc_pos -> FStar_Interactive_JsonHelper.assoct FStar_Pervasives_Native.option) = diff --git a/ocaml/fstar-lib/generated/FStar_Interactive_ReplState.ml b/ocaml/fstar-lib/generated/FStar_Interactive_ReplState.ml index 0a505ac7e74..e8306abedb2 100644 --- a/ocaml/fstar-lib/generated/FStar_Interactive_ReplState.ml +++ b/ocaml/fstar-lib/generated/FStar_Interactive_ReplState.ml @@ -1,184 +1 @@ -open Prims -type repl_depth_t = (FStar_TypeChecker_Env.tcenv_depth_t * Prims.int) -type optmod_t = FStar_Syntax_Syntax.modul FStar_Pervasives_Native.option -type timed_fname = - { - tf_fname: Prims.string ; - tf_modtime: FStar_Compiler_Util.time } -let (__proj__Mktimed_fname__item__tf_fname : timed_fname -> Prims.string) = - fun projectee -> - match projectee with | { tf_fname; tf_modtime;_} -> tf_fname -let (__proj__Mktimed_fname__item__tf_modtime : - timed_fname -> FStar_Compiler_Util.time) = - fun projectee -> - match projectee with | { tf_fname; tf_modtime;_} -> tf_modtime -type repl_task = - | LDInterleaved of (timed_fname * timed_fname) - | LDSingle of timed_fname - | LDInterfaceOfCurrentFile of timed_fname - | PushFragment of (FStar_Parser_ParseIt.input_frag, FStar_Parser_AST.decl) - FStar_Pervasives.either - | Noop -let (uu___is_LDInterleaved : repl_task -> Prims.bool) = - fun projectee -> - match projectee with | LDInterleaved _0 -> true | uu___ -> false -let (__proj__LDInterleaved__item___0 : - repl_task -> (timed_fname * timed_fname)) = - fun projectee -> match projectee with | LDInterleaved _0 -> _0 -let (uu___is_LDSingle : repl_task -> Prims.bool) = - fun projectee -> - match projectee with | LDSingle _0 -> true | uu___ -> false -let (__proj__LDSingle__item___0 : repl_task -> timed_fname) = - fun projectee -> match projectee with | LDSingle _0 -> _0 -let (uu___is_LDInterfaceOfCurrentFile : repl_task -> Prims.bool) = - fun projectee -> - match projectee with - | LDInterfaceOfCurrentFile _0 -> true - | uu___ -> false -let (__proj__LDInterfaceOfCurrentFile__item___0 : repl_task -> timed_fname) = - fun projectee -> match projectee with | LDInterfaceOfCurrentFile _0 -> _0 -let (uu___is_PushFragment : repl_task -> Prims.bool) = - fun projectee -> - match projectee with | PushFragment _0 -> true | uu___ -> false -let (__proj__PushFragment__item___0 : - repl_task -> - (FStar_Parser_ParseIt.input_frag, FStar_Parser_AST.decl) - FStar_Pervasives.either) - = fun projectee -> match projectee with | PushFragment _0 -> _0 -let (uu___is_Noop : repl_task -> Prims.bool) = - fun projectee -> match projectee with | Noop -> true | uu___ -> false -type repl_state = - { - repl_line: Prims.int ; - repl_column: Prims.int ; - repl_fname: Prims.string ; - repl_deps_stack: (repl_depth_t * (repl_task * repl_state)) Prims.list ; - repl_curmod: optmod_t ; - repl_env: FStar_TypeChecker_Env.env ; - repl_stdin: FStar_Compiler_Util.stream_reader ; - repl_names: FStar_Interactive_CompletionTable.table } -let (__proj__Mkrepl_state__item__repl_line : repl_state -> Prims.int) = - fun projectee -> - match projectee with - | { repl_line; repl_column; repl_fname; repl_deps_stack; repl_curmod; - repl_env; repl_stdin; repl_names;_} -> repl_line -let (__proj__Mkrepl_state__item__repl_column : repl_state -> Prims.int) = - fun projectee -> - match projectee with - | { repl_line; repl_column; repl_fname; repl_deps_stack; repl_curmod; - repl_env; repl_stdin; repl_names;_} -> repl_column -let (__proj__Mkrepl_state__item__repl_fname : repl_state -> Prims.string) = - fun projectee -> - match projectee with - | { repl_line; repl_column; repl_fname; repl_deps_stack; repl_curmod; - repl_env; repl_stdin; repl_names;_} -> repl_fname -let (__proj__Mkrepl_state__item__repl_deps_stack : - repl_state -> (repl_depth_t * (repl_task * repl_state)) Prims.list) = - fun projectee -> - match projectee with - | { repl_line; repl_column; repl_fname; repl_deps_stack; repl_curmod; - repl_env; repl_stdin; repl_names;_} -> repl_deps_stack -let (__proj__Mkrepl_state__item__repl_curmod : repl_state -> optmod_t) = - fun projectee -> - match projectee with - | { repl_line; repl_column; repl_fname; repl_deps_stack; repl_curmod; - repl_env; repl_stdin; repl_names;_} -> repl_curmod -let (__proj__Mkrepl_state__item__repl_env : - repl_state -> FStar_TypeChecker_Env.env) = - fun projectee -> - match projectee with - | { repl_line; repl_column; repl_fname; repl_deps_stack; repl_curmod; - repl_env; repl_stdin; repl_names;_} -> repl_env -let (__proj__Mkrepl_state__item__repl_stdin : - repl_state -> FStar_Compiler_Util.stream_reader) = - fun projectee -> - match projectee with - | { repl_line; repl_column; repl_fname; repl_deps_stack; repl_curmod; - repl_env; repl_stdin; repl_names;_} -> repl_stdin -let (__proj__Mkrepl_state__item__repl_names : - repl_state -> FStar_Interactive_CompletionTable.table) = - fun projectee -> - match projectee with - | { repl_line; repl_column; repl_fname; repl_deps_stack; repl_curmod; - repl_env; repl_stdin; repl_names;_} -> repl_names -type repl_stack_entry_t = (repl_depth_t * (repl_task * repl_state)) -type repl_stack_t = (repl_depth_t * (repl_task * repl_state)) Prims.list -type grepl_state = - { - grepl_repls: repl_state FStar_Compiler_Util.psmap ; - grepl_stdin: FStar_Compiler_Util.stream_reader } -let (__proj__Mkgrepl_state__item__grepl_repls : - grepl_state -> repl_state FStar_Compiler_Util.psmap) = - fun projectee -> - match projectee with | { grepl_repls; grepl_stdin;_} -> grepl_repls -let (__proj__Mkgrepl_state__item__grepl_stdin : - grepl_state -> FStar_Compiler_Util.stream_reader) = - fun projectee -> - match projectee with | { grepl_repls; grepl_stdin;_} -> grepl_stdin -let (t0 : FStar_Compiler_Util.time) = FStar_Compiler_Util.now () -let (dummy_tf_of_fname : Prims.string -> timed_fname) = - fun fname -> { tf_fname = fname; tf_modtime = t0 } -let (string_of_timed_fname : timed_fname -> Prims.string) = - fun uu___ -> - match uu___ with - | { tf_fname = fname; tf_modtime = modtime;_} -> - if modtime = t0 - then FStar_Compiler_Util.format1 "{ %s }" fname - else - (let uu___2 = FStar_Compiler_Util.string_of_time modtime in - FStar_Compiler_Util.format2 "{ %s; %s }" fname uu___2) -let (string_of_repl_task : repl_task -> Prims.string) = - fun uu___ -> - match uu___ with - | LDInterleaved (intf, impl) -> - let uu___1 = string_of_timed_fname intf in - let uu___2 = string_of_timed_fname impl in - FStar_Compiler_Util.format2 "LDInterleaved (%s, %s)" uu___1 uu___2 - | LDSingle intf_or_impl -> - let uu___1 = string_of_timed_fname intf_or_impl in - FStar_Compiler_Util.format1 "LDSingle %s" uu___1 - | LDInterfaceOfCurrentFile intf -> - let uu___1 = string_of_timed_fname intf in - FStar_Compiler_Util.format1 "LDInterfaceOfCurrentFile %s" uu___1 - | PushFragment (FStar_Pervasives.Inl frag) -> - FStar_Compiler_Util.format1 "PushFragment { code = %s }" - frag.FStar_Parser_ParseIt.frag_text - | PushFragment (FStar_Pervasives.Inr d) -> - let uu___1 = FStar_Parser_AST.decl_to_string d in - FStar_Compiler_Util.format1 "PushFragment { decl = %s }" uu___1 - | Noop -> "Noop {}" -let (string_of_repl_stack_entry : repl_stack_entry_t -> Prims.string) = - fun uu___ -> - match uu___ with - | ((depth, i), (task, state)) -> - let uu___1 = - let uu___2 = FStar_Compiler_Util.string_of_int i in - let uu___3 = let uu___4 = string_of_repl_task task in [uu___4] in - uu___2 :: uu___3 in - FStar_Compiler_Util.format "{depth=%s; task=%s}" uu___1 -let (string_of_repl_stack : repl_stack_entry_t Prims.list -> Prims.string) = - fun s -> - let uu___ = FStar_Compiler_List.map string_of_repl_stack_entry s in - FStar_String.concat ";\n\t\t" uu___ -let (repl_state_to_string : repl_state -> Prims.string) = - fun r -> - let uu___ = - let uu___1 = FStar_Compiler_Util.string_of_int r.repl_line in - let uu___2 = - let uu___3 = FStar_Compiler_Util.string_of_int r.repl_column in - let uu___4 = - let uu___5 = - let uu___6 = - match r.repl_curmod with - | FStar_Pervasives_Native.None -> "None" - | FStar_Pervasives_Native.Some m -> - FStar_Ident.string_of_lid m.FStar_Syntax_Syntax.name in - let uu___7 = - let uu___8 = string_of_repl_stack r.repl_deps_stack in [uu___8] in - uu___6 :: uu___7 in - (r.repl_fname) :: uu___5 in - uu___3 :: uu___4 in - uu___1 :: uu___2 in - FStar_Compiler_Util.format - "{\n\trepl_line=%s;\n\trepl_column=%s;\n\trepl_fname=%s;\n\trepl_cur_mod=%s;\n\t\\ \n repl_deps_stack={%s}\n}" - uu___ \ No newline at end of file +open Prims \ No newline at end of file diff --git a/src/basic/boot/FStar.Compiler.Util.fsti b/src/basic/boot/FStar.Compiler.Util.fsti index de17a546c87..ef4e86770a0 100644 --- a/src/basic/boot/FStar.Compiler.Util.fsti +++ b/src/basic/boot/FStar.Compiler.Util.fsti @@ -187,6 +187,7 @@ type stream_reader val open_stdin : unit -> stream_reader val read_line: stream_reader -> option string val nread : stream_reader -> int -> option string +val poll_stdin : float -> bool (* not relying on representation *) type string_builder diff --git a/src/fstar/FStar.Interactive.Ide.Types.fst b/src/fstar/FStar.Interactive.Ide.Types.fst index a1fd7594e42..6344b6f72d9 100755 --- a/src/fstar/FStar.Interactive.Ide.Types.fst +++ b/src/fstar/FStar.Interactive.Ide.Types.fst @@ -26,8 +26,6 @@ open FStar.Getopt open FStar.Ident open FStar.Errors open FStar.Interactive.JsonHelper -open FStar.Interactive.QueryHelper -open FStar.Interactive.PushHelper open FStar.Universal open FStar.TypeChecker.Env @@ -40,8 +38,8 @@ module DsEnv = FStar.Syntax.DsEnv module TcErr = FStar.TypeChecker.Err module TcEnv = FStar.TypeChecker.Env module CTable = FStar.Interactive.CompletionTable -module QH = FStar.Interactive.QueryHelper -module RS = FStar.Interactive.ReplState +module PI = FStar.Parser.ParseIt +module U = FStar.Compiler.Util (***********************) (* Global state setup *) @@ -50,72 +48,21 @@ let initial_range = Range.mk_range "" (Range.mk_pos 1 0) (Range.mk_pos 1 0) -(*****************************************) -(* Reading queries and writing responses *) -(*****************************************) - -let js_pushkind s : push_kind = match js_str s with - | "syntax" -> SyntaxCheck - | "lax" -> LaxCheck - | "full" -> FullCheck - | _ -> js_fail "push_kind" s - -let js_reductionrule s = match js_str s with - | "beta" -> FStar.TypeChecker.Env.Beta - | "delta" -> FStar.TypeChecker.Env.UnfoldUntil SS.delta_constant - | "iota" -> FStar.TypeChecker.Env.Iota - | "zeta" -> FStar.TypeChecker.Env.Zeta - | "reify" -> FStar.TypeChecker.Env.Reify - | "pure-subterms" -> FStar.TypeChecker.Env.PureSubtermsWithinComputations - | _ -> js_fail "reduction rule" s type completion_context = | CKCode | CKOption of bool (* #set-options (false) or #reset-options (true) *) | CKModuleOrNamespace of bool (* modules *) * bool (* namespaces *) -let js_optional_completion_context k = - match k with - | None -> CKCode - | Some k -> - match js_str k with - | "symbol" // Backwards compatibility - | "code" -> CKCode - | "set-options" -> CKOption false - | "reset-options" -> CKOption true - | "open" - | "let-open" -> CKModuleOrNamespace (true, true) - | "include" - | "module-alias" -> CKModuleOrNamespace (true, false) - | _ -> - js_fail "completion context (code, set-options, reset-options, \ -open, let-open, include, module-alias)" k - type lookup_context = | LKSymbolOnly | LKModule | LKOption | LKCode -let js_optional_lookup_context k = - match k with - | None -> LKSymbolOnly // Backwards-compatible default - | Some k -> - match js_str k with - | "symbol-only" -> LKSymbolOnly - | "code" -> LKCode - | "set-options" - | "reset-options" -> LKOption - | "open" - | "let-open" - | "include" - | "module-alias" -> LKModule - | _ -> - js_fail "lookup context (symbol-only, code, set-options, reset-options, \ -open, let-open, include, module-alias)" k - type position = string * int * int +type push_kind = | SyntaxCheck | LaxCheck | FullCheck type push_query = { push_kind: push_kind; @@ -128,7 +75,29 @@ type lookup_symbol_range = json type query_status = | QueryOK | QueryNOK | QueryViolatesProtocol -type callback_t = RS.repl_state -> (query_status * list json) * either RS.repl_state int +(* Types concerning repl *) +type repl_depth_t = TcEnv.tcenv_depth_t * int +type optmod_t = option Syntax.Syntax.modul + +type timed_fname = + { tf_fname: string; + tf_modtime: time } + +(** Every snapshot pushed in the repl stack is annotated with one of these. The +``LD``-prefixed (“Load Dependency”) onces are useful when loading or updating +dependencies, as they carry enough information to determine whether a dependency +is stale. **) +type repl_task = + | LDInterleaved of timed_fname * timed_fname (* (interface * implementation) *) + | LDSingle of timed_fname (* interface or implementation *) + | LDInterfaceOfCurrentFile of timed_fname (* interface *) + | PushFragment of either PI.input_frag FStar.Parser.AST.decl (* code fragment *) + | Noop (* Used by compute *) + +type full_buffer_request_kind = + | Full : full_buffer_request_kind + | Cache : full_buffer_request_kind + | ReloadDeps : full_buffer_request_kind type query' = | Exit @@ -144,10 +113,90 @@ type query' = | Search of string | GenericError of string | ProtocolViolation of string -| FullBuffer of string & bool //if true, check the buffer, otherwise just find verified prefix +| FullBuffer of string & full_buffer_request_kind | Callback of callback_t | Format of string +| Cancel of option position and query = { qq: query'; qid: string } +and callback_t = repl_state -> (query_status * list json) * either repl_state int +and repl_state = { + repl_line: int; + repl_column: int; + repl_fname: string; + repl_deps_stack: repl_stack_t; + repl_curmod: optmod_t; + repl_env: TcEnv.env; + repl_stdin: stream_reader; + repl_names: CTable.table; + repl_buffered_input_queries: list query +} +and repl_stack_t = list repl_stack_entry_t +and repl_stack_entry_t = repl_depth_t * (repl_task * repl_state) + +// Global repl_state, keeping state of different buffers +type grepl_state = { grepl_repls: U.psmap repl_state; grepl_stdin: stream_reader } + + +(*************************) +(* REPL tasks and states *) +(*************************) + +let t0 = Util.now () + +(** Create a timed_fname with a dummy modtime **) +let dummy_tf_of_fname fname = + { tf_fname = fname; + tf_modtime = t0 } + +let string_of_timed_fname { tf_fname = fname; tf_modtime = modtime } = + if modtime = t0 then Util.format1 "{ %s }" fname + else Util.format2 "{ %s; %s }" fname (string_of_time modtime) + +let string_of_repl_task = function + | LDInterleaved (intf, impl) -> + Util.format2 "LDInterleaved (%s, %s)" (string_of_timed_fname intf) (string_of_timed_fname impl) + | LDSingle intf_or_impl -> + Util.format1 "LDSingle %s" (string_of_timed_fname intf_or_impl) + | LDInterfaceOfCurrentFile intf -> + Util.format1 "LDInterfaceOfCurrentFile %s" (string_of_timed_fname intf) + | PushFragment (Inl frag) -> + Util.format1 "PushFragment { code = %s }" frag.frag_text + | PushFragment (Inr d) -> + Util.format1 "PushFragment { decl = %s }" (FStar.Parser.AST.decl_to_string d) + | Noop -> "Noop {}" + +module BU = FStar.Compiler.Util + +let string_of_repl_stack_entry + : repl_stack_entry_t -> string + = fun ((depth, i), (task, state)) -> + BU.format "{depth=%s; task=%s}" + [string_of_int i; + string_of_repl_task task] + + +let string_of_repl_stack s = + String.concat ";\n\t\t" + (List.map string_of_repl_stack_entry s) + +let repl_state_to_string (r:repl_state) + : string + = BU.format + "{\n\t\ + repl_line=%s;\n\t\ + repl_column=%s;\n\t\ + repl_fname=%s;\n\t\ + repl_cur_mod=%s;\n\t\ + repl_deps_stack={%s}\n\ + }" + [string_of_int r.repl_line; + string_of_int r.repl_column; + r.repl_fname; + (match r.repl_curmod with + | None -> "None" + | Some m -> Ident.string_of_lid m.name); + string_of_repl_stack r.repl_deps_stack] + let push_query_to_string pq = let pk = @@ -185,6 +234,7 @@ let query_to_string q = match q.qq with | FullBuffer _ -> "FullBuffer" | Callback _ -> "Callback" | Format _ -> "Format" +| Cancel _ -> "Cancel" let query_needs_current_module = function | Exit | DescribeProtocol | DescribeRepl | Segment _ @@ -201,7 +251,7 @@ let interactive_protocol_features = "lookup"; "lookup/context"; "lookup/documentation"; "lookup/definition"; "peek"; "pop"; "push"; "search"; "segment"; "vfs-add"; "tactic-ranges"; "interrupt"; "progress"; - "full-buffer"; "format"] + "full-buffer"; "format"; "cancel"] let json_of_issue_level i = JsonStr (match i with @@ -225,3 +275,57 @@ let json_of_issue issue = | Some r when def_range r <> use_range r -> [json_of_def_range r] | _ -> [])))] + +(*****************************************) +(* Reading queries and writing responses *) +(*****************************************) + +let js_pushkind s : push_kind = match js_str s with + | "syntax" -> SyntaxCheck + | "lax" -> LaxCheck + | "full" -> FullCheck + | _ -> js_fail "push_kind" s + +let js_reductionrule s = match js_str s with + | "beta" -> FStar.TypeChecker.Env.Beta + | "delta" -> FStar.TypeChecker.Env.UnfoldUntil SS.delta_constant + | "iota" -> FStar.TypeChecker.Env.Iota + | "zeta" -> FStar.TypeChecker.Env.Zeta + | "reify" -> FStar.TypeChecker.Env.Reify + | "pure-subterms" -> FStar.TypeChecker.Env.PureSubtermsWithinComputations + | _ -> js_fail "reduction rule" s + +let js_optional_completion_context k = + match k with + | None -> CKCode + | Some k -> + match js_str k with + | "symbol" // Backwards compatibility + | "code" -> CKCode + | "set-options" -> CKOption false + | "reset-options" -> CKOption true + | "open" + | "let-open" -> CKModuleOrNamespace (true, true) + | "include" + | "module-alias" -> CKModuleOrNamespace (true, false) + | _ -> + js_fail "completion context (code, set-options, reset-options, \ +open, let-open, include, module-alias)" k + +let js_optional_lookup_context k = + match k with + | None -> LKSymbolOnly // Backwards-compatible default + | Some k -> + match js_str k with + | "symbol-only" -> LKSymbolOnly + | "code" -> LKCode + | "set-options" + | "reset-options" -> LKOption + | "open" + | "let-open" + | "include" + | "module-alias" -> LKModule + | _ -> + js_fail "lookup context (symbol-only, code, set-options, reset-options, \ +open, let-open, include, module-alias)" k + diff --git a/src/fstar/FStar.Interactive.Ide.fst b/src/fstar/FStar.Interactive.Ide.fst index 0f81cc59a2d..3e4304c935d 100644 --- a/src/fstar/FStar.Interactive.Ide.fst +++ b/src/fstar/FStar.Interactive.Ide.fst @@ -218,6 +218,12 @@ let unpack_interactive_query json = assoc err "line" loc |> js_int, assoc err "column" loc |> js_int in + let parse_full_buffer_kind = function + | "full" -> Full + | "cache" -> Cache + | "reload-deps" -> ReloadDeps + | _ -> raise (InvalidQuery "Invalid full-buffer kind") + in { qid = qid; qq = match query with | "exit" -> Exit @@ -230,7 +236,7 @@ let unpack_interactive_query json = push_line = arg "line" |> js_int; push_column = arg "column" |> js_int; push_peek_only = query = "peek" }) - | "full-buffer" -> FullBuffer (arg "code" |> js_str, (arg "kind" |> js_str) = "full") + | "full-buffer" -> FullBuffer (arg "code" |> js_str, parse_full_buffer_kind (arg "kind" |> js_str)) | "autocomplete" -> AutoComplete (arg "partial-symbol" |> js_str, try_arg "context" |> js_optional_completion_context) | "lookup" -> Lookup (arg "symbol" |> js_str, @@ -264,11 +270,40 @@ let parse_interactive_query query_str : query = | None -> { qid = "?"; qq = ProtocolViolation "Json parsing failed." } | Some request -> deserialize_interactive_query request -let read_interactive_query stream : query = - match Util.read_line stream with - | None -> exit 0 - | Some line -> parse_interactive_query line - +let buffer_input_queries (st:repl_state) : repl_state = + let rec aux qs (st:repl_state) : repl_state = + let done qs st = + {st with repl_buffered_input_queries = + st.repl_buffered_input_queries @ List.rev qs} + in + if not (Util.poll_stdin (float_of_string "0.0")) + then done qs st + else ( + match Util.read_line st.repl_stdin with + | None -> + done qs st + + | Some line -> + let q = parse_interactive_query line in + match q.qq with + | Cancel _ -> + //Cancel drains all buffered queries + {st with repl_buffered_input_queries = [q] } + | _ -> aux (q :: qs) st + ) + in + aux [] st + +let read_interactive_query (st:repl_state) : query & repl_state = + match st.repl_buffered_input_queries with + | [] -> ( + match Util.read_line st.repl_stdin with + | None -> exit 0 + | Some line -> parse_interactive_query line, st + ) + | q :: qs -> + q, { st with repl_buffered_input_queries = qs } + let json_of_opt json_of_a opt_a = Util.dflt JsonNull (Util.map_option json_of_a opt_a) @@ -939,6 +974,7 @@ let rec fold_query (f:repl_state -> query -> run_query_result) let responses = responses @ resp in match status, st' with | QueryOK, Inl st -> + let st = buffer_input_queries st in fold_query f l st responses | _ -> (status, responses), st' @@ -1023,7 +1059,7 @@ let rec go st : int = FStar.Compiler.Util.print1 "Repl stack is:\n%s\n" (string_of_repl_stack (!repl_stack)) ); - let query = read_interactive_query st.repl_stdin in + let query, st = read_interactive_query st in let (status, responses), state_opt = validate_and_run_query st query in List.iter (write_response query.qid status) responses; match state_opt with @@ -1062,10 +1098,15 @@ let build_initial_repl_state (filename: string) = let env = init_env FStar.Parser.Dep.empty_deps in let env = FStar.TypeChecker.Env.set_range env initial_range in - { repl_line = 1; repl_column = 0; repl_fname = filename; - repl_curmod = None; repl_env = env; repl_deps_stack = []; + { repl_line = 1; + repl_column = 0; + repl_fname = filename; + repl_curmod = None; + repl_env = env; + repl_deps_stack = []; repl_stdin = open_stdin (); - repl_names = CompletionTable.empty } + repl_names = CompletionTable.empty; + repl_buffered_input_queries = [] } let interactive_mode' init_st = write_hello (); diff --git a/src/fstar/FStar.Interactive.Incremental.fst b/src/fstar/FStar.Interactive.Incremental.fst index 4c41a89cdda..366516a2316 100755 --- a/src/fstar/FStar.Interactive.Incremental.fst +++ b/src/fstar/FStar.Interactive.Incremental.fst @@ -124,7 +124,9 @@ let response_success (d:decl) ("query-id", JsonStr q); ("status", JsonStr "success"); ("contents", contents)]) - + +let repl_task (_, (p, _)) = p + let inspect_repl_stack (s:repl_stack_t) (ds:list decl) (write_full_buffer_fragment_progress: fragment_progress -> unit) @@ -177,6 +179,22 @@ let inspect_repl_stack (s:repl_stack_t) in matching_prefix entries ds +let reload_deps repl_stack = + let pop_until_deps entries + : qst (list query) + = match BU.prefix_until + (fun e -> match repl_task e with + | PushFragment _ | Noop -> false + | _ -> true) + entries + with + | None -> return [] + | Some (prefix, _, _) -> + let! pop = as_query Pop in + return (List.map (fun _ -> pop) prefix) + in + pop_until_deps repl_stack + let parse_code (code:string) = P.parse (Incremental { frag_fname = Range.file_of_range initial_range; @@ -200,7 +218,7 @@ let syntax_issue (raw_error, msg, range) = let run_full_buffer (st:repl_state) (qid:string) (code:string) - (full:bool) + (request_type:full_buffer_request_kind) (write_full_buffer_fragment_progress: fragment_progress -> unit) : list query = let parse_result = parse_code code in @@ -213,20 +231,26 @@ let run_full_buffer (st:repl_state) in let qs = match parse_result with - | IncrementalFragment (decls, _, err_opt) -> - let queries = - run_qst (inspect_repl_stack (!repl_stack) decls write_full_buffer_fragment_progress) qid - in - if full then log_syntax_issues err_opt; - if Options.debug_any() - then ( - BU.print1 "Generating queries\n%s\n" - (String.concat "\n" (List.map query_to_string queries)) - ); - if full then queries else [] + | IncrementalFragment (decls, _, err_opt) -> ( + match request_type with + | Full | Cache -> + let queries = + run_qst (inspect_repl_stack (!repl_stack) decls write_full_buffer_fragment_progress) qid + in + if request_type = Full then log_syntax_issues err_opt; + if Options.debug_any() + then ( + BU.print1 "Generating queries\n%s\n" + (String.concat "\n" (List.map query_to_string queries)) + ); + if request_type = Full then queries else [] + + | ReloadDeps -> + run_qst (reload_deps (!repl_stack)) qid + ) | ParseError err -> - if full then log_syntax_issues (Some err); + if request_type = Full then log_syntax_issues (Some err); [] | _ -> failwith "Unexpected parse result" @@ -252,4 +276,4 @@ let format_code (st:repl_state) (code:string) | ParseError err -> Inr [syntax_issue err] | _ -> - failwith "Unexpected parse result" \ No newline at end of file + failwith "Unexpected parse result" diff --git a/src/fstar/FStar.Interactive.Incremental.fsti b/src/fstar/FStar.Interactive.Incremental.fsti index 1064578f313..d8321f5443e 100755 --- a/src/fstar/FStar.Interactive.Incremental.fsti +++ b/src/fstar/FStar.Interactive.Incremental.fsti @@ -44,7 +44,7 @@ type fragment_progress = val run_full_buffer (st:repl_state) (qid:string) (code:string) - (full:bool) + (full:full_buffer_request_kind) (write_full_buffer_fragment_progress: fragment_progress -> unit) : list query diff --git a/src/fstar/FStar.Interactive.Lsp.fst b/src/fstar/FStar.Interactive.Lsp.fst index c8eac86789f..b6928251cbe 100644 --- a/src/fstar/FStar.Interactive.Lsp.fst +++ b/src/fstar/FStar.Interactive.Lsp.fst @@ -23,8 +23,8 @@ open FStar.Compiler.Util open FStar.Compiler.Range open FStar.Errors open FStar.Universal +open FStar.Interactive.Ide.Types open FStar.Interactive.JsonHelper -open FStar.Interactive.ReplState module U = FStar.Compiler.Util module QH = FStar.Interactive.QueryHelper @@ -113,9 +113,15 @@ let repl_state_init (fname: string) : repl_state = let intial_range = Range.mk_range fname (Range.mk_pos 1 0) (Range.mk_pos 1 0) in let env = init_env FStar.Parser.Dep.empty_deps in let env = TcEnv.set_range env intial_range in - { repl_line = 1; repl_column = 0; repl_fname = fname; - repl_curmod = None; repl_env = env; repl_deps_stack = []; - repl_stdin = open_stdin (); repl_names = CompletionTable.empty } + { repl_line = 1; + repl_column = 0; + repl_fname = fname; + repl_curmod = None; + repl_env = env; + repl_deps_stack = []; + repl_stdin = open_stdin (); + repl_names = CompletionTable.empty; + repl_buffered_input_queries = [] } type optresponse = option assoct // Contains [("result", ...)], [("error", ...)], but is not // the full response; call json_of_response for that diff --git a/src/fstar/FStar.Interactive.PushHelper.fst b/src/fstar/FStar.Interactive.PushHelper.fst index 3088970a2b9..ff872a02db2 100644 --- a/src/fstar/FStar.Interactive.PushHelper.fst +++ b/src/fstar/FStar.Interactive.PushHelper.fst @@ -30,7 +30,7 @@ open FStar.Universal open FStar.Parser.ParseIt open FStar.TypeChecker.Env open FStar.Interactive.JsonHelper -open FStar.Interactive.ReplState +open FStar.Interactive.Ide.Types module U = FStar.Compiler.Util module SS = FStar.Syntax.Syntax diff --git a/src/fstar/FStar.Interactive.PushHelper.fsti b/src/fstar/FStar.Interactive.PushHelper.fsti index 10221554b7c..49f4f1b0167 100644 --- a/src/fstar/FStar.Interactive.PushHelper.fsti +++ b/src/fstar/FStar.Interactive.PushHelper.fsti @@ -26,13 +26,12 @@ open FStar.Compiler.Util open FStar.Ident open FStar.TypeChecker.Env open FStar.Interactive.JsonHelper -open FStar.Interactive.ReplState +open FStar.Interactive.Ide.Types module DsEnv = FStar.Syntax.DsEnv module CTable = FStar.Interactive.CompletionTable module TcEnv = FStar.TypeChecker.Env -type push_kind = | SyntaxCheck | LaxCheck | FullCheck type ctx_depth_t = int * int * solver_depth_t * int type deps_t = FStar.Parser.Dep.deps type either_replst = either repl_state repl_state @@ -66,4 +65,4 @@ val commit_name_tracking : repl_state -> list name_tracking_event -> repl_state // Lax-check the whole file; used on didOpen and didSave // returns a diagnostic (only on error) along with the repl_state -val full_lax : string -> repl_state -> option assoct * repl_state +val full_lax : string -> repl_state -> option assoct * repl_state \ No newline at end of file diff --git a/src/fstar/FStar.Interactive.QueryHelper.fsti b/src/fstar/FStar.Interactive.QueryHelper.fsti index 85552a5e9ed..6c5a77413bd 100644 --- a/src/fstar/FStar.Interactive.QueryHelper.fsti +++ b/src/fstar/FStar.Interactive.QueryHelper.fsti @@ -24,7 +24,7 @@ open FStar.Compiler.Range open FStar.Compiler.Util open FStar.TypeChecker.Env open FStar.Interactive.JsonHelper -open FStar.Interactive.ReplState +open FStar.Interactive.Ide.Types module TcErr = FStar.TypeChecker.Err module TcEnv = FStar.TypeChecker.Env diff --git a/src/fstar/FStar.Interactive.ReplState.fst b/src/fstar/FStar.Interactive.ReplState.fst index ccb73adc1ba..362f1a890a7 100755 --- a/src/fstar/FStar.Interactive.ReplState.fst +++ b/src/fstar/FStar.Interactive.ReplState.fst @@ -21,7 +21,7 @@ open FStar.Errors open FStar.Compiler.Util open FStar.Compiler.Range open FStar.TypeChecker.Env - +(* module U = FStar.Compiler.Util module TcEnv = FStar.TypeChecker.Env module PI = FStar.Parser.ParseIt @@ -47,17 +47,22 @@ type repl_task = | PushFragment of either PI.input_frag FStar.Parser.AST.decl (* code fragment *) | Noop (* Used by compute *) -type repl_state = { repl_line: int; repl_column: int; repl_fname: string; - repl_deps_stack: repl_stack_t; - repl_curmod: optmod_t; - repl_env: TcEnv.env; - repl_stdin: stream_reader; - repl_names: CTable.table } -and repl_stack_t = list repl_stack_entry_t -and repl_stack_entry_t = repl_depth_t * (repl_task * repl_state) +type repl_state (a:Type) = { + repl_line: int; + repl_column: int; + repl_fname: string; + repl_deps_stack: repl_stack_t a; + repl_curmod: optmod_t; + repl_env: TcEnv.env; + repl_stdin: stream_reader; + repl_names: CTable.table; + repl_buffered_input_queries: list a +} +and repl_stack_t a = list (repl_stack_entry_t a) +and repl_stack_entry_t a = repl_depth_t * (repl_task * repl_state a) // Global repl_state, keeping state of different buffers -type grepl_state = { grepl_repls: U.psmap repl_state; grepl_stdin: stream_reader } +type grepl_state = { grepl_repls: U.psmap (repl_state unit); grepl_stdin: stream_reader } (*************************) @@ -91,7 +96,7 @@ let string_of_repl_task = function module BU = FStar.Compiler.Util let string_of_repl_stack_entry - : repl_stack_entry_t -> string + : repl_stack_entry_t _ -> string = fun ((depth, i), (task, state)) -> BU.format "{depth=%s; task=%s}" [string_of_int i; @@ -102,7 +107,7 @@ let string_of_repl_stack s = String.concat ";\n\t\t" (List.map string_of_repl_stack_entry s) -let repl_state_to_string (r:repl_state) +let repl_state_to_string (r:repl_state 'a) : string = BU.format "{\n\t\ @@ -119,3 +124,4 @@ let repl_state_to_string (r:repl_state) | None -> "None" | Some m -> Ident.string_of_lid m.name); string_of_repl_stack r.repl_deps_stack] +*) \ No newline at end of file From 857a5684d471ffcc62a54c2bd2b279bdca2234d5 Mon Sep 17 00:00:00 2001 From: Nikhil Swamy Date: Sun, 19 Mar 2023 21:35:52 -0700 Subject: [PATCH 16/48] Recording the raw document content to echo back in incremental IDE requets; handling cancel requests' --- ocaml/fstar-lib/FStar_Parser_ParseIt.ml | 45 ++++++-- ocaml/fstar-lib/FStar_Parser_ParseIt.mli | 7 +- ocaml/fstar-lib/FStar_Sedlexing.ml | 2 + .../generated/FStar_Interactive_Ide.ml | 108 ++++++++++++++++-- .../generated/FStar_Interactive_Ide_Types.ml | 12 +- .../FStar_Interactive_Incremental.ml | 89 +++++++++------ .../generated/FStar_Parser_Driver.ml | 11 +- src/fstar/FStar.Interactive.Ide.Types.fst | 10 +- src/fstar/FStar.Interactive.Ide.fst | 53 +++++++-- src/fstar/FStar.Interactive.Incremental.fst | 15 +-- src/fstar/FStar.Interactive.Incremental.fsti | 2 +- src/parser/FStar.Parser.Driver.fst | 2 +- src/parser/FStar.Parser.Driver.fsti | 1 + src/parser/FStar.Parser.ParseIt.fsti | 7 +- 14 files changed, 279 insertions(+), 85 deletions(-) diff --git a/ocaml/fstar-lib/FStar_Parser_ParseIt.ml b/ocaml/fstar-lib/FStar_Parser_ParseIt.ml index 2d5badc2da2..9bc10ea9f5f 100644 --- a/ocaml/fstar-lib/FStar_Parser_ParseIt.ml +++ b/ocaml/fstar-lib/FStar_Parser_ParseIt.ml @@ -99,9 +99,15 @@ type parse_frag = type parse_error = (Codes.raw_error * string * FStar_Compiler_Range.range) + +type code_fragment = { + range: FStar_Compiler_Range.range; + code: string; +} + type parse_result = | ASTFragment of (FStar_Parser_AST.inputFragment * (string * FStar_Compiler_Range.range) list) - | IncrementalFragment of (FStar_Parser_AST.decl list * (string * FStar_Compiler_Range.range) list * parse_error option) + | IncrementalFragment of ((FStar_Parser_AST.decl * code_fragment) list * (string * FStar_Compiler_Range.range) list * parse_error option) | Term of FStar_Parser_AST.term | ParseError of parse_error @@ -109,33 +115,48 @@ let parse fn = FStar_Parser_Util.warningHandler := (function | e -> Printf.printf "There was some warning (TODO)\n"); - let lexbuf, filename = match fn with + let lexbuf, filename, contents = match fn with | Filename f -> check_extension f; let f', contents = read_file f in - (try create contents f' 1 0, f' + (try create contents f' 1 0, f', contents with _ -> raise_err (Fatal_InvalidUTF8Encoding, U.format1 "File %s has invalid UTF-8 encoding.\n" f')) | Incremental s | Toplevel s | Fragment s -> - create s.frag_text s.frag_fname (Z.to_int s.frag_line) (Z.to_int s.frag_col), "" + create s.frag_text s.frag_fname (Z.to_int s.frag_line) (Z.to_int s.frag_col), "", s.frag_text in let lexer () = let tok = FStar_Parser_LexFStar.token lexbuf in (tok, lexbuf.start_p, lexbuf.cur_p) in + let range_of_positions start fin = + let start_pos = FStar_Parser_Util.pos_of_lexpos start in + let end_pos = FStar_Parser_Util.pos_of_lexpos fin in + FStar_Compiler_Range.mk_range filename start_pos end_pos + in let err_of_parse_error () = - let pos = FStar_Parser_Util.pos_of_lexpos lexbuf.cur_p in - let r = FStar_Compiler_Range.mk_range filename pos pos in - Fatal_SyntaxError, "Syntax error", r + let pos = lexbuf.cur_p in + Fatal_SyntaxError, + "Syntax error", + range_of_positions pos pos in - let parse_incremental_decls () = let parse_one_decl = MenhirLib.Convert.Simplified.traditional2revised FStar_Parser_Parse.oneDeclOrEOF in + let contents_at = + let lines = U.splitlines contents in + fun (start_pos:Lexing.position) (end_pos:Lexing.position) -> + let suffix = FStar_Compiler_Util.nth_tail (Z.of_int (start_pos.pos_lnum - 1)) lines in + let text, _ = FStar_Compiler_Util.first_N (Z.of_int (end_pos.pos_lnum - start_pos.pos_lnum)) suffix in + let range = range_of_positions start_pos end_pos in + { range; + code = FStar_String.concat "\n" text } + in let open FStar_Pervasives in let rec parse decls = let _, _, r = err_of_parse_error () in + let start_pos = current_pos lexbuf in let d = try FStar_Ident.reset_gensym(); @@ -144,7 +165,7 @@ let parse fn = | FStar_Errors.Error(e, msg, r, _ctx) -> Inr (e, msg, r) - | Parsing.Parse_error as _e -> + | Parsing.Parse_error as _e -> Inr (err_of_parse_error ()) in match d with @@ -152,7 +173,9 @@ let parse fn = | Inl (Some d) -> if not (FStar_Parser_AST.decl_syntax_is_delimited d) then rollback lexbuf; - parse (d::decls) + let end_pos = current_pos lexbuf in + let raw_contents = contents_at start_pos end_pos in + parse ((d, raw_contents)::decls) | Inr err -> List.rev decls, Some err in parse [] @@ -161,7 +184,7 @@ let parse fn = let decls, err_opt = parse_incremental_decls () in match err_opt with | None -> - FStar_Parser_AST.as_frag decls + FStar_Parser_AST.as_frag (List.map fst decls) | Some (e, msg, r) -> raise (FStar_Errors.Error(e, msg, r, [])) in diff --git a/ocaml/fstar-lib/FStar_Parser_ParseIt.mli b/ocaml/fstar-lib/FStar_Parser_ParseIt.mli index d0284cc18bd..8f2db314a6f 100755 --- a/ocaml/fstar-lib/FStar_Parser_ParseIt.mli +++ b/ocaml/fstar-lib/FStar_Parser_ParseIt.mli @@ -26,9 +26,14 @@ type parse_frag = type parse_error = (Codes.raw_error * string * FStar_Compiler_Range.range) +type code_fragment = { + range : FStar_Compiler_Range.range; + code: string; +} + type parse_result = | ASTFragment of (FStar_Parser_AST.inputFragment * (string * FStar_Compiler_Range.range) list) - | IncrementalFragment of (FStar_Parser_AST.decl list * (string * FStar_Compiler_Range.range) list * parse_error option) + | IncrementalFragment of ((FStar_Parser_AST.decl * code_fragment) list * (string * FStar_Compiler_Range.range) list * parse_error option) | Term of FStar_Parser_AST.term | ParseError of parse_error diff --git a/ocaml/fstar-lib/FStar_Sedlexing.ml b/ocaml/fstar-lib/FStar_Sedlexing.ml index d7f00a2f578..1c4fe22f2f9 100644 --- a/ocaml/fstar-lib/FStar_Sedlexing.ml +++ b/ocaml/fstar-lib/FStar_Sedlexing.ml @@ -53,6 +53,8 @@ let create (s:string) fn loffset coffset = mark_val = 0; } +let current_pos b = b.start_p + let start b = b.mark <- b.cur; b.mark_val <- (-1); diff --git a/ocaml/fstar-lib/generated/FStar_Interactive_Ide.ml b/ocaml/fstar-lib/generated/FStar_Interactive_Ide.ml index 8a5f0754ad8..7984ba9497a 100644 --- a/ocaml/fstar-lib/generated/FStar_Interactive_Ide.ml +++ b/ocaml/fstar-lib/generated/FStar_Interactive_Ide.ml @@ -513,6 +513,20 @@ let (unpack_interactive_query : FStar_Compiler_Effect.op_Bar_Greater uu___3 FStar_Interactive_JsonHelper.js_str in FStar_Interactive_Ide_Types.Format uu___2 + | "cancel" -> + let uu___2 = + let uu___3 = + let uu___4 = + let uu___5 = arg "cancel-line" in + FStar_Compiler_Effect.op_Bar_Greater uu___5 + FStar_Interactive_JsonHelper.js_int in + let uu___5 = + let uu___6 = arg "cancel-column" in + FStar_Compiler_Effect.op_Bar_Greater uu___6 + FStar_Interactive_JsonHelper.js_int in + ("", uu___4, uu___5) in + FStar_Pervasives_Native.Some uu___3 in + FStar_Interactive_Ide_Types.Cancel uu___2 | uu___2 -> let uu___3 = FStar_Compiler_Util.format1 "Unknown query '%s'" query in @@ -1361,6 +1375,17 @@ let (rephrase_dependency_error : FStar_Errors.issue -> FStar_Errors.issue) = let (write_full_buffer_fragment_progress : FStar_Interactive_Incremental.fragment_progress -> unit) = fun di -> + let json_of_code_fragment cf = + let uu___ = + let uu___1 = + let uu___2 = + FStar_Compiler_Range.json_of_def_range + cf.FStar_Parser_ParseIt.range in + ("range", uu___2) in + [uu___1; + ("code", + (FStar_Compiler_Util.JsonStr (cf.FStar_Parser_ParseIt.code)))] in + FStar_Compiler_Util.JsonAssoc uu___ in match di with | FStar_Interactive_Incremental.FragmentStarted d -> let uu___ = @@ -1372,14 +1397,19 @@ let (write_full_buffer_fragment_progress : [uu___1] in write_progress (FStar_Pervasives_Native.Some "full-buffer-fragment-started") uu___ - | FStar_Interactive_Incremental.FragmentSuccess d -> + | FStar_Interactive_Incremental.FragmentSuccess (d, cf) -> let uu___ = let uu___1 = let uu___2 = FStar_Compiler_Range.json_of_def_range d.FStar_Parser_AST.drange in ("ranges", uu___2) in - [uu___1] in + let uu___2 = + let uu___3 = + let uu___4 = json_of_code_fragment cf in + ("code-fragment", uu___4) in + [uu___3] in + uu___1 :: uu___2 in write_progress (FStar_Pervasives_Native.Some "full-buffer-fragment-ok") uu___ | FStar_Interactive_Incremental.FragmentFailed d -> @@ -1560,7 +1590,8 @@ let (run_push_without_deps : FStar_Parser_ParseIt.frag_line = line; FStar_Parser_ParseIt.frag_col = column } - | FStar_Pervasives.Inr decl -> FStar_Pervasives.Inr decl in + | FStar_Pervasives.Inr (decl, _code) -> + FStar_Pervasives.Inr decl in let st1 = set_nosynth_flag st peek_only in let uu___2 = run_repl_transaction st1 push_kind peek_only @@ -1581,14 +1612,15 @@ let (run_push_without_deps : | FStar_Errors.ENotImplemented -> true | uu___3 -> false) errs in ((match code_or_decl with - | FStar_Pervasives.Inr d -> + | FStar_Pervasives.Inr ds -> if Prims.op_Negation has_error then write_full_buffer_fragment_progress - (FStar_Interactive_Incremental.FragmentSuccess d) + (FStar_Interactive_Incremental.FragmentSuccess ds) else write_full_buffer_fragment_progress - (FStar_Interactive_Incremental.FragmentFailed d) + (FStar_Interactive_Incremental.FragmentFailed + (FStar_Pervasives_Native.fst ds)) | uu___4 -> ()); (let json_errors = let uu___4 = @@ -2085,7 +2117,10 @@ let run_with_parsed_and_tc_term : (FStar_Parser_ParseIt.Incremental frag) in match uu___ with | FStar_Parser_ParseIt.IncrementalFragment - (decls, uu___1, _err) -> FStar_Pervasives_Native.Some decls + (decls, uu___1, _err) -> + let uu___2 = + FStar_Compiler_List.map FStar_Pervasives_Native.fst decls in + FStar_Pervasives_Native.Some uu___2 | uu___1 -> FStar_Pervasives_Native.None in let desugar env decls = let uu___ = @@ -2468,6 +2503,58 @@ type run_query_result = ((FStar_Interactive_Ide_Types.query_status * FStar_Compiler_Util.json Prims.list) * (FStar_Interactive_Ide_Types.repl_state, Prims.int) FStar_Pervasives.either) +let (maybe_cancel_queries : + FStar_Interactive_Ide_Types.repl_state -> + FStar_Interactive_Ide_Types.query Prims.list -> + (FStar_Interactive_Ide_Types.query Prims.list * + FStar_Interactive_Ide_Types.repl_state)) + = + fun st -> + fun l -> + match st.FStar_Interactive_Ide_Types.repl_buffered_input_queries with + | { + FStar_Interactive_Ide_Types.qq = FStar_Interactive_Ide_Types.Cancel + p; + FStar_Interactive_Ide_Types.qid = uu___;_}::rest -> + let st1 = + { + FStar_Interactive_Ide_Types.repl_line = + (st.FStar_Interactive_Ide_Types.repl_line); + FStar_Interactive_Ide_Types.repl_column = + (st.FStar_Interactive_Ide_Types.repl_column); + FStar_Interactive_Ide_Types.repl_fname = + (st.FStar_Interactive_Ide_Types.repl_fname); + FStar_Interactive_Ide_Types.repl_deps_stack = + (st.FStar_Interactive_Ide_Types.repl_deps_stack); + FStar_Interactive_Ide_Types.repl_curmod = + (st.FStar_Interactive_Ide_Types.repl_curmod); + FStar_Interactive_Ide_Types.repl_env = + (st.FStar_Interactive_Ide_Types.repl_env); + FStar_Interactive_Ide_Types.repl_stdin = + (st.FStar_Interactive_Ide_Types.repl_stdin); + FStar_Interactive_Ide_Types.repl_names = + (st.FStar_Interactive_Ide_Types.repl_names); + FStar_Interactive_Ide_Types.repl_buffered_input_queries = rest + } in + (match p with + | FStar_Pervasives_Native.None -> ([], st1) + | FStar_Pervasives_Native.Some p1 -> + let query_ahead_of p2 q = + let uu___1 = p2 in + match uu___1 with + | (uu___2, l1, c) -> + (match q.FStar_Interactive_Ide_Types.qq with + | FStar_Interactive_Ide_Types.Push pq -> + pq.FStar_Interactive_Ide_Types.push_line >= l1 + | uu___3 -> false) in + let l1 = + let uu___1 = + FStar_Compiler_Util.prefix_until (query_ahead_of p1) l in + match uu___1 with + | FStar_Pervasives_Native.None -> l + | FStar_Pervasives_Native.Some (l2, uu___2, uu___3) -> l2 in + (l1, st1)) + | uu___ -> (l, st) let rec (fold_query : (FStar_Interactive_Ide_Types.repl_state -> FStar_Interactive_Ide_Types.query -> run_query_result) @@ -2493,7 +2580,9 @@ let rec (fold_query : | (FStar_Interactive_Ide_Types.QueryOK, FStar_Pervasives.Inl st1) -> let st2 = buffer_input_queries st1 in - fold_query f l1 st2 responses1 + let uu___1 = maybe_cancel_queries st2 l1 in + (match uu___1 with + | (l2, st3) -> fold_query f l2 st3 responses1) | uu___1 -> ((status, responses1), st'))) let (validate_query : FStar_Interactive_Ide_Types.repl_state -> @@ -2579,6 +2668,9 @@ let rec (run_query : | FStar_Interactive_Ide_Types.Callback f -> f st | FStar_Interactive_Ide_Types.Format code -> let uu___ = run_format_code st code in as_json_list uu___ + | FStar_Interactive_Ide_Types.Cancel uu___ -> + ((FStar_Interactive_Ide_Types.QueryOK, []), + (FStar_Pervasives.Inl st)) and (validate_and_run_query : FStar_Interactive_Ide_Types.repl_state -> FStar_Interactive_Ide_Types.query -> run_query_result) diff --git a/ocaml/fstar-lib/generated/FStar_Interactive_Ide_Types.ml b/ocaml/fstar-lib/generated/FStar_Interactive_Ide_Types.ml index 88ecda335bb..59ef2d6e3c0 100644 --- a/ocaml/fstar-lib/generated/FStar_Interactive_Ide_Types.ml +++ b/ocaml/fstar-lib/generated/FStar_Interactive_Ide_Types.ml @@ -53,7 +53,10 @@ type push_query = push_column: Prims.int ; push_peek_only: Prims.bool ; push_code_or_decl: - (Prims.string, FStar_Parser_AST.decl) FStar_Pervasives.either } + (Prims.string, + (FStar_Parser_AST.decl * FStar_Parser_ParseIt.code_fragment)) + FStar_Pervasives.either + } let (__proj__Mkpush_query__item__push_kind : push_query -> push_kind) = fun projectee -> match projectee with @@ -75,7 +78,10 @@ let (__proj__Mkpush_query__item__push_peek_only : push_query -> Prims.bool) = | { push_kind = push_kind1; push_line; push_column; push_peek_only; push_code_or_decl;_} -> push_peek_only let (__proj__Mkpush_query__item__push_code_or_decl : - push_query -> (Prims.string, FStar_Parser_AST.decl) FStar_Pervasives.either) + push_query -> + (Prims.string, + (FStar_Parser_AST.decl * FStar_Parser_ParseIt.code_fragment)) + FStar_Pervasives.either) = fun projectee -> match projectee with @@ -430,7 +436,7 @@ let (push_query_to_string : push_query -> Prims.string) = let code_or_decl = match pq.push_code_or_decl with | FStar_Pervasives.Inl code -> code - | FStar_Pervasives.Inr decl -> FStar_Parser_AST.decl_to_string decl in + | FStar_Pervasives.Inr (_decl, code) -> code.FStar_Parser_ParseIt.code in let uu___ = let uu___1 = let uu___2 = FStar_Compiler_Util.string_of_int pq.push_line in diff --git a/ocaml/fstar-lib/generated/FStar_Interactive_Incremental.ml b/ocaml/fstar-lib/generated/FStar_Interactive_Incremental.ml index 7effb4eefed..046efebbc9f 100644 --- a/ocaml/fstar-lib/generated/FStar_Interactive_Incremental.ml +++ b/ocaml/fstar-lib/generated/FStar_Interactive_Incremental.ml @@ -1,7 +1,8 @@ open Prims type fragment_progress = | FragmentStarted of FStar_Parser_AST.decl - | FragmentSuccess of FStar_Parser_AST.decl + | FragmentSuccess of (FStar_Parser_AST.decl * + FStar_Parser_ParseIt.code_fragment) | FragmentFailed of FStar_Parser_AST.decl | FragmentError of FStar_Errors.issue Prims.list let (uu___is_FragmentStarted : fragment_progress -> Prims.bool) = @@ -14,8 +15,9 @@ let (uu___is_FragmentSuccess : fragment_progress -> Prims.bool) = fun projectee -> match projectee with | FragmentSuccess _0 -> true | uu___ -> false let (__proj__FragmentSuccess__item___0 : - fragment_progress -> FStar_Parser_AST.decl) = - fun projectee -> match projectee with | FragmentSuccess _0 -> _0 + fragment_progress -> + (FStar_Parser_AST.decl * FStar_Parser_ParseIt.code_fragment)) + = fun projectee -> match projectee with | FragmentSuccess _0 -> _0 let (uu___is_FragmentFailed : fragment_progress -> Prims.bool) = fun projectee -> match projectee with | FragmentFailed _0 -> true | uu___ -> false @@ -81,41 +83,46 @@ let (as_query : return uu___1) let (push_decl : (fragment_progress -> unit) -> - FStar_Parser_AST.decl -> FStar_Interactive_Ide_Types.query Prims.list qst) + (FStar_Parser_AST.decl * FStar_Parser_ParseIt.code_fragment) -> + FStar_Interactive_Ide_Types.query Prims.list qst) = fun write_full_buffer_fragment_progress -> - fun d -> - let pq = - let uu___ = + fun ds -> + let uu___ = ds in + match uu___ with + | (d, s) -> + let pq = + let uu___1 = + let uu___2 = + FStar_Compiler_Range.start_of_range d.FStar_Parser_AST.drange in + FStar_Compiler_Range.line_of_pos uu___2 in + let uu___2 = + let uu___3 = + FStar_Compiler_Range.start_of_range d.FStar_Parser_AST.drange in + FStar_Compiler_Range.col_of_pos uu___3 in + { + FStar_Interactive_Ide_Types.push_kind = + FStar_Interactive_Ide_Types.FullCheck; + FStar_Interactive_Ide_Types.push_line = uu___1; + FStar_Interactive_Ide_Types.push_column = uu___2; + FStar_Interactive_Ide_Types.push_peek_only = false; + FStar_Interactive_Ide_Types.push_code_or_decl = + (FStar_Pervasives.Inr ds) + } in + let progress st = + write_full_buffer_fragment_progress (FragmentStarted d); + ((FStar_Interactive_Ide_Types.QueryOK, []), + (FStar_Pervasives.Inl st)) in let uu___1 = - FStar_Compiler_Range.start_of_range d.FStar_Parser_AST.drange in - FStar_Compiler_Range.line_of_pos uu___1 in - let uu___1 = - let uu___2 = - FStar_Compiler_Range.start_of_range d.FStar_Parser_AST.drange in - FStar_Compiler_Range.col_of_pos uu___2 in - { - FStar_Interactive_Ide_Types.push_kind = - FStar_Interactive_Ide_Types.FullCheck; - FStar_Interactive_Ide_Types.push_line = uu___; - FStar_Interactive_Ide_Types.push_column = uu___1; - FStar_Interactive_Ide_Types.push_peek_only = false; - FStar_Interactive_Ide_Types.push_code_or_decl = - (FStar_Pervasives.Inr d) - } in - let progress st = - write_full_buffer_fragment_progress (FragmentStarted d); - ((FStar_Interactive_Ide_Types.QueryOK, []), - (FStar_Pervasives.Inl st)) in - let uu___ = as_query (FStar_Interactive_Ide_Types.Callback progress) in - op_let_Bang uu___ - (fun cb -> - let uu___1 = as_query (FStar_Interactive_Ide_Types.Push pq) in - op_let_Bang uu___1 (fun push -> return [cb; push])) + as_query (FStar_Interactive_Ide_Types.Callback progress) in + op_let_Bang uu___1 + (fun cb -> + let uu___2 = as_query (FStar_Interactive_Ide_Types.Push pq) in + op_let_Bang uu___2 (fun push -> return [cb; push])) let (push_decls : (fragment_progress -> unit) -> - FStar_Parser_AST.decl Prims.list -> - FStar_Interactive_Ide_Types.query Prims.list qst) + (FStar_Parser_AST.decl * FStar_Parser_ParseIt.code_fragment) Prims.list + -> FStar_Interactive_Ide_Types.query Prims.list qst) = fun write_full_buffer_fragment_progress -> fun ds -> @@ -152,7 +159,8 @@ let repl_task : fun uu___ -> match uu___ with | (uu___1, (p, uu___2)) -> p let (inspect_repl_stack : FStar_Interactive_Ide_Types.repl_stack_t -> - FStar_Parser_AST.decl Prims.list -> + (FStar_Parser_AST.decl * FStar_Parser_ParseIt.code_fragment) Prims.list + -> (fragment_progress -> unit) -> FStar_Interactive_Ide_Types.query Prims.list qst) = @@ -196,7 +204,9 @@ let (inspect_repl_stack : (FStar_Compiler_List.op_At pops pushes))) | FStar_Interactive_Ide_Types.PushFragment (FStar_Pervasives.Inr d') -> - let uu___1 = FStar_Parser_AST_Comparison.eq_decl d d' in + let uu___1 = + FStar_Parser_AST_Comparison.eq_decl + (FStar_Pervasives_Native.fst d) d' in if uu___1 then (write_full_buffer_fragment_progress @@ -399,9 +409,12 @@ let (format_code : let formatted_code = let uu___1 = FStar_Compiler_List.map - (fun d -> - let uu___2 = FStar_Parser_ToDocument.decl_to_document d in - doc_to_string uu___2) decls in + (fun uu___2 -> + match uu___2 with + | (d, uu___3) -> + let uu___4 = + FStar_Parser_ToDocument.decl_to_document d in + doc_to_string uu___4) decls in FStar_Compiler_Effect.op_Bar_Greater uu___1 (FStar_String.concat "\n\n") in FStar_Pervasives.Inl formatted_code diff --git a/ocaml/fstar-lib/generated/FStar_Parser_Driver.ml b/ocaml/fstar-lib/generated/FStar_Parser_Driver.ml index b07a889a5d1..f14930bc8bd 100644 --- a/ocaml/fstar-lib/generated/FStar_Parser_Driver.ml +++ b/ocaml/fstar-lib/generated/FStar_Parser_Driver.ml @@ -6,6 +6,8 @@ type fragment = | Empty | Modul of FStar_Parser_AST.modul | Decls of FStar_Parser_AST.decl Prims.list + | DeclsWithContent of (FStar_Parser_AST.decl * + FStar_Parser_ParseIt.code_fragment) Prims.list let (uu___is_Empty : fragment -> Prims.bool) = fun projectee -> match projectee with | Empty -> true | uu___ -> false let (uu___is_Modul : fragment -> Prims.bool) = @@ -16,6 +18,13 @@ let (uu___is_Decls : fragment -> Prims.bool) = fun projectee -> match projectee with | Decls _0 -> true | uu___ -> false let (__proj__Decls__item___0 : fragment -> FStar_Parser_AST.decl Prims.list) = fun projectee -> match projectee with | Decls _0 -> _0 +let (uu___is_DeclsWithContent : fragment -> Prims.bool) = + fun projectee -> + match projectee with | DeclsWithContent _0 -> true | uu___ -> false +let (__proj__DeclsWithContent__item___0 : + fragment -> + (FStar_Parser_AST.decl * FStar_Parser_ParseIt.code_fragment) Prims.list) + = fun projectee -> match projectee with | DeclsWithContent _0 -> _0 let (parse_fragment : FStar_Parser_ParseIt.input_frag -> fragment) = fun frag -> let uu___ = @@ -28,7 +37,7 @@ let (parse_fragment : FStar_Parser_ParseIt.input_frag -> fragment) = | FStar_Parser_ParseIt.ASTFragment (FStar_Pervasives.Inr decls, uu___1) -> Decls decls | FStar_Parser_ParseIt.IncrementalFragment (decls, uu___1, uu___2) -> - Decls decls + DeclsWithContent decls | FStar_Parser_ParseIt.ParseError (e, msg, r) -> FStar_Errors.raise_error (e, msg) r | FStar_Parser_ParseIt.Term uu___1 -> diff --git a/src/fstar/FStar.Interactive.Ide.Types.fst b/src/fstar/FStar.Interactive.Ide.Types.fst index 6344b6f72d9..417702138af 100755 --- a/src/fstar/FStar.Interactive.Ide.Types.fst +++ b/src/fstar/FStar.Interactive.Ide.Types.fst @@ -65,11 +65,15 @@ type position = string * int * int type push_kind = | SyntaxCheck | LaxCheck | FullCheck type push_query = - { push_kind: push_kind; + { + push_kind: push_kind; push_line: int; push_column: int; push_peek_only: bool; - push_code_or_decl: either string FStar.Parser.AST.decl} + //Either a string: Just the raw content of a document fragment + //Or a parsed document fragment and the raw content it corresponds to + push_code_or_decl: either string (FStar.Parser.AST.decl * PI.code_fragment) + } type lookup_symbol_range = json @@ -208,7 +212,7 @@ let push_query_to_string pq = let code_or_decl = match pq.push_code_or_decl with | Inl code -> code - | Inr decl -> FStar.Parser.AST.decl_to_string decl + | Inr (_decl, code) -> code.code in FStar.Compiler.Util.format "{ push_kind = %s; push_line = %s; \ push_column = %s; push_peek_only = %s; push_code_or_decl = %s }" diff --git a/src/fstar/FStar.Interactive.Ide.fst b/src/fstar/FStar.Interactive.Ide.fst index 3e4304c935d..b3fb90a1255 100644 --- a/src/fstar/FStar.Interactive.Ide.fst +++ b/src/fstar/FStar.Interactive.Ide.fst @@ -100,7 +100,7 @@ If `must_rollback` is set, always pop. Returns a pair: a boolean indicating success, and a new REPL state. **) let run_repl_transaction st push_kind must_rollback task = let st = push_repl "run_repl_transaction" push_kind task st in - let env, finish_name_tracking = track_name_changes st.repl_env in // begin name tracking… + let env, finish_name_tracking = track_name_changes st.repl_env in // begin name tracking … let check_success () = get_err_count () = 0 && not must_rollback in @@ -112,7 +112,7 @@ let run_repl_transaction st push_kind must_rollback task = | Some (curmod, env) when check_success () -> curmod, env, true | _ -> st.repl_curmod, env, false in - let env, name_events = finish_name_tracking env in // …end name tracking + let env, name_events = finish_name_tracking env in // …end name tracking let st = if success then let st = { st with repl_env = env; repl_curmod = curmod } in @@ -127,7 +127,7 @@ let run_repl_transaction st push_kind must_rollback task = Returns a new REPL state, wrapped in ``Inl`` to indicate complete success or ``Inr`` to indicate a partial failure. That new state has an updated ``repl_deps_stack``, containing loaded dependencies in reverse order compared to -the original list of tasks: the first dependencies (prims, …) come first; the +the original list of tasks: the first dependencies (prims, ...) come first; the current file's interface comes last. The original value of the ``repl_deps_stack`` field in ``st`` is used to skip @@ -253,6 +253,7 @@ let unpack_interactive_query json = | "vfs-add" -> VfsAdd (try_arg "filename" |> Util.map_option js_str, arg "contents" |> js_str) | "format" -> Format (arg "code" |> js_str) + | "cancel" -> Cancel (Some("", arg "cancel-line" |> js_int, arg "cancel-column" |> js_int)) | _ -> ProtocolViolation (Util.format1 "Unknown query '%s'" query) } with | InvalidQuery msg -> { qid = qid; qq = ProtocolViolation msg } @@ -587,13 +588,18 @@ let rephrase_dependency_error issue = let write_full_buffer_fragment_progress (di:Incremental.fragment_progress) = let open FStar.Interactive.Incremental in + let json_of_code_fragment (cf:FStar.Parser.ParseIt.code_fragment) = + JsonAssoc ["range", json_of_def_range cf.range; + "code", JsonStr cf.code] + in match di with | FragmentStarted d -> write_progress (Some "full-buffer-fragment-started") ["ranges", json_of_def_range d.FStar.Parser.AST.drange] - | FragmentSuccess d -> + | FragmentSuccess (d, cf) -> write_progress (Some "full-buffer-fragment-ok") - ["ranges", json_of_def_range d.FStar.Parser.AST.drange] + ["ranges", json_of_def_range d.FStar.Parser.AST.drange; + "code-fragment", json_of_code_fragment cf] | FragmentFailed d -> write_progress (Some "full-buffer-fragment-failed") ["ranges", json_of_def_range d.FStar.Parser.AST.drange] @@ -627,7 +633,7 @@ let run_push_without_deps st query match code_or_decl with | Inl text -> Inl { frag_fname = ""; frag_text = text; frag_line = line; frag_col = column } - | Inr decl -> + | Inr (decl, _code) -> Inr decl in let st = set_nosynth_flag st peek_only in @@ -646,10 +652,10 @@ let run_push_without_deps st query in let _ = match code_or_decl with - | Inr d -> + | Inr ds -> if not has_error - then write_full_buffer_fragment_progress (Incremental.FragmentSuccess d) - else write_full_buffer_fragment_progress (Incremental.FragmentFailed d) + then write_full_buffer_fragment_progress (Incremental.FragmentSuccess ds) + else write_full_buffer_fragment_progress (Incremental.FragmentFailed (fst ds)) | _ -> () in let json_errors = JsonList (errs |> List.map json_of_issue) in @@ -792,7 +798,7 @@ let run_with_parsed_and_tc_term st term line column continuation = let parse frag = match FStar.Parser.ParseIt.parse (FStar.Parser.ParseIt.Incremental frag) with - | FStar.Parser.ParseIt.IncrementalFragment (decls, _, _err) -> Some decls + | FStar.Parser.ParseIt.IncrementalFragment (decls, _, _err) -> Some (List.map fst decls) | _ -> None in let desugar env decls = @@ -962,6 +968,29 @@ let as_json_list (q: (query_status & json) & either repl_state int) let run_query_result = (query_status * list json) * either repl_state int +let maybe_cancel_queries st l = + match st.repl_buffered_input_queries with + | { qq = Cancel p } :: rest -> ( + let st = { st with repl_buffered_input_queries = rest } in + match p with + | None -> //If no range, then cancel all remaining queries + [], st + | Some p -> //Cancel all queries that are within the range + let query_ahead_of p q = + let _, l, c = p in + match q.qq with + | Push pq -> pq.push_line >= l + | _ -> false + in + let l = + match BU.prefix_until (query_ahead_of p) l with + | None -> l + | Some (l, _, _) -> l + in + l, st + ) + | _ -> l, st + let rec fold_query (f:repl_state -> query -> run_query_result) (l:list query) (st:repl_state) @@ -975,6 +1004,7 @@ let rec fold_query (f:repl_state -> query -> run_query_result) match status, st' with | QueryOK, Inl st -> let st = buffer_input_queries st in + let l, st = maybe_cancel_queries st l in fold_query f l st responses | _ -> (status, responses), st' @@ -1017,6 +1047,9 @@ let rec run_query st (q: query) : (query_status * list json) * either repl_state f st | Format code -> as_json_list (run_format_code st code) + | Cancel _ -> + //This should be handled in the fold_query loop above + (QueryOK, []), Inl st and validate_and_run_query st query = let query = validate_query st query in repl_current_qid := Some query.qid; diff --git a/src/fstar/FStar.Interactive.Incremental.fst b/src/fstar/FStar.Interactive.Incremental.fst index 366516a2316..267c0d37c5b 100755 --- a/src/fstar/FStar.Interactive.Incremental.fst +++ b/src/fstar/FStar.Interactive.Incremental.fst @@ -88,15 +88,16 @@ let as_query (q:query') } let push_decl (write_full_buffer_fragment_progress: fragment_progress -> unit) - (d:decl) + (ds:decl * code_fragment) : qst (list query) = let open FStar.Compiler.Range in + let d, s = ds in let pq = { push_kind = FullCheck; push_line = line_of_pos (start_of_range d.drange); push_column = col_of_pos (start_of_range d.drange); push_peek_only = false; - push_code_or_decl = Inr d + push_code_or_decl = Inr ds } in let progress st = write_full_buffer_fragment_progress (FragmentStarted d); @@ -107,7 +108,7 @@ let push_decl (write_full_buffer_fragment_progress: fragment_progress -> unit) return [cb; push] let push_decls (write_full_buffer_fragment_progress : fragment_progress -> unit) - (ds:list decl) + (ds:list (decl & code_fragment)) : qst (list query) = let! qs = map (push_decl write_full_buffer_fragment_progress) ds in return (List.flatten qs) @@ -128,7 +129,7 @@ let response_success (d:decl) let repl_task (_, (p, _)) = p let inspect_repl_stack (s:repl_stack_t) - (ds:list decl) + (ds:list (decl & code_fragment)) (write_full_buffer_fragment_progress: fragment_progress -> unit) : qst (list query) // & list json) = let entries = List.rev s in @@ -144,7 +145,7 @@ let inspect_repl_stack (s:repl_stack_t) | Some (prefix, first_push, rest) -> let entries = first_push :: rest in let repl_task (_, (p, _)) = p in - let rec matching_prefix entries ds + let rec matching_prefix entries (ds:list (decl & code_fragment)) : qst (list query) = match entries, ds with | [], [] -> @@ -159,7 +160,7 @@ let inspect_repl_stack (s:repl_stack_t) let! pushes = push_decls (d::ds) in return (pops @ pushes) | PushFragment (Inr d') -> - if eq_decl d d' + if eq_decl (fst d) d' then ( write_full_buffer_fragment_progress (FragmentSuccess d); matching_prefix entries ds @@ -266,7 +267,7 @@ let format_code (st:repl_state) (code:string) in let formatted_code = List.map - (fun d -> doc_to_string (FStar.Parser.ToDocument.decl_to_document d)) + (fun (d, _) -> doc_to_string (FStar.Parser.ToDocument.decl_to_document d)) decls |> String.concat "\n\n" in diff --git a/src/fstar/FStar.Interactive.Incremental.fsti b/src/fstar/FStar.Interactive.Incremental.fsti index d8321f5443e..19638da7e99 100755 --- a/src/fstar/FStar.Interactive.Incremental.fsti +++ b/src/fstar/FStar.Interactive.Incremental.fsti @@ -27,7 +27,7 @@ open FStar.Interactive.Ide.Types type fragment_progress = | FragmentStarted of decl - | FragmentSuccess of decl + | FragmentSuccess of (decl & FStar.Parser.ParseIt.code_fragment) | FragmentFailed of decl | FragmentError of list issue diff --git a/src/parser/FStar.Parser.Driver.fst b/src/parser/FStar.Parser.Driver.fst index ebda66f3f0b..da9ab27b881 100644 --- a/src/parser/FStar.Parser.Driver.fst +++ b/src/parser/FStar.Parser.Driver.fst @@ -36,7 +36,7 @@ let parse_fragment (frag: ParseIt.input_frag) : fragment = | ASTFragment (Inr decls, _) -> //interactive mode: more decls Decls decls | IncrementalFragment (decls, _, _) -> - Decls decls + DeclsWithContent decls | ParseError (e, msg, r) -> raise_error (e, msg) r | Term _ -> diff --git a/src/parser/FStar.Parser.Driver.fsti b/src/parser/FStar.Parser.Driver.fsti index 7a0a695291b..fec9a908730 100644 --- a/src/parser/FStar.Parser.Driver.fsti +++ b/src/parser/FStar.Parser.Driver.fsti @@ -25,6 +25,7 @@ type fragment = | Empty | Modul of AST.modul // an entire module or interface -- unspecified | Decls of list AST.decl // a partial set of declarations + | DeclsWithContent of list (AST.decl & ParseIt.code_fragment) val parse_fragment : ParseIt.input_frag -> fragment diff --git a/src/parser/FStar.Parser.ParseIt.fsti b/src/parser/FStar.Parser.ParseIt.fsti index 9ad31c188d2..593c1026987 100644 --- a/src/parser/FStar.Parser.ParseIt.fsti +++ b/src/parser/FStar.Parser.ParseIt.fsti @@ -44,9 +44,14 @@ type parse_frag = type parse_error = (Errors.raw_error * string * Range.range) +type code_fragment = { + code: string; + range: FStar.Compiler.Range.range; +} + type parse_result = | ASTFragment of (AST.inputFragment * list (string * Range.range)) - | IncrementalFragment of (list AST.decl & list (string * Range.range) & option parse_error) + | IncrementalFragment of (list (AST.decl & code_fragment) & list (string * Range.range) & option parse_error) | Term of AST.term | ParseError of parse_error From 08855459c2642917815ccf914d55d6f16d96feb0 Mon Sep 17 00:00:00 2001 From: Aseem Rastogi Date: Mon, 20 Mar 2023 14:46:57 +0530 Subject: [PATCH 17/48] Const_reify syntax change --- src/parser/FStar.Parser.AST.Comparison.fst | 6 +++++- 1 file changed, 5 insertions(+), 1 deletion(-) diff --git a/src/parser/FStar.Parser.AST.Comparison.fst b/src/parser/FStar.Parser.AST.Comparison.fst index 21a2def23b3..b9cb3382cf5 100755 --- a/src/parser/FStar.Parser.AST.Comparison.fst +++ b/src/parser/FStar.Parser.AST.Comparison.fst @@ -44,6 +44,10 @@ let eq_option (f: 'a -> 'a -> bool) (t1 t2:option 'a) | Some t1, Some t2 -> f t1 t2 | _ -> false + +// +// TODO: There is an eq_const in FStar.Const.fst, can we use that? +// let eq_sconst (c1 c2: sconst) : bool = match c1, c2 with | Const_effect, Const_effect -> true @@ -54,7 +58,7 @@ let eq_sconst (c1 c2: sconst) : bool = | Const_string (s1, _), Const_string (s2, _) -> s1 = s2 | Const_real s1, Const_real s2 -> s1 = s2 | Const_range r1, Const_range r2 -> r1 = r2 - | Const_reify, Const_reify -> true + | Const_reify _, Const_reify _ -> true | Const_reflect l1, Const_reflect l2 -> Ident.lid_equals l1 l2 | _ -> false From 5fcec9dfab8098e8f481604d6bdf4a854a049a4e Mon Sep 17 00:00:00 2001 From: Aseem Rastogi Date: Mon, 20 Mar 2023 14:47:00 +0530 Subject: [PATCH 18/48] snap --- .../FStar_InteractiveHelpers_Base.ml | 4 +- .../FStar_InteractiveHelpers_Effectful.ml | 21 +- .../FStar_InteractiveHelpers_ExploreTerm.ml | 265 +- .../FStar_InteractiveHelpers_PostProcess.ml | 109 +- .../generated/FStar_Parser_AST_Comparison.ml | 3 +- .../fstar-lib/generated/FStar_Parser_Parse.ml | 2 +- ocaml/fstar-lib/generated/FStar_Tactics_BV.ml | 8 +- .../generated/FStar_Tactics_Canon.ml | 8 +- .../FStar_Tactics_CanonCommMonoid.ml | 68 +- .../FStar_Tactics_CanonCommMonoidSimple.ml | 188 +- ...tar_Tactics_CanonCommMonoidSimple_Equiv.ml | 99 +- .../FStar_Tactics_CanonCommSemiring.ml | 56 +- .../generated/FStar_Tactics_CanonMonoid.ml | 148 +- .../generated/FStar_Tactics_Derived.ml | 1062 ++++--- .../generated/FStar_Tactics_Effect.ml | 4 +- .../generated/FStar_Tactics_Logic.ml | 225 +- .../FStar_Tactics_PatternMatching.ml | 678 ++--- .../generated/FStar_Tactics_Simplifier.ml | 324 +- .../generated/FStar_Tactics_SyntaxHelpers.ml | 6 +- .../generated/FStar_Tactics_Typeclasses.ml | 73 +- .../generated/Steel_Effect_Common.ml | 2688 ++++++++--------- .../generated/Steel_ST_GenElim_Base.ml | 498 +-- 22 files changed, 3146 insertions(+), 3391 deletions(-) diff --git a/ocaml/fstar-lib/generated/FStar_InteractiveHelpers_Base.ml b/ocaml/fstar-lib/generated/FStar_InteractiveHelpers_Base.ml index e8b1b42a385..bea6412242a 100644 --- a/ocaml/fstar-lib/generated/FStar_InteractiveHelpers_Base.ml +++ b/ocaml/fstar-lib/generated/FStar_InteractiveHelpers_Base.ml @@ -946,9 +946,7 @@ let (__proj__MetaAnalysis__item__uu___ : Prims.exn -> Prims.string) = fun projectee -> match projectee with | MetaAnalysis uu___ -> uu___ let mfail : 'uuuuu . Prims.string -> ('uuuuu, unit) FStar_Tactics_Effect.tac_repr = - fun uu___ -> - (fun str -> Obj.magic (FStar_Tactics_Effect.raise (MetaAnalysis str))) - uu___ + fun str -> FStar_Tactics_Effect.raise (MetaAnalysis str) let (print_dbg : Prims.bool -> Prims.string -> (unit, unit) FStar_Tactics_Effect.tac_repr) = fun uu___1 -> diff --git a/ocaml/fstar-lib/generated/FStar_InteractiveHelpers_Effectful.ml b/ocaml/fstar-lib/generated/FStar_InteractiveHelpers_Effectful.ml index b6e37006c5d..d113c2ed518 100644 --- a/ocaml/fstar-lib/generated/FStar_InteractiveHelpers_Effectful.ml +++ b/ocaml/fstar-lib/generated/FStar_InteractiveHelpers_Effectful.ml @@ -1431,20 +1431,13 @@ let (compute_eterm_info : parameters)))) uu___2))) uu___2)) (fun uu___1 -> - (fun uu___1 -> - match uu___1 with - | FStar_Tactics_Common.TacticFailure msg -> - Obj.magic - (Obj.repr - (FStar_InteractiveHelpers_Base.mfail - (Prims.strcat - "compute_eterm_info: failure: '" - (Prims.strcat msg "'")))) - | e1 -> - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.raise e1))) - uu___1))) uu___) + match uu___1 with + | FStar_Tactics_Common.TacticFailure msg -> + FStar_InteractiveHelpers_Base.mfail + (Prims.strcat + "compute_eterm_info: failure: '" + (Prims.strcat msg "'")) + | e1 -> FStar_Tactics_Effect.raise e1))) uu___) let (has_refinement : FStar_InteractiveHelpers_ExploreTerm.type_info -> Prims.bool) = fun ty -> diff --git a/ocaml/fstar-lib/generated/FStar_InteractiveHelpers_ExploreTerm.ml b/ocaml/fstar-lib/generated/FStar_InteractiveHelpers_ExploreTerm.ml index fa2b8a6733f..7428ba9e9f6 100644 --- a/ocaml/fstar-lib/generated/FStar_InteractiveHelpers_ExploreTerm.ml +++ b/ocaml/fstar-lib/generated/FStar_InteractiveHelpers_ExploreTerm.ml @@ -1556,21 +1556,13 @@ let rec (inst_comp : match () with | () -> inst_comp_once e c t) (fun uu___ -> - (fun uu___ -> - match uu___ with - | FStar_InteractiveHelpers_Base.MetaAnalysis - msg -> - Obj.magic - (Obj.repr - (FStar_InteractiveHelpers_Base.mfail - (Prims.strcat - "inst_comp: error: " - msg))) - | err -> - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.raise - err))) uu___))) + match uu___ with + | FStar_InteractiveHelpers_Base.MetaAnalysis + msg -> + FStar_InteractiveHelpers_Base.mfail + (Prims.strcat "inst_comp: error: " + msg) + | err -> FStar_Tactics_Effect.raise err))) (fun uu___ -> (fun c' -> Obj.magic (inst_comp e c' tl')) uu___)))) uu___2 uu___1 uu___ @@ -1669,68 +1661,54 @@ let (_abs_update_typ : "_abs_update_typ: inconsistent state")) uu___1))) uu___1)) (fun uu___ -> - (fun uu___ -> - match uu___ with - | FStar_InteractiveHelpers_Base.MetaAnalysis msg -> - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.InteractiveHelpers.ExploreTerm.fst" - (Prims.of_int (303)) (Prims.of_int (10)) - (Prims.of_int (303)) (Prims.of_int (93))) - (Prims.mk_range - "FStar.InteractiveHelpers.ExploreTerm.fst" - (Prims.of_int (303)) (Prims.of_int (4)) - (Prims.of_int (303)) (Prims.of_int (93))) - (Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.InteractiveHelpers.ExploreTerm.fst" - (Prims.of_int (303)) - (Prims.of_int (61)) - (Prims.of_int (303)) - (Prims.of_int (92))) - (Prims.mk_range "prims.fst" - (Prims.of_int (606)) - (Prims.of_int (19)) - (Prims.of_int (606)) - (Prims.of_int (31))) - (Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.InteractiveHelpers.ExploreTerm.fst" - (Prims.of_int (303)) - (Prims.of_int (61)) - (Prims.of_int (303)) - (Prims.of_int (78))) - (Prims.mk_range "prims.fst" - (Prims.of_int (606)) - (Prims.of_int (19)) - (Prims.of_int (606)) - (Prims.of_int (31))) - (Obj.magic - (FStar_Tactics_Builtins.term_to_string - ty)) - (fun uu___1 -> - FStar_Tactics_Effect.lift_div_tac - (fun uu___2 -> - Prims.strcat uu___1 - (Prims.strcat ":\n" msg))))) - (fun uu___1 -> - FStar_Tactics_Effect.lift_div_tac - (fun uu___2 -> - Prims.strcat - "_abs_update_typ: could not find an arrow in: " - uu___1)))) - (fun uu___1 -> + match uu___ with + | FStar_InteractiveHelpers_Base.MetaAnalysis msg -> + FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.InteractiveHelpers.ExploreTerm.fst" + (Prims.of_int (303)) (Prims.of_int (10)) + (Prims.of_int (303)) (Prims.of_int (93))) + (Prims.mk_range + "FStar.InteractiveHelpers.ExploreTerm.fst" + (Prims.of_int (303)) (Prims.of_int (4)) + (Prims.of_int (303)) (Prims.of_int (93))) + (Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.InteractiveHelpers.ExploreTerm.fst" + (Prims.of_int (303)) (Prims.of_int (61)) + (Prims.of_int (303)) (Prims.of_int (92))) + (Prims.mk_range "prims.fst" (Prims.of_int (606)) + (Prims.of_int (19)) (Prims.of_int (606)) + (Prims.of_int (31))) + (Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.InteractiveHelpers.ExploreTerm.fst" + (Prims.of_int (303)) (Prims.of_int (61)) + (Prims.of_int (303)) (Prims.of_int (78))) + (Prims.mk_range "prims.fst" + (Prims.of_int (606)) (Prims.of_int (19)) + (Prims.of_int (606)) (Prims.of_int (31))) + (Obj.magic + (FStar_Tactics_Builtins.term_to_string ty)) (fun uu___1 -> - Obj.magic - (FStar_InteractiveHelpers_Base.mfail - uu___1)) uu___1))) - | err -> - Obj.magic (Obj.repr (FStar_Tactics_Effect.raise err))) - uu___) + FStar_Tactics_Effect.lift_div_tac + (fun uu___2 -> + Prims.strcat uu___1 + (Prims.strcat ":\n" msg))))) + (fun uu___1 -> + FStar_Tactics_Effect.lift_div_tac + (fun uu___2 -> + Prims.strcat + "_abs_update_typ: could not find an arrow in: " + uu___1)))) + (fun uu___1 -> + (fun uu___1 -> + Obj.magic + (FStar_InteractiveHelpers_Base.mfail uu___1)) + uu___1) + | err -> FStar_Tactics_Effect.raise err) let (abs_update_typ_or_comp : FStar_Reflection_Types.binder -> typ_or_comp -> @@ -1802,12 +1780,14 @@ let (abs_update_opt_typ_or_comp : | FStar_InteractiveHelpers_Base.MetaAnalysis msg -> Obj.magic - (FStar_Tactics_Effect.lift_div_tac - (fun uu___1 -> - FStar_Pervasives_Native.None)) + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___1 -> + FStar_Pervasives_Native.None))) | err -> Obj.magic - (FStar_Tactics_Effect.raise err)) + (Obj.repr + (FStar_Tactics_Effect.raise err))) uu___)))) uu___2 uu___1 uu___ let rec (_flush_typ_or_comp_comp : Prims.bool -> @@ -2335,77 +2315,62 @@ let (flush_typ_or_comp : uu___1) | TC_Comp (c, pl, n) -> flush_comp pl n c)) (fun uu___ -> - (fun uu___ -> - match uu___ with - | FStar_InteractiveHelpers_Base.MetaAnalysis msg - -> - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.InteractiveHelpers.ExploreTerm.fst" - (Prims.of_int (379)) - (Prims.of_int (15)) - (Prims.of_int (379)) - (Prims.of_int (90))) - (Prims.mk_range - "FStar.InteractiveHelpers.ExploreTerm.fst" - (Prims.of_int (379)) - (Prims.of_int (9)) - (Prims.of_int (379)) - (Prims.of_int (90))) - (Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.InteractiveHelpers.ExploreTerm.fst" - (Prims.of_int (379)) - (Prims.of_int (50)) - (Prims.of_int (379)) - (Prims.of_int (89))) - (Prims.mk_range "prims.fst" - (Prims.of_int (606)) - (Prims.of_int (19)) - (Prims.of_int (606)) - (Prims.of_int (31))) - (Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.InteractiveHelpers.ExploreTerm.fst" - (Prims.of_int (379)) - (Prims.of_int (50)) - (Prims.of_int (379)) - (Prims.of_int (75))) - (Prims.mk_range - "prims.fst" - (Prims.of_int (606)) - (Prims.of_int (19)) - (Prims.of_int (606)) - (Prims.of_int (31))) - (Obj.magic - (typ_or_comp_to_string - tyc)) - (fun uu___1 -> - FStar_Tactics_Effect.lift_div_tac - (fun uu___2 -> - Prims.strcat - uu___1 - (Prims.strcat - ":\n" msg))))) - (fun uu___1 -> - FStar_Tactics_Effect.lift_div_tac - (fun uu___2 -> - Prims.strcat - "flush_typ_or_comp failed on: " - uu___1)))) - (fun uu___1 -> + match uu___ with + | FStar_InteractiveHelpers_Base.MetaAnalysis msg -> + FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.InteractiveHelpers.ExploreTerm.fst" + (Prims.of_int (379)) (Prims.of_int (15)) + (Prims.of_int (379)) (Prims.of_int (90))) + (Prims.mk_range + "FStar.InteractiveHelpers.ExploreTerm.fst" + (Prims.of_int (379)) (Prims.of_int (9)) + (Prims.of_int (379)) (Prims.of_int (90))) + (Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.InteractiveHelpers.ExploreTerm.fst" + (Prims.of_int (379)) + (Prims.of_int (50)) + (Prims.of_int (379)) + (Prims.of_int (89))) + (Prims.mk_range "prims.fst" + (Prims.of_int (606)) + (Prims.of_int (19)) + (Prims.of_int (606)) + (Prims.of_int (31))) + (Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.InteractiveHelpers.ExploreTerm.fst" + (Prims.of_int (379)) + (Prims.of_int (50)) + (Prims.of_int (379)) + (Prims.of_int (75))) + (Prims.mk_range "prims.fst" + (Prims.of_int (606)) + (Prims.of_int (19)) + (Prims.of_int (606)) + (Prims.of_int (31))) + (Obj.magic + (typ_or_comp_to_string tyc)) (fun uu___1 -> - Obj.magic - (FStar_InteractiveHelpers_Base.mfail - uu___1)) uu___1))) - | err -> - Obj.magic - (Obj.repr (FStar_Tactics_Effect.raise err))) - uu___))) uu___) + FStar_Tactics_Effect.lift_div_tac + (fun uu___2 -> + Prims.strcat uu___1 + (Prims.strcat ":\n" msg))))) + (fun uu___1 -> + FStar_Tactics_Effect.lift_div_tac + (fun uu___2 -> + Prims.strcat + "flush_typ_or_comp failed on: " + uu___1)))) + (fun uu___1 -> + (fun uu___1 -> + Obj.magic + (FStar_InteractiveHelpers_Base.mfail + uu___1)) uu___1) + | err -> FStar_Tactics_Effect.raise err))) uu___) let (safe_arg_typ_or_comp : Prims.bool -> FStar_Reflection_Types.env -> diff --git a/ocaml/fstar-lib/generated/FStar_InteractiveHelpers_PostProcess.ml b/ocaml/fstar-lib/generated/FStar_InteractiveHelpers_PostProcess.ml index a2430d7f851..a3e884e6b41 100644 --- a/ocaml/fstar-lib/generated/FStar_InteractiveHelpers_PostProcess.ml +++ b/ocaml/fstar-lib/generated/FStar_InteractiveHelpers_PostProcess.ml @@ -2146,29 +2146,22 @@ let (pp_analyze_effectful_term : (fun uu___2 -> Obj.magic (end_proof ())) uu___2))) uu___2)) (fun uu___1 -> - (fun uu___1 -> - match uu___1 with - | FStar_InteractiveHelpers_Base.MetaAnalysis msg -> - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.InteractiveHelpers.PostProcess.fst" - (Prims.of_int (256)) (Prims.of_int (29)) - (Prims.of_int (256)) (Prims.of_int (49))) - (Prims.mk_range - "FStar.InteractiveHelpers.PostProcess.fst" - (Prims.of_int (256)) (Prims.of_int (51)) - (Prims.of_int (256)) (Prims.of_int (63))) - (Obj.magic - (FStar_InteractiveHelpers_Output.printout_failure - msg)) - (fun uu___2 -> - (fun uu___2 -> Obj.magic (end_proof ())) - uu___2))) - | err -> - Obj.magic (Obj.repr (FStar_Tactics_Effect.raise err))) - uu___1) + match uu___1 with + | FStar_InteractiveHelpers_Base.MetaAnalysis msg -> + FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.InteractiveHelpers.PostProcess.fst" + (Prims.of_int (256)) (Prims.of_int (29)) + (Prims.of_int (256)) (Prims.of_int (49))) + (Prims.mk_range + "FStar.InteractiveHelpers.PostProcess.fst" + (Prims.of_int (256)) (Prims.of_int (51)) + (Prims.of_int (256)) (Prims.of_int (63))) + (Obj.magic + (FStar_InteractiveHelpers_Output.printout_failure msg)) + (fun uu___2 -> + (fun uu___2 -> Obj.magic (end_proof ())) uu___2) + | err -> FStar_Tactics_Effect.raise err) let _ = FStar_Tactics_Native.register_tactic "FStar.InteractiveHelpers.PostProcess.pp_analyze_effectful_term" @@ -2612,27 +2605,20 @@ let (pp_split_assert_conjs : (fun uu___2 -> Obj.magic (end_proof ())) uu___2))) uu___2)) (fun uu___1 -> - (fun uu___1 -> - match uu___1 with - | FStar_InteractiveHelpers_Base.MetaAnalysis msg -> - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.InteractiveHelpers.PostProcess.fst" - (Prims.of_int (356)) (Prims.of_int (29)) - (Prims.of_int (356)) (Prims.of_int (49))) - (Prims.mk_range - "FStar.InteractiveHelpers.PostProcess.fst" - (Prims.of_int (356)) (Prims.of_int (51)) - (Prims.of_int (356)) (Prims.of_int (63))) - (Obj.magic - (FStar_InteractiveHelpers_Output.printout_failure - msg)) - (fun uu___2 -> - (fun uu___2 -> Obj.magic (end_proof ())) uu___2))) - | err -> Obj.magic (Obj.repr (FStar_Tactics_Effect.raise err))) - uu___1) + match uu___1 with + | FStar_InteractiveHelpers_Base.MetaAnalysis msg -> + FStar_Tactics_Effect.tac_bind + (Prims.mk_range "FStar.InteractiveHelpers.PostProcess.fst" + (Prims.of_int (356)) (Prims.of_int (29)) + (Prims.of_int (356)) (Prims.of_int (49))) + (Prims.mk_range "FStar.InteractiveHelpers.PostProcess.fst" + (Prims.of_int (356)) (Prims.of_int (51)) + (Prims.of_int (356)) (Prims.of_int (63))) + (Obj.magic + (FStar_InteractiveHelpers_Output.printout_failure msg)) + (fun uu___2 -> + (fun uu___2 -> Obj.magic (end_proof ())) uu___2) + | err -> FStar_Tactics_Effect.raise err) let _ = FStar_Tactics_Native.register_tactic "FStar.InteractiveHelpers.PostProcess.pp_split_assert_conjs" @@ -6927,27 +6913,20 @@ let (pp_unfold_in_assert_or_assume : (fun uu___2 -> Obj.magic (end_proof ())) uu___2))) uu___2)) (fun uu___1 -> - (fun uu___1 -> - match uu___1 with - | FStar_InteractiveHelpers_Base.MetaAnalysis msg -> - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.InteractiveHelpers.PostProcess.fst" - (Prims.of_int (748)) (Prims.of_int (29)) - (Prims.of_int (748)) (Prims.of_int (49))) - (Prims.mk_range - "FStar.InteractiveHelpers.PostProcess.fst" - (Prims.of_int (748)) (Prims.of_int (51)) - (Prims.of_int (748)) (Prims.of_int (63))) - (Obj.magic - (FStar_InteractiveHelpers_Output.printout_failure - msg)) - (fun uu___2 -> - (fun uu___2 -> Obj.magic (end_proof ())) uu___2))) - | err -> Obj.magic (Obj.repr (FStar_Tactics_Effect.raise err))) - uu___1) + match uu___1 with + | FStar_InteractiveHelpers_Base.MetaAnalysis msg -> + FStar_Tactics_Effect.tac_bind + (Prims.mk_range "FStar.InteractiveHelpers.PostProcess.fst" + (Prims.of_int (748)) (Prims.of_int (29)) + (Prims.of_int (748)) (Prims.of_int (49))) + (Prims.mk_range "FStar.InteractiveHelpers.PostProcess.fst" + (Prims.of_int (748)) (Prims.of_int (51)) + (Prims.of_int (748)) (Prims.of_int (63))) + (Obj.magic + (FStar_InteractiveHelpers_Output.printout_failure msg)) + (fun uu___2 -> + (fun uu___2 -> Obj.magic (end_proof ())) uu___2) + | err -> FStar_Tactics_Effect.raise err) let _ = FStar_Tactics_Native.register_tactic "FStar.InteractiveHelpers.PostProcess.pp_unfold_in_assert_or_assume" diff --git a/ocaml/fstar-lib/generated/FStar_Parser_AST_Comparison.ml b/ocaml/fstar-lib/generated/FStar_Parser_AST_Comparison.ml index 43ddb6bcd56..1f13b917564 100644 --- a/ocaml/fstar-lib/generated/FStar_Parser_AST_Comparison.ml +++ b/ocaml/fstar-lib/generated/FStar_Parser_AST_Comparison.ml @@ -39,7 +39,8 @@ let (eq_sconst : FStar_Const.sconst -> FStar_Const.sconst -> Prims.bool) = (s2, uu___1)) -> s1 = s2 | (FStar_Const.Const_real s1, FStar_Const.Const_real s2) -> s1 = s2 | (FStar_Const.Const_range r1, FStar_Const.Const_range r2) -> r1 = r2 - | (FStar_Const.Const_reify, FStar_Const.Const_reify) -> true + | (FStar_Const.Const_reify uu___, FStar_Const.Const_reify uu___1) -> + true | (FStar_Const.Const_reflect l1, FStar_Const.Const_reflect l2) -> FStar_Ident.lid_equals l1 l2 | uu___ -> false diff --git a/ocaml/fstar-lib/generated/FStar_Parser_Parse.ml b/ocaml/fstar-lib/generated/FStar_Parser_Parse.ml index 972b81479b8..3887a8d6259 100644 --- a/ocaml/fstar-lib/generated/FStar_Parser_Parse.ml +++ b/ocaml/fstar-lib/generated/FStar_Parser_Parse.ml @@ -12280,7 +12280,7 @@ let es = ( xs ) in Obj.repr( # 3134 "parse.mly" (let _1 = () in - ( Const_reify )) + ( Const_reify None )) # 12271 "parse.ml" : 'constant)) ; (fun __caml_parser_env -> diff --git a/ocaml/fstar-lib/generated/FStar_Tactics_BV.ml b/ocaml/fstar-lib/generated/FStar_Tactics_BV.ml index 496adf26e40..cc7174ea77a 100644 --- a/ocaml/fstar-lib/generated/FStar_Tactics_BV.ml +++ b/ocaml/fstar-lib/generated/FStar_Tactics_BV.ml @@ -892,9 +892,11 @@ let (arith_to_bv_tac : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = "arith_to_bv_tac: unexpected: " uu___4)))) (fun uu___4 -> - FStar_Tactics_Derived.fail - uu___4))) uu___3))) - uu___3))) uu___2)) + (fun uu___4 -> + Obj.magic + (FStar_Tactics_Derived.fail + uu___4)) uu___4))) + uu___3))) uu___3))) uu___2)) let (bv_tac : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = fun uu___ -> FStar_Tactics_Derived.focus diff --git a/ocaml/fstar-lib/generated/FStar_Tactics_Canon.ml b/ocaml/fstar-lib/generated/FStar_Tactics_Canon.ml index ef14886be48..93e57a67c65 100644 --- a/ocaml/fstar-lib/generated/FStar_Tactics_Canon.ml +++ b/ocaml/fstar-lib/generated/FStar_Tactics_Canon.ml @@ -1028,8 +1028,10 @@ let (canon_point_entry : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) "impossible: " uu___4)))) (fun uu___4 -> - FStar_Tactics_Derived.fail - uu___4))) uu___2))) uu___2))) - uu___1) + (fun uu___4 -> + Obj.magic + (FStar_Tactics_Derived.fail + uu___4)) uu___4))) + uu___2))) uu___2))) uu___1) let (canon : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = fun uu___ -> FStar_Tactics_Derived.pointwise canon_point_entry \ No newline at end of file diff --git a/ocaml/fstar-lib/generated/FStar_Tactics_CanonCommMonoid.ml b/ocaml/fstar-lib/generated/FStar_Tactics_CanonCommMonoid.ml index a6482cacc8d..d71fd4339c0 100644 --- a/ocaml/fstar-lib/generated/FStar_Tactics_CanonCommMonoid.ml +++ b/ocaml/fstar-lib/generated/FStar_Tactics_CanonCommMonoid.ml @@ -1669,30 +1669,28 @@ let canon_monoid_aux : t), t1, t2) -> Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.CanonCommMonoid.fst" - (Prims.of_int (340)) - (Prims.of_int (9)) - (Prims.of_int (340)) - (Prims.of_int (25))) - (Prims.mk_range - "FStar.Tactics.CanonCommMonoid.fst" - (Prims.of_int (340)) - (Prims.of_int (6)) - (Prims.of_int (411)) - (Prims.of_int (69))) - (Obj.magic - (FStar_Tactics_Builtins.term_eq_old - t ta)) + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.CanonCommMonoid.fst" + (Prims.of_int (340)) + (Prims.of_int (9)) + (Prims.of_int (340)) + (Prims.of_int (25))) + (Prims.mk_range + "FStar.Tactics.CanonCommMonoid.fst" + (Prims.of_int (340)) + (Prims.of_int (6)) + (Prims.of_int (411)) + (Prims.of_int (69))) + (Obj.magic + (FStar_Tactics_Builtins.term_eq_old + t ta)) + (fun uu___2 -> (fun uu___2 -> - (fun uu___2 -> - if uu___2 - then - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind + if uu___2 + then + Obj.magic + (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.CanonCommMonoid.fst" (Prims.of_int (341)) @@ -1733,7 +1731,6 @@ let canon_monoid_aux : (r1::r2::[], vm) -> Obj.magic - (Obj.repr (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.CanonCommMonoid.fst" @@ -2000,25 +1997,22 @@ let canon_monoid_aux : uu___4))) uu___4))) uu___4))) - uu___4))) + uu___4)) | uu___4 -> Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "Unexpected"))) - uu___3))) - else - Obj.magic - (Obj.repr (FStar_Tactics_Derived.fail - "Goal should be an equality at the right monoid type"))) - uu___2))) + "Unexpected")) + uu___3)) + else + Obj.magic + (FStar_Tactics_Derived.fail + "Goal should be an equality at the right monoid type")) + uu___2)) | uu___2 -> Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "Goal should be an equality"))) + (FStar_Tactics_Derived.fail + "Goal should be an equality")) uu___1))) uu___) let canon_monoid_with : 'b . diff --git a/ocaml/fstar-lib/generated/FStar_Tactics_CanonCommMonoidSimple.ml b/ocaml/fstar-lib/generated/FStar_Tactics_CanonCommMonoidSimple.ml index 16012c10b64..d4730d05f43 100644 --- a/ocaml/fstar-lib/generated/FStar_Tactics_CanonCommMonoidSimple.ml +++ b/ocaml/fstar-lib/generated/FStar_Tactics_CanonCommMonoidSimple.ml @@ -624,97 +624,89 @@ let canon_monoid : (FStar_Pervasives_Native.Some t), t1, t2) -> Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.CanonCommMonoidSimple.fst" - (Prims.of_int (255)) (Prims.of_int (9)) - (Prims.of_int (255)) - (Prims.of_int (28))) - (Prims.mk_range - "FStar.Tactics.CanonCommMonoidSimple.fst" - (Prims.of_int (255)) (Prims.of_int (6)) - (Prims.of_int (273)) - (Prims.of_int (69))) - (Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.CanonCommMonoidSimple.fst" - (Prims.of_int (255)) - (Prims.of_int (19)) - (Prims.of_int (255)) - (Prims.of_int (28))) - (Prims.mk_range - "FStar.Tactics.CanonCommMonoidSimple.fst" - (Prims.of_int (255)) - (Prims.of_int (9)) - (Prims.of_int (255)) - (Prims.of_int (28))) - (FStar_Tactics_Effect.lift_div_tac - (fun uu___2 -> - (fun uu___2 -> - Obj.magic - (failwith - "Cannot evaluate open quotation at runtime")) - uu___2)) + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.CanonCommMonoidSimple.fst" + (Prims.of_int (255)) (Prims.of_int (9)) + (Prims.of_int (255)) (Prims.of_int (28))) + (Prims.mk_range + "FStar.Tactics.CanonCommMonoidSimple.fst" + (Prims.of_int (255)) (Prims.of_int (6)) + (Prims.of_int (273)) (Prims.of_int (69))) + (Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.CanonCommMonoidSimple.fst" + (Prims.of_int (255)) + (Prims.of_int (19)) + (Prims.of_int (255)) + (Prims.of_int (28))) + (Prims.mk_range + "FStar.Tactics.CanonCommMonoidSimple.fst" + (Prims.of_int (255)) + (Prims.of_int (9)) + (Prims.of_int (255)) + (Prims.of_int (28))) + (FStar_Tactics_Effect.lift_div_tac (fun uu___2 -> (fun uu___2 -> - Obj.magic (term_eq t uu___2)) - uu___2))) - (fun uu___2 -> + Obj.magic + (failwith + "Cannot evaluate open quotation at runtime")) + uu___2)) (fun uu___2 -> - if uu___2 - then - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.CanonCommMonoidSimple.fst" - (Prims.of_int (256)) - (Prims.of_int (27)) - (Prims.of_int (256)) - (Prims.of_int (67))) - (Prims.mk_range - "FStar.Tactics.CanonCommMonoidSimple.fst" - (Prims.of_int (256)) - (Prims.of_int (8)) - (Prims.of_int (271)) - (Prims.of_int (22))) - (Obj.magic - (reification m [] - (const - (FStar_Algebra_CommMonoid.__proj__CM__item__unit - m)) t1)) - (fun uu___3 -> - (fun uu___3 -> - match uu___3 with - | (r1, ts, am) -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.CanonCommMonoidSimple.fst" - (Prims.of_int (257)) - (Prims.of_int (26)) - (Prims.of_int (257)) - (Prims.of_int (48))) - (Prims.mk_range - "FStar.Tactics.CanonCommMonoidSimple.fst" - (Prims.of_int (257)) - (Prims.of_int (8)) - (Prims.of_int (271)) - (Prims.of_int (22))) - (Obj.magic - (reification - m ts am - t2)) - (fun uu___4 - -> - (fun - uu___4 -> - match uu___4 - with - | - (r2, + (fun uu___2 -> + Obj.magic (term_eq t uu___2)) + uu___2))) + (fun uu___2 -> + (fun uu___2 -> + if uu___2 + then + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.CanonCommMonoidSimple.fst" + (Prims.of_int (256)) + (Prims.of_int (27)) + (Prims.of_int (256)) + (Prims.of_int (67))) + (Prims.mk_range + "FStar.Tactics.CanonCommMonoidSimple.fst" + (Prims.of_int (256)) + (Prims.of_int (8)) + (Prims.of_int (271)) + (Prims.of_int (22))) + (Obj.magic + (reification m [] + (const + (FStar_Algebra_CommMonoid.__proj__CM__item__unit + m)) t1)) + (fun uu___3 -> + (fun uu___3 -> + match uu___3 with + | (r1, ts, am) -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.CanonCommMonoidSimple.fst" + (Prims.of_int (257)) + (Prims.of_int (26)) + (Prims.of_int (257)) + (Prims.of_int (48))) + (Prims.mk_range + "FStar.Tactics.CanonCommMonoidSimple.fst" + (Prims.of_int (257)) + (Prims.of_int (8)) + (Prims.of_int (271)) + (Prims.of_int (22))) + (Obj.magic + (reification m + ts am t2)) + (fun uu___4 -> + (fun uu___4 -> + match uu___4 + with + | (r2, uu___5, am1) -> Obj.magic @@ -908,17 +900,15 @@ let canon_monoid : uu___8))) uu___7))) uu___6))) - uu___4))) - uu___3))) - else - Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "Goal should be an equality at the right monoid type"))) - uu___2))) + uu___4))) + uu___3)) + else + Obj.magic + (FStar_Tactics_Derived.fail + "Goal should be an equality at the right monoid type")) + uu___2)) | uu___2 -> Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "Goal should be an equality"))) uu___1))) + (FStar_Tactics_Derived.fail + "Goal should be an equality")) uu___1))) uu___) \ No newline at end of file diff --git a/ocaml/fstar-lib/generated/FStar_Tactics_CanonCommMonoidSimple_Equiv.ml b/ocaml/fstar-lib/generated/FStar_Tactics_CanonCommMonoidSimple_Equiv.ml index 5d9bec7196b..6bd65b0452a 100644 --- a/ocaml/fstar-lib/generated/FStar_Tactics_CanonCommMonoidSimple_Equiv.ml +++ b/ocaml/fstar-lib/generated/FStar_Tactics_CanonCommMonoidSimple_Equiv.ml @@ -1075,74 +1075,67 @@ let (canon_monoid : (match rel_xy with | (rel_xy1, uu___2)::[] -> Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.CanonCommMonoidSimple.Equiv.fst" - (Prims.of_int (381)) - (Prims.of_int (21)) - (Prims.of_int (381)) - (Prims.of_int (43))) - (Prims.mk_range - "FStar.Tactics.CanonCommMonoidSimple.Equiv.fst" - (Prims.of_int (380)) - (Prims.of_int (21)) - (Prims.of_int (391)) - (Prims.of_int (6))) - (FStar_Tactics_Effect.lift_div_tac - (fun uu___3 -> - FStar_Reflection_Derived_Lemmas.collect_app_ref - rel_xy1)) + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.CanonCommMonoidSimple.Equiv.fst" + (Prims.of_int (381)) + (Prims.of_int (21)) + (Prims.of_int (381)) + (Prims.of_int (43))) + (Prims.mk_range + "FStar.Tactics.CanonCommMonoidSimple.Equiv.fst" + (Prims.of_int (380)) + (Prims.of_int (21)) + (Prims.of_int (391)) + (Prims.of_int (6))) + (FStar_Tactics_Effect.lift_div_tac (fun uu___3 -> - (fun uu___3 -> - match uu___3 with - | (rel, xy) -> - if - (FStar_List_Tot_Base.length - xy) - >= - (Prims.of_int (2)) - then - Obj.magic - (Obj.repr - (match - ((FStar_List_Tot_Base.index + FStar_Reflection_Derived_Lemmas.collect_app_ref + rel_xy1)) + (fun uu___3 -> + (fun uu___3 -> + match uu___3 with + | (rel, xy) -> + if + (FStar_List_Tot_Base.length + xy) + >= + (Prims.of_int (2)) + then + (match + ((FStar_List_Tot_Base.index xy ((FStar_List_Tot_Base.length xy) - (Prims.of_int (2)))), - (FStar_List_Tot_Base.index + (FStar_List_Tot_Base.index xy ((FStar_List_Tot_Base.length xy) - Prims.int_one))) - with - | - ((lhs, - FStar_Reflection_Data.Q_Explicit), - (rhs, - FStar_Reflection_Data.Q_Explicit)) - -> - Obj.repr + with + | ((lhs, + FStar_Reflection_Data.Q_Explicit), + (rhs, + FStar_Reflection_Data.Q_Explicit)) + -> + Obj.magic (canon_lhs_rhs eq m lhs rhs) - | - uu___4 -> - Obj.repr - (FStar_Tactics_Derived.fail - "Goal should have been an application of a binary relation to 2 explicit arguments"))) - else - Obj.magic - (Obj.repr + | uu___4 -> + Obj.magic (FStar_Tactics_Derived.fail - "Goal should have been an application of a binary relation to n implicit and 2 explicit arguments"))) - uu___3))) + "Goal should have been an application of a binary relation to 2 explicit arguments")) + else + Obj.magic + (FStar_Tactics_Derived.fail + "Goal should have been an application of a binary relation to n implicit and 2 explicit arguments")) + uu___3)) | uu___2 -> Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "Goal should be squash applied to a binary relation")))) + (FStar_Tactics_Derived.fail + "Goal should be squash applied to a binary relation"))) uu___1))) uu___1))) uu___) let _ = FStar_Tactics_Native.register_tactic diff --git a/ocaml/fstar-lib/generated/FStar_Tactics_CanonCommSemiring.ml b/ocaml/fstar-lib/generated/FStar_Tactics_CanonCommSemiring.ml index 547fce6978e..a7843079ede 100644 --- a/ocaml/fstar-lib/generated/FStar_Tactics_CanonCommSemiring.ml +++ b/ocaml/fstar-lib/generated/FStar_Tactics_CanonCommSemiring.ml @@ -2136,31 +2136,29 @@ let canon_semiring_aux : t), t1, t2) -> Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.CanonCommSemiring.fst" - (Prims.of_int (1634)) - (Prims.of_int (9)) - (Prims.of_int (1634)) - (Prims.of_int (21))) - (Prims.mk_range - "FStar.Tactics.CanonCommSemiring.fst" - (Prims.of_int (1634)) - (Prims.of_int (6)) - (Prims.of_int (1671)) - (Prims.of_int (73))) - (Obj.magic - (term_eq - t ta)) + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.CanonCommSemiring.fst" + (Prims.of_int (1634)) + (Prims.of_int (9)) + (Prims.of_int (1634)) + (Prims.of_int (21))) + (Prims.mk_range + "FStar.Tactics.CanonCommSemiring.fst" + (Prims.of_int (1634)) + (Prims.of_int (6)) + (Prims.of_int (1671)) + (Prims.of_int (73))) + (Obj.magic + (term_eq t + ta)) + (fun uu___3 + -> (fun - uu___3 -> - (fun uu___3 -> if uu___3 then Obj.magic - (Obj.repr (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.CanonCommSemiring.fst" @@ -2193,7 +2191,6 @@ let canon_semiring_aux : (e1::e2::[], vm) -> Obj.magic - (Obj.repr (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.CanonCommSemiring.fst" @@ -2432,25 +2429,22 @@ let canon_semiring_aux : uu___5))) uu___5))) uu___5))) - uu___5))) + uu___5)) | uu___5 -> Obj.magic - (Obj.repr (FStar_Tactics_Derived.fail - "Unexpected"))) - uu___4))) + "Unexpected")) + uu___4)) else Obj.magic - (Obj.repr (FStar_Tactics_Derived.fail - "Found equality, but terms do not have the expected type"))) - uu___3))) + "Found equality, but terms do not have the expected type")) + uu___3)) | uu___3 -> Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "Goal should be an equality"))) + (FStar_Tactics_Derived.fail + "Goal should be an equality")) uu___2))) uu___2))) uu___1)) let canon_semiring : 'a . 'a cr -> (unit, unit) FStar_Tactics_Effect.tac_repr diff --git a/ocaml/fstar-lib/generated/FStar_Tactics_CanonMonoid.ml b/ocaml/fstar-lib/generated/FStar_Tactics_CanonMonoid.ml index be918e0117b..648100aa857 100644 --- a/ocaml/fstar-lib/generated/FStar_Tactics_CanonMonoid.ml +++ b/ocaml/fstar-lib/generated/FStar_Tactics_CanonMonoid.ml @@ -448,89 +448,85 @@ let canon_monoid : me2) -> Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.CanonMonoid.fst" - (Prims.of_int (110)) - (Prims.of_int (9)) - (Prims.of_int (110)) - (Prims.of_int (32))) - (Prims.mk_range - "FStar.Tactics.CanonMonoid.fst" - (Prims.of_int (110)) - (Prims.of_int (6)) - (Prims.of_int (119)) - (Prims.of_int (69))) - (Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.CanonMonoid.fst" - (Prims.of_int (110)) - (Prims.of_int (23)) - (Prims.of_int (110)) - (Prims.of_int (32))) - (Prims.mk_range - "FStar.Tactics.CanonMonoid.fst" - (Prims.of_int (110)) - (Prims.of_int (9)) - (Prims.of_int (110)) - (Prims.of_int (32))) - (FStar_Tactics_Effect.lift_div_tac - (fun uu___2 -> - (fun uu___2 -> - Obj.magic - (failwith - "Cannot evaluate open quotation at runtime")) - uu___2)) + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.CanonMonoid.fst" + (Prims.of_int (110)) + (Prims.of_int (9)) + (Prims.of_int (110)) + (Prims.of_int (32))) + (Prims.mk_range + "FStar.Tactics.CanonMonoid.fst" + (Prims.of_int (110)) + (Prims.of_int (6)) + (Prims.of_int (119)) + (Prims.of_int (69))) + (Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.CanonMonoid.fst" + (Prims.of_int (110)) + (Prims.of_int (23)) + (Prims.of_int (110)) + (Prims.of_int (32))) + (Prims.mk_range + "FStar.Tactics.CanonMonoid.fst" + (Prims.of_int (110)) + (Prims.of_int (9)) + (Prims.of_int (110)) + (Prims.of_int (32))) + (FStar_Tactics_Effect.lift_div_tac (fun uu___2 -> (fun uu___2 -> Obj.magic - (FStar_Tactics_Builtins.term_eq_old - t uu___2)) - uu___2))) - (fun uu___2 -> + (failwith + "Cannot evaluate open quotation at runtime")) + uu___2)) (fun uu___2 -> - if uu___2 - then - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.CanonMonoid.fst" - (Prims.of_int (111)) - (Prims.of_int (17)) - (Prims.of_int (111)) - (Prims.of_int (34))) - (Prims.mk_range - "FStar.Tactics.CanonMonoid.fst" - (Prims.of_int (112)) - (Prims.of_int (8)) - (Prims.of_int (118)) - (Prims.of_int (51))) - (Obj.magic - (reification - m me1)) - (fun uu___3 -> - (fun r1 -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range + (fun uu___2 -> + Obj.magic + (FStar_Tactics_Builtins.term_eq_old + t uu___2)) uu___2))) + (fun uu___2 -> + (fun uu___2 -> + if uu___2 + then + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.CanonMonoid.fst" + (Prims.of_int (111)) + (Prims.of_int (17)) + (Prims.of_int (111)) + (Prims.of_int (34))) + (Prims.mk_range + "FStar.Tactics.CanonMonoid.fst" + (Prims.of_int (112)) + (Prims.of_int (8)) + (Prims.of_int (118)) + (Prims.of_int (51))) + (Obj.magic + (reification m me1)) + (fun uu___3 -> + (fun r1 -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "FStar.Tactics.CanonMonoid.fst" (Prims.of_int (112)) (Prims.of_int (17)) (Prims.of_int (112)) (Prims.of_int (34))) - (Prims.mk_range + (Prims.mk_range "FStar.Tactics.CanonMonoid.fst" (Prims.of_int (113)) (Prims.of_int (8)) (Prims.of_int (118)) (Prims.of_int (51))) - (Obj.magic + (Obj.magic (reification m me2)) - (fun + (fun uu___3 -> (fun r2 -> @@ -621,16 +617,14 @@ let canon_monoid : uu___4))) uu___3))) uu___3))) - uu___3))) - else - Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "Goal should be an equality at the right monoid type"))) - uu___2))) + uu___3)) + else + Obj.magic + (FStar_Tactics_Derived.fail + "Goal should be an equality at the right monoid type")) + uu___2)) | uu___2 -> Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "Goal should be an equality"))) + (FStar_Tactics_Derived.fail + "Goal should be an equality")) uu___1))) uu___1))) uu___) \ No newline at end of file diff --git a/ocaml/fstar-lib/generated/FStar_Tactics_Derived.ml b/ocaml/fstar-lib/generated/FStar_Tactics_Derived.ml index c8d79f468ed..4c8008fdce0 100644 --- a/ocaml/fstar-lib/generated/FStar_Tactics_Derived.ml +++ b/ocaml/fstar-lib/generated/FStar_Tactics_Derived.ml @@ -36,7 +36,7 @@ let (goals : (Prims.of_int (42)) (Prims.of_int (40)) (Prims.of_int (50))) (Prims.mk_range "FStar.Tactics.Derived.fst" (Prims.of_int (40)) (Prims.of_int (33)) (Prims.of_int (40)) (Prims.of_int (50))) - (FStar_Tactics_Effect.get ()) + (Obj.magic (FStar_Tactics_Effect.get ())) (fun uu___1 -> FStar_Tactics_Effect.lift_div_tac (fun uu___2 -> FStar_Tactics_Types.goals_of uu___1)) @@ -50,11 +50,11 @@ let (smt_goals : (Prims.of_int (50)) (Prims.of_int (41)) (Prims.of_int (58))) (Prims.mk_range "FStar.Tactics.Derived.fst" (Prims.of_int (41)) (Prims.of_int (37)) (Prims.of_int (41)) (Prims.of_int (58))) - (FStar_Tactics_Effect.get ()) + (Obj.magic (FStar_Tactics_Effect.get ())) (fun uu___1 -> FStar_Tactics_Effect.lift_div_tac (fun uu___2 -> FStar_Tactics_Types.smt_goals_of uu___1)) -let fail : 'a . Prims.string -> ('a, Obj.t) FStar_Tactics_Effect.tac_repr = +let fail : 'a . Prims.string -> ('a, unit) FStar_Tactics_Effect.tac_repr = fun m -> FStar_Tactics_Effect.raise (FStar_Tactics_Common.TacticFailure m) let fail_silently : 'a . Prims.string -> ('a, unit) FStar_Tactics_Effect.tac_repr = @@ -66,7 +66,10 @@ let fail_silently : (Prims.of_int (4)) (Prims.of_int (50)) (Prims.of_int (30))) (Obj.magic (FStar_Tactics_Builtins.set_urgency Prims.int_zero)) (fun uu___ -> - FStar_Tactics_Effect.raise (FStar_Tactics_Common.TacticFailure m)) + (fun uu___ -> + Obj.magic + (FStar_Tactics_Effect.raise + (FStar_Tactics_Common.TacticFailure m))) uu___) let (_cur_goal : unit -> (FStar_Tactics_Types.goal, unit) FStar_Tactics_Effect.tac_repr) = fun uu___ -> @@ -77,9 +80,14 @@ let (_cur_goal : (Prims.of_int (4)) (Prims.of_int (56)) (Prims.of_int (15))) (Obj.magic (goals ())) (fun uu___1 -> - match uu___1 with - | [] -> fail "no more goals" - | g::uu___2 -> FStar_Tactics_Effect.lift_div_tac (fun uu___3 -> g)) + (fun uu___1 -> + match uu___1 with + | [] -> Obj.magic (Obj.repr (fail "no more goals")) + | g::uu___2 -> + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac (fun uu___3 -> g)))) + uu___1) let (cur_env : unit -> (FStar_Reflection_Types.env, unit) FStar_Tactics_Effect.tac_repr) = fun uu___ -> @@ -131,7 +139,7 @@ let (cur_goal_safe : (Prims.of_int (18)) (Prims.of_int (72)) (Prims.of_int (26))) (Prims.mk_range "FStar.Tactics.Derived.fst" (Prims.of_int (72)) (Prims.of_int (9)) (Prims.of_int (72)) (Prims.of_int (26))) - (FStar_Tactics_Effect.get ()) + (Obj.magic (FStar_Tactics_Effect.get ())) (fun uu___1 -> FStar_Tactics_Effect.lift_div_tac (fun uu___2 -> FStar_Tactics_Types.goals_of uu___1)))) @@ -267,16 +275,14 @@ let (trivial : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = match uu___2 with | FStar_Reflection_Formula.True_ -> Obj.magic - (Obj.repr - (exact - (FStar_Reflection_Builtins.pack_ln - (FStar_Reflection_Data.Tv_Const - FStar_Reflection_Data.C_Unit)))) + (exact + (FStar_Reflection_Builtins.pack_ln + (FStar_Reflection_Data.Tv_Const + FStar_Reflection_Data.C_Unit))) | uu___3 -> Obj.magic - (Obj.repr - (FStar_Tactics_Effect.raise - Goal_not_trivial))) uu___2))) + (FStar_Tactics_Effect.raise + Goal_not_trivial)) uu___2))) uu___2))) uu___1) let (dismiss : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = fun uu___ -> @@ -289,9 +295,8 @@ let (dismiss : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = (fun uu___1 -> (fun uu___1 -> match uu___1 with - | [] -> Obj.magic (Obj.repr (fail "dismiss: no more goals")) - | uu___2::gs -> - Obj.magic (Obj.repr (FStar_Tactics_Builtins.set_goals gs))) + | [] -> Obj.magic (fail "dismiss: no more goals") + | uu___2::gs -> Obj.magic (FStar_Tactics_Builtins.set_goals gs)) uu___1) let (flip : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = fun uu___ -> @@ -315,17 +320,13 @@ let (flip : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = (fun uu___1 -> (fun uu___1 -> match uu___1 with - | [] -> - Obj.magic - (Obj.repr (fail "flip: less than two goals")) + | [] -> Obj.magic (fail "flip: less than two goals") | uu___2::[] -> - Obj.magic - (Obj.repr (fail "flip: less than two goals")) + Obj.magic (fail "flip: less than two goals") | g1::g2::gs1 -> Obj.magic - (Obj.repr - (FStar_Tactics_Builtins.set_goals (g2 :: g1 - :: gs1)))) uu___1))) uu___1) + (FStar_Tactics_Builtins.set_goals (g2 :: g1 :: + gs1))) uu___1))) uu___1) let (qed : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = fun uu___ -> FStar_Tactics_Effect.tac_bind @@ -335,9 +336,13 @@ let (qed : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = (Prims.of_int (4)) (Prims.of_int (130)) (Prims.of_int (32))) (Obj.magic (goals ())) (fun uu___1 -> - match uu___1 with - | [] -> FStar_Tactics_Effect.lift_div_tac (fun uu___2 -> ()) - | uu___2 -> fail "qed: not done!") + (fun uu___1 -> + match uu___1 with + | [] -> + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac (fun uu___2 -> ()))) + | uu___2 -> Obj.magic (Obj.repr (fail "qed: not done!"))) uu___1) let (debug : Prims.string -> (unit, unit) FStar_Tactics_Effect.tac_repr) = fun m -> FStar_Tactics_Effect.tac_bind @@ -386,24 +391,22 @@ let (smt : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = (fun uu___1 -> (fun uu___1 -> match uu___1 with - | ([], uu___2) -> - Obj.magic (Obj.repr (fail "smt: no active goals")) + | ([], uu___2) -> Obj.magic (fail "smt: no active goals") | (g::gs, gs') -> Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "FStar.Tactics.Derived.fst" - (Prims.of_int (147)) (Prims.of_int (8)) - (Prims.of_int (147)) (Prims.of_int (20))) - (Prims.mk_range "FStar.Tactics.Derived.fst" - (Prims.of_int (148)) (Prims.of_int (8)) - (Prims.of_int (148)) (Prims.of_int (32))) - (Obj.magic (FStar_Tactics_Builtins.set_goals gs)) + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "FStar.Tactics.Derived.fst" + (Prims.of_int (147)) (Prims.of_int (8)) + (Prims.of_int (147)) (Prims.of_int (20))) + (Prims.mk_range "FStar.Tactics.Derived.fst" + (Prims.of_int (148)) (Prims.of_int (8)) + (Prims.of_int (148)) (Prims.of_int (32))) + (Obj.magic (FStar_Tactics_Builtins.set_goals gs)) + (fun uu___2 -> (fun uu___2 -> - (fun uu___2 -> - Obj.magic - (FStar_Tactics_Builtins.set_smt_goals (g :: - gs'))) uu___2)))) uu___1) + Obj.magic + (FStar_Tactics_Builtins.set_smt_goals (g :: gs'))) + uu___2))) uu___1) let (idtac : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = fun uu___ -> (fun uu___ -> @@ -422,10 +425,9 @@ let (later : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = match uu___1 with | g::gs -> Obj.magic - (Obj.repr - (FStar_Tactics_Builtins.set_goals - (FStar_List_Tot_Base.op_At gs [g]))) - | uu___2 -> Obj.magic (Obj.repr (fail "later: no goals"))) uu___1) + (FStar_Tactics_Builtins.set_goals + (FStar_List_Tot_Base.op_At gs [g])) + | uu___2 -> Obj.magic (fail "later: no goals")) uu___1) let (apply : FStar_Reflection_Types.term -> (unit, unit) FStar_Tactics_Effect.tac_repr) = fun t -> FStar_Tactics_Builtins.t_apply true false false t @@ -534,21 +536,29 @@ let (topdown_rewrite : (Prims.of_int (254)) (Prims.of_int (10))) (match i with | uu___2 when uu___2 = Prims.int_zero -> - FStar_Tactics_Effect.lift_div_tac - (fun uu___3 -> - FStar_Tactics_Types.Continue) + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___3 -> + FStar_Tactics_Types.Continue))) | uu___2 when uu___2 = Prims.int_one -> - FStar_Tactics_Effect.lift_div_tac - (fun uu___3 -> - FStar_Tactics_Types.Skip) + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___3 -> + FStar_Tactics_Types.Skip))) | uu___2 when uu___2 = (Prims.of_int (2)) -> - FStar_Tactics_Effect.lift_div_tac - (fun uu___3 -> - FStar_Tactics_Types.Abort) + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___3 -> + FStar_Tactics_Types.Abort))) | uu___2 -> - fail - "topdown_rewrite: bad value from ctrl") + Obj.magic + (Obj.repr + (fail + "topdown_rewrite: bad value from ctrl"))) (fun f -> FStar_Tactics_Effect.lift_div_tac (fun uu___2 -> (b, f))))) uu___1))) @@ -671,8 +681,11 @@ let divide : (Prims.mk_range "FStar.Tactics.Derived.fst" (Prims.of_int (289)) (Prims.of_int (4)) (Prims.of_int (301)) (Prims.of_int (10))) (if n < Prims.int_zero - then fail "divide: negative n" - else FStar_Tactics_Effect.lift_div_tac (fun uu___1 -> ())) + then Obj.magic (Obj.repr (fail "divide: negative n")) + else + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac (fun uu___1 -> ())))) (fun uu___ -> (fun uu___ -> Obj.magic @@ -1093,91 +1106,84 @@ let focus : (fun uu___ -> (fun uu___ -> match uu___ with - | [] -> Obj.magic (Obj.repr (fail "focus: no goals")) + | [] -> Obj.magic (fail "focus: no goals") | g::gs -> Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "FStar.Tactics.Derived.fst" - (Prims.of_int (314)) (Prims.of_int (18)) - (Prims.of_int (314)) (Prims.of_int (30))) - (Prims.mk_range "FStar.Tactics.Derived.fst" - (Prims.of_int (315)) (Prims.of_int (8)) - (Prims.of_int (318)) (Prims.of_int (9))) - (Obj.magic (smt_goals ())) - (fun uu___1 -> - (fun sgs -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.Derived.fst" - (Prims.of_int (315)) (Prims.of_int (8)) - (Prims.of_int (315)) - (Prims.of_int (21))) - (Prims.mk_range - "FStar.Tactics.Derived.fst" - (Prims.of_int (315)) - (Prims.of_int (23)) - (Prims.of_int (318)) (Prims.of_int (9))) - (Obj.magic - (FStar_Tactics_Builtins.set_goals [g])) + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "FStar.Tactics.Derived.fst" + (Prims.of_int (314)) (Prims.of_int (18)) + (Prims.of_int (314)) (Prims.of_int (30))) + (Prims.mk_range "FStar.Tactics.Derived.fst" + (Prims.of_int (315)) (Prims.of_int (8)) + (Prims.of_int (318)) (Prims.of_int (9))) + (Obj.magic (smt_goals ())) + (fun uu___1 -> + (fun sgs -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "FStar.Tactics.Derived.fst" + (Prims.of_int (315)) (Prims.of_int (8)) + (Prims.of_int (315)) (Prims.of_int (21))) + (Prims.mk_range "FStar.Tactics.Derived.fst" + (Prims.of_int (315)) (Prims.of_int (23)) + (Prims.of_int (318)) (Prims.of_int (9))) + (Obj.magic + (FStar_Tactics_Builtins.set_goals [g])) + (fun uu___1 -> (fun uu___1 -> - (fun uu___1 -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.Derived.fst" - (Prims.of_int (315)) - (Prims.of_int (23)) - (Prims.of_int (315)) - (Prims.of_int (39))) - (Prims.mk_range - "FStar.Tactics.Derived.fst" - (Prims.of_int (316)) - (Prims.of_int (8)) - (Prims.of_int (318)) - (Prims.of_int (9))) - (Obj.magic - (FStar_Tactics_Builtins.set_smt_goals - [])) + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.Derived.fst" + (Prims.of_int (315)) + (Prims.of_int (23)) + (Prims.of_int (315)) + (Prims.of_int (39))) + (Prims.mk_range + "FStar.Tactics.Derived.fst" + (Prims.of_int (316)) + (Prims.of_int (8)) + (Prims.of_int (318)) + (Prims.of_int (9))) + (Obj.magic + (FStar_Tactics_Builtins.set_smt_goals + [])) + (fun uu___2 -> (fun uu___2 -> - (fun uu___2 -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.Derived.fst" - (Prims.of_int (316)) - (Prims.of_int (16)) - (Prims.of_int (316)) - (Prims.of_int (20))) - (Prims.mk_range - "FStar.Tactics.Derived.fst" - (Prims.of_int (317)) - (Prims.of_int (8)) - (Prims.of_int (318)) - (Prims.of_int (9))) - (Obj.magic (t ())) - (fun uu___3 -> - (fun x -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - ( - Prims.mk_range + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.Derived.fst" + (Prims.of_int (316)) + (Prims.of_int (16)) + (Prims.of_int (316)) + (Prims.of_int (20))) + (Prims.mk_range + "FStar.Tactics.Derived.fst" + (Prims.of_int (317)) + (Prims.of_int (8)) + (Prims.of_int (318)) + (Prims.of_int (9))) + (Obj.magic (t ())) + (fun uu___3 -> + (fun x -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "FStar.Tactics.Derived.fst" (Prims.of_int (317)) (Prims.of_int (8)) (Prims.of_int (317)) (Prims.of_int (33))) - ( - Prims.mk_range + (Prims.mk_range "FStar.Tactics.Derived.fst" (Prims.of_int (317)) (Prims.of_int (35)) (Prims.of_int (318)) (Prims.of_int (9))) + (Obj.magic ( - Obj.magic - (FStar_Tactics_Effect.tac_bind + FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.Derived.fst" (Prims.of_int (317)) @@ -1221,9 +1227,8 @@ let focus : (FStar_Tactics_Builtins.set_goals uu___3)) uu___3))) - ( - fun - uu___3 -> + (fun uu___3 + -> (fun uu___3 -> Obj.magic @@ -1294,9 +1299,8 @@ let focus : uu___5 -> x)))) uu___3))) - uu___3))) - uu___2))) uu___1))) uu___1)))) - uu___) + uu___3))) uu___2))) + uu___1))) uu___1))) uu___) let (dump1 : Prims.string -> (unit, unit) FStar_Tactics_Effect.tac_repr) = fun m -> focus (fun uu___ -> FStar_Tactics_Builtins.dump m) let rec mapAll : @@ -1782,8 +1786,10 @@ let (guard : Prims.bool -> (unit, unit) FStar_Tactics_Effect.tac_repr) = fun uu___ -> (fun b -> if Prims.op_Negation b - then Obj.magic (fail "guard failed") - else Obj.magic (FStar_Tactics_Effect.lift_div_tac (fun uu___1 -> ()))) + then Obj.magic (Obj.repr (fail "guard failed")) + else + Obj.magic + (Obj.repr (FStar_Tactics_Effect.lift_div_tac (fun uu___1 -> ())))) uu___ let try_with : 'a . @@ -1857,8 +1863,7 @@ let first : = fun ts -> FStar_List_Tot_Base.fold_right op_Less_Bar_Greater ts - (fun uu___ -> (fun uu___ -> Obj.magic (fail "no tactics to try")) uu___) - () + (fun uu___ -> fail "no tactics to try") () let rec repeat : 'a . (unit -> ('a, unit) FStar_Tactics_Effect.tac_repr) -> @@ -2136,9 +2141,8 @@ let (skip_guard : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = (Obj.magic (is_guard ())) (fun uu___1 -> (fun uu___1 -> - if uu___1 - then Obj.magic (Obj.repr (smt ())) - else Obj.magic (Obj.repr (fail ""))) uu___1) + if uu___1 then Obj.magic (smt ()) else Obj.magic (fail "")) + uu___1) let (guards_to_smt : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = fun uu___ -> FStar_Tactics_Effect.tac_bind @@ -2383,57 +2387,50 @@ let rec (__assumption_aux : FStar_Reflection_Types.binders -> (unit, unit) FStar_Tactics_Effect.tac_repr) = - fun uu___ -> - (fun bs -> - match bs with - | [] -> Obj.magic (Obj.repr (fail "no assumption matches goal")) - | b::bs1 -> - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "FStar.Tactics.Derived.fst" - (Prims.of_int (539)) (Prims.of_int (16)) - (Prims.of_int (539)) (Prims.of_int (32))) - (Prims.mk_range "FStar.Tactics.Derived.fst" - (Prims.of_int (540)) (Prims.of_int (8)) - (Prims.of_int (543)) (Prims.of_int (27))) - (Obj.magic (binder_to_term b)) - (fun uu___ -> - (fun t -> - Obj.magic - (try_with - (fun uu___ -> match () with | () -> exact t) - (fun uu___ -> - try_with - (fun uu___1 -> - match () with - | () -> - FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.Derived.fst" - (Prims.of_int (541)) - (Prims.of_int (13)) - (Prims.of_int (541)) - (Prims.of_int (48))) - (Prims.mk_range - "FStar.Tactics.Derived.fst" - (Prims.of_int (542)) - (Prims.of_int (13)) - (Prims.of_int (542)) - (Prims.of_int (20))) - (Obj.magic - (apply - (FStar_Reflection_Builtins.pack_ln - (FStar_Reflection_Data.Tv_FVar - (FStar_Reflection_Builtins.pack_fv - ["FStar"; - "Squash"; - "return_squash"]))))) - (fun uu___2 -> - (fun uu___2 -> - Obj.magic (exact t)) uu___2)) - (fun uu___1 -> __assumption_aux bs1)))) - uu___)))) uu___ + fun bs -> + match bs with + | [] -> fail "no assumption matches goal" + | b::bs1 -> + FStar_Tactics_Effect.tac_bind + (Prims.mk_range "FStar.Tactics.Derived.fst" (Prims.of_int (539)) + (Prims.of_int (16)) (Prims.of_int (539)) (Prims.of_int (32))) + (Prims.mk_range "FStar.Tactics.Derived.fst" (Prims.of_int (540)) + (Prims.of_int (8)) (Prims.of_int (543)) (Prims.of_int (27))) + (Obj.magic (binder_to_term b)) + (fun uu___ -> + (fun t -> + Obj.magic + (try_with (fun uu___ -> match () with | () -> exact t) + (fun uu___ -> + try_with + (fun uu___1 -> + match () with + | () -> + FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.Derived.fst" + (Prims.of_int (541)) + (Prims.of_int (13)) + (Prims.of_int (541)) + (Prims.of_int (48))) + (Prims.mk_range + "FStar.Tactics.Derived.fst" + (Prims.of_int (542)) + (Prims.of_int (13)) + (Prims.of_int (542)) + (Prims.of_int (20))) + (Obj.magic + (apply + (FStar_Reflection_Builtins.pack_ln + (FStar_Reflection_Data.Tv_FVar + (FStar_Reflection_Builtins.pack_fv + ["FStar"; + "Squash"; + "return_squash"]))))) + (fun uu___2 -> + (fun uu___2 -> Obj.magic (exact t)) + uu___2)) + (fun uu___1 -> __assumption_aux bs1)))) uu___) let (assumption : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = fun uu___ -> FStar_Tactics_Effect.tac_bind @@ -2525,8 +2522,7 @@ let (rewrite' : (fun uu___2 -> Obj.magic (FStar_Tactics_Builtins.rewrite b)) uu___2))) uu___1))) - (fun uu___ -> (fun uu___ -> Obj.magic (fail "rewrite' failed")) uu___) - () + (fun uu___ -> fail "rewrite' failed") () let rec (try_rewrite_equality : FStar_Reflection_Types.term -> FStar_Reflection_Types.binders -> @@ -2642,26 +2638,23 @@ let (unfold_def : match uu___ with | FStar_Reflection_Data.Tv_FVar fv -> Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "FStar.Tactics.Derived.fst" - (Prims.of_int (601)) (Prims.of_int (16)) - (Prims.of_int (601)) (Prims.of_int (42))) - (Prims.mk_range "FStar.Tactics.Derived.fst" - (Prims.of_int (602)) (Prims.of_int (8)) - (Prims.of_int (602)) (Prims.of_int (30))) - (FStar_Tactics_Effect.lift_div_tac - (fun uu___1 -> - FStar_Reflection_Builtins.implode_qn - (FStar_Reflection_Builtins.inspect_fv fv))) + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "FStar.Tactics.Derived.fst" + (Prims.of_int (601)) (Prims.of_int (16)) + (Prims.of_int (601)) (Prims.of_int (42))) + (Prims.mk_range "FStar.Tactics.Derived.fst" + (Prims.of_int (602)) (Prims.of_int (8)) + (Prims.of_int (602)) (Prims.of_int (30))) + (FStar_Tactics_Effect.lift_div_tac (fun uu___1 -> - (fun n -> - Obj.magic - (FStar_Tactics_Builtins.norm - [FStar_Pervasives.delta_fully [n]])) - uu___1))) - | uu___1 -> - Obj.magic (Obj.repr (fail "unfold_def: term is not a fv"))) + FStar_Reflection_Builtins.implode_qn + (FStar_Reflection_Builtins.inspect_fv fv))) + (fun uu___1 -> + (fun n -> + Obj.magic + (FStar_Tactics_Builtins.norm + [FStar_Pervasives.delta_fully [n]])) uu___1)) + | uu___1 -> Obj.magic (fail "unfold_def: term is not a fv")) uu___) let (l_to_r : FStar_Reflection_Types.term Prims.list -> @@ -2892,83 +2885,79 @@ let (grewrite_eq : | FStar_Reflection_Formula.Comp (FStar_Reflection_Formula.Eq uu___3, l, r) -> Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.Derived.fst" - (Prims.of_int (655)) - (Prims.of_int (6)) - (Prims.of_int (655)) - (Prims.of_int (18))) - (Prims.mk_range - "FStar.Tactics.Derived.fst" - (Prims.of_int (656)) - (Prims.of_int (6)) - (Prims.of_int (657)) - (Prims.of_int (56))) - (Obj.magic (grewrite l r)) + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.Derived.fst" + (Prims.of_int (655)) + (Prims.of_int (6)) + (Prims.of_int (655)) + (Prims.of_int (18))) + (Prims.mk_range + "FStar.Tactics.Derived.fst" + (Prims.of_int (656)) + (Prims.of_int (6)) + (Prims.of_int (657)) + (Prims.of_int (56))) + (Obj.magic (grewrite l r)) + (fun uu___4 -> (fun uu___4 -> - (fun uu___4 -> - Obj.magic - (iseq - [idtac; - (fun uu___5 -> - FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.Derived.fst" - (Prims.of_int (656)) - (Prims.of_int (30)) - (Prims.of_int (656)) - (Prims.of_int (55))) - (Prims.mk_range - "FStar.Tactics.Derived.fst" - (Prims.of_int (657)) - (Prims.of_int (30)) - (Prims.of_int (657)) - (Prims.of_int (54))) - (Obj.magic - (apply_lemma - (FStar_Reflection_Builtins.pack_ln - (FStar_Reflection_Data.Tv_FVar - (FStar_Reflection_Builtins.pack_fv - ["FStar"; - "Tactics"; - "Derived"; - "__un_sq_eq"]))))) + Obj.magic + (iseq + [idtac; + (fun uu___5 -> + FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.Derived.fst" + (Prims.of_int (656)) + (Prims.of_int (30)) + (Prims.of_int (656)) + (Prims.of_int (55))) + (Prims.mk_range + "FStar.Tactics.Derived.fst" + (Prims.of_int (657)) + (Prims.of_int (30)) + (Prims.of_int (657)) + (Prims.of_int (54))) + (Obj.magic + (apply_lemma + (FStar_Reflection_Builtins.pack_ln + (FStar_Reflection_Data.Tv_FVar + (FStar_Reflection_Builtins.pack_fv + ["FStar"; + "Tactics"; + "Derived"; + "__un_sq_eq"]))))) + (fun uu___6 -> (fun uu___6 -> - (fun uu___6 -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.Derived.fst" - (Prims.of_int (657)) - (Prims.of_int (36)) - (Prims.of_int (657)) - (Prims.of_int (54))) - (Prims.mk_range - "FStar.Tactics.Derived.fst" - (Prims.of_int (657)) - (Prims.of_int (30)) - (Prims.of_int (657)) - (Prims.of_int (54))) - (Obj.magic - (binder_to_term + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.Derived.fst" + (Prims.of_int (657)) + (Prims.of_int (36)) + (Prims.of_int (657)) + (Prims.of_int (54))) + (Prims.mk_range + "FStar.Tactics.Derived.fst" + (Prims.of_int (657)) + (Prims.of_int (30)) + (Prims.of_int (657)) + (Prims.of_int (54))) + (Obj.magic + (binder_to_term b)) + (fun uu___7 -> (fun uu___7 -> - (fun - uu___7 -> Obj.magic (exact uu___7)) uu___7))) - uu___6))])) - uu___4))) + uu___6))])) uu___4)) | uu___3 -> Obj.magic - (Obj.repr - (fail - "grewrite_eq: binder type is not an equality"))) + (fail + "grewrite_eq: binder type is not an equality")) uu___2))) uu___) let rec (apply_squash_or_lem : Prims.nat -> @@ -3000,121 +2989,101 @@ let rec (apply_squash_or_lem : (fun uu___1 -> try_with (fun uu___2 -> match () with | () -> apply_lemma t) (fun uu___2 -> - (fun uu___2 -> - if d <= Prims.int_zero - then - Obj.magic (Obj.repr (fail "mapply: out of fuel")) - else - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "FStar.Tactics.Derived.fst" - (Prims.of_int (688)) (Prims.of_int (13)) - (Prims.of_int (688)) (Prims.of_int (30))) - (Prims.mk_range "FStar.Tactics.Derived.fst" - (Prims.of_int (689)) (Prims.of_int (4)) - (Prims.of_int (737)) (Prims.of_int (41))) - (Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.Derived.fst" - (Prims.of_int (688)) - (Prims.of_int (16)) - (Prims.of_int (688)) - (Prims.of_int (28))) - (Prims.mk_range - "FStar.Tactics.Derived.fst" - (Prims.of_int (688)) - (Prims.of_int (13)) - (Prims.of_int (688)) - (Prims.of_int (30))) - (Obj.magic (cur_env ())) - (fun uu___4 -> - (fun uu___4 -> - Obj.magic - (FStar_Tactics_Builtins.tc - uu___4 t)) uu___4))) + if d <= Prims.int_zero + then fail "mapply: out of fuel" + else + FStar_Tactics_Effect.tac_bind + (Prims.mk_range "FStar.Tactics.Derived.fst" + (Prims.of_int (688)) (Prims.of_int (13)) + (Prims.of_int (688)) (Prims.of_int (30))) + (Prims.mk_range "FStar.Tactics.Derived.fst" + (Prims.of_int (689)) (Prims.of_int (4)) + (Prims.of_int (737)) (Prims.of_int (41))) + (Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "FStar.Tactics.Derived.fst" + (Prims.of_int (688)) (Prims.of_int (16)) + (Prims.of_int (688)) (Prims.of_int (28))) + (Prims.mk_range "FStar.Tactics.Derived.fst" + (Prims.of_int (688)) (Prims.of_int (13)) + (Prims.of_int (688)) (Prims.of_int (30))) + (Obj.magic (cur_env ())) + (fun uu___4 -> (fun uu___4 -> - (fun ty -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.Derived.fst" - (Prims.of_int (689)) - (Prims.of_int (17)) - (Prims.of_int (689)) - (Prims.of_int (31))) - (Prims.mk_range - "FStar.Tactics.Derived.fst" - (Prims.of_int (689)) - (Prims.of_int (4)) - (Prims.of_int (737)) - (Prims.of_int (41))) - (Obj.magic - (FStar_Tactics_SyntaxHelpers.collect_arr - ty)) - (fun uu___4 -> - (fun uu___4 -> - match uu___4 with - | (tys, c) -> - (match FStar_Reflection_Builtins.inspect_comp - c - with - | FStar_Reflection_Data.C_Lemma - (pre, post, - uu___5) - -> - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - ( - Prims.mk_range - "FStar.Tactics.Derived.fst" - (Prims.of_int (693)) - (Prims.of_int (18)) - (Prims.of_int (693)) - (Prims.of_int (32))) - ( - Prims.mk_range - "FStar.Tactics.Derived.fst" - (Prims.of_int (694)) - (Prims.of_int (7)) - (Prims.of_int (702)) - (Prims.of_int (41))) - ( - FStar_Tactics_Effect.lift_div_tac - (fun - uu___6 -> - FStar_Reflection_Builtins.pack_ln - (FStar_Reflection_Data.Tv_App - (post, + Obj.magic + (FStar_Tactics_Builtins.tc uu___4 t)) + uu___4))) + (fun uu___4 -> + (fun ty -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.Derived.fst" + (Prims.of_int (689)) + (Prims.of_int (17)) + (Prims.of_int (689)) + (Prims.of_int (31))) + (Prims.mk_range + "FStar.Tactics.Derived.fst" + (Prims.of_int (689)) + (Prims.of_int (4)) + (Prims.of_int (737)) + (Prims.of_int (41))) + (Obj.magic + (FStar_Tactics_SyntaxHelpers.collect_arr + ty)) + (fun uu___4 -> + (fun uu___4 -> + match uu___4 with + | (tys, c) -> + (match FStar_Reflection_Builtins.inspect_comp + c + with + | FStar_Reflection_Data.C_Lemma + (pre, post, uu___5) -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.Derived.fst" + (Prims.of_int (693)) + (Prims.of_int (18)) + (Prims.of_int (693)) + (Prims.of_int (32))) + (Prims.mk_range + "FStar.Tactics.Derived.fst" + (Prims.of_int (694)) + (Prims.of_int (7)) + (Prims.of_int (702)) + (Prims.of_int (41))) + (FStar_Tactics_Effect.lift_div_tac + (fun uu___6 -> + FStar_Reflection_Builtins.pack_ln + (FStar_Reflection_Data.Tv_App + (post, ((FStar_Reflection_Builtins.pack_ln (FStar_Reflection_Data.Tv_Const FStar_Reflection_Data.C_Unit)), FStar_Reflection_Data.Q_Explicit))))) - ( - fun - uu___6 -> - (fun - post1 -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range + (fun uu___6 -> + (fun post1 -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "FStar.Tactics.Derived.fst" (Prims.of_int (694)) (Prims.of_int (18)) (Prims.of_int (694)) (Prims.of_int (35))) - (Prims.mk_range + (Prims.mk_range "FStar.Tactics.Derived.fst" (Prims.of_int (696)) (Prims.of_int (7)) (Prims.of_int (702)) (Prims.of_int (41))) - (Obj.magic + (Obj.magic (norm_term [] post1)) - (fun + (fun uu___6 -> (fun post2 -> @@ -3145,7 +3114,6 @@ let rec (apply_squash_or_lem : FStar_Reflection_Formula.Implies (p, q) -> Obj.magic - (Obj.repr (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.Derived.fst" @@ -3177,47 +3145,42 @@ let rec (apply_squash_or_lem : (d - Prims.int_one) t)) - uu___7))) + uu___7)) | uu___7 -> Obj.magic - (Obj.repr (fail - "mapply: can't apply (1)"))) - uu___6))) + "mapply: can't apply (1)")) uu___6))) uu___6))) - | FStar_Reflection_Data.C_Total - rt -> - Obj.magic - (Obj.repr - (match - FStar_Reflection_Derived.unsquash_term - rt - with - | FStar_Pervasives_Native.Some - rt1 -> - FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.Derived.fst" - (Prims.of_int (710)) - (Prims.of_int (18)) - (Prims.of_int (710)) - (Prims.of_int (33))) - (Prims.mk_range - "FStar.Tactics.Derived.fst" - (Prims.of_int (712)) - (Prims.of_int (9)) - (Prims.of_int (718)) - (Prims.of_int (43))) - (Obj.magic - (norm_term - [] rt1)) - (fun - uu___5 -> - (fun rt2 - -> - Obj.magic + uu___6)) + | FStar_Reflection_Data.C_Total + rt -> + (match FStar_Reflection_Derived.unsquash_term + rt + with + | FStar_Pervasives_Native.Some + rt1 -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.Derived.fst" + (Prims.of_int (710)) + (Prims.of_int (18)) + (Prims.of_int (710)) + (Prims.of_int (33))) + (Prims.mk_range + "FStar.Tactics.Derived.fst" + (Prims.of_int (712)) + (Prims.of_int (9)) + (Prims.of_int (718)) + (Prims.of_int (43))) + (Obj.magic + (norm_term [] + rt1)) + (fun uu___5 -> + (fun rt2 -> + Obj.magic (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.Derived.fst" @@ -3244,7 +3207,6 @@ let rec (apply_squash_or_lem : FStar_Reflection_Formula.Implies (p, q) -> Obj.magic - (Obj.repr (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.Derived.fst" @@ -3276,38 +3238,36 @@ let rec (apply_squash_or_lem : (d - Prims.int_one) t)) - uu___6))) + uu___6)) | uu___6 -> Obj.magic - (Obj.repr (fail - "mapply: can't apply (1)"))) + "mapply: can't apply (1)")) uu___5))) - uu___5) - | FStar_Pervasives_Native.None - -> - FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.Derived.fst" - (Prims.of_int (725)) - (Prims.of_int (18)) - (Prims.of_int (725)) - (Prims.of_int (33))) - (Prims.mk_range - "FStar.Tactics.Derived.fst" - (Prims.of_int (727)) - (Prims.of_int (9)) - (Prims.of_int (734)) - (Prims.of_int (20))) - (Obj.magic - (norm_term - [] rt)) - (fun - uu___5 -> - (fun rt1 - -> - Obj.magic + uu___5)) + | FStar_Pervasives_Native.None + -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.Derived.fst" + (Prims.of_int (725)) + (Prims.of_int (18)) + (Prims.of_int (725)) + (Prims.of_int (33))) + (Prims.mk_range + "FStar.Tactics.Derived.fst" + (Prims.of_int (727)) + (Prims.of_int (9)) + (Prims.of_int (734)) + (Prims.of_int (20))) + (Obj.magic + (norm_term [] + rt)) + (fun uu___5 -> + (fun rt1 -> + Obj.magic (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.Derived.fst" @@ -3398,13 +3358,12 @@ let rec (apply_squash_or_lem : (apply t)) uu___7))) uu___5))) - uu___5))) - | uu___5 -> - Obj.magic - (Obj.repr - (fail - "mapply: can't apply (2)")))) - uu___4))) uu___4)))) uu___2))) + uu___5))) + | uu___5 -> + Obj.magic + (fail + "mapply: can't apply (2)"))) + uu___4))) uu___4)))) let (mapply : FStar_Reflection_Types.term -> (unit, unit) FStar_Tactics_Effect.tac_repr) = fun t -> apply_squash_or_lem (Prims.of_int (10)) t @@ -3518,10 +3477,7 @@ let finish_by : (Prims.of_int (773)) (Prims.of_int (9))) (Obj.magic (or_else qed - (fun uu___ -> - (fun uu___ -> - Obj.magic (fail "finish_by: not finished")) - uu___))) + (fun uu___ -> fail "finish_by: not finished"))) (fun uu___ -> FStar_Tactics_Effect.lift_div_tac (fun uu___1 -> x)))) uu___) @@ -3676,12 +3632,11 @@ let (tlabel : Prims.string -> (unit, unit) FStar_Tactics_Effect.tac_repr) = (fun uu___ -> (fun uu___ -> match uu___ with - | [] -> Obj.magic (Obj.repr (fail "tlabel: no goals")) + | [] -> Obj.magic (fail "tlabel: no goals") | h::t -> Obj.magic - (Obj.repr - (FStar_Tactics_Builtins.set_goals - ((FStar_Tactics_Types.set_label l h) :: t)))) uu___) + (FStar_Tactics_Builtins.set_goals + ((FStar_Tactics_Types.set_label l h) :: t))) uu___) let (tlabel' : Prims.string -> (unit, unit) FStar_Tactics_Effect.tac_repr) = fun l -> FStar_Tactics_Effect.tac_bind @@ -3693,27 +3648,26 @@ let (tlabel' : Prims.string -> (unit, unit) FStar_Tactics_Effect.tac_repr) = (fun uu___ -> (fun uu___ -> match uu___ with - | [] -> Obj.magic (Obj.repr (fail "tlabel': no goals")) + | [] -> Obj.magic (fail "tlabel': no goals") | h::t -> Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "FStar.Tactics.Derived.fst" - (Prims.of_int (821)) (Prims.of_int (16)) - (Prims.of_int (821)) (Prims.of_int (45))) - (Prims.mk_range "FStar.Tactics.Derived.fst" - (Prims.of_int (822)) (Prims.of_int (8)) - (Prims.of_int (822)) (Prims.of_int (26))) - (FStar_Tactics_Effect.lift_div_tac - (fun uu___1 -> - FStar_Tactics_Types.set_label - (Prims.strcat l - (FStar_Tactics_Types.get_label h)) h)) + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "FStar.Tactics.Derived.fst" + (Prims.of_int (821)) (Prims.of_int (16)) + (Prims.of_int (821)) (Prims.of_int (45))) + (Prims.mk_range "FStar.Tactics.Derived.fst" + (Prims.of_int (822)) (Prims.of_int (8)) + (Prims.of_int (822)) (Prims.of_int (26))) + (FStar_Tactics_Effect.lift_div_tac (fun uu___1 -> - (fun h1 -> - Obj.magic - (FStar_Tactics_Builtins.set_goals (h1 :: t))) - uu___1)))) uu___) + FStar_Tactics_Types.set_label + (Prims.strcat l + (FStar_Tactics_Types.get_label h)) h)) + (fun uu___1 -> + (fun h1 -> + Obj.magic + (FStar_Tactics_Builtins.set_goals (h1 :: t))) + uu___1))) uu___) let (focus_all : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = fun uu___ -> FStar_Tactics_Effect.tac_bind @@ -3795,11 +3749,9 @@ let (bump_nth : Prims.pos -> (unit, unit) FStar_Tactics_Effect.tac_repr) = (fun uu___ -> match uu___ with | FStar_Pervasives_Native.None -> - Obj.magic (Obj.repr (fail "bump_nth: not that many goals")) + Obj.magic (fail "bump_nth: not that many goals") | FStar_Pervasives_Native.Some (h, t) -> - Obj.magic - (Obj.repr (FStar_Tactics_Builtins.set_goals (h :: t)))) - uu___) + Obj.magic (FStar_Tactics_Builtins.set_goals (h :: t))) uu___) let rec (destruct_list : FStar_Reflection_Types.term -> (FStar_Reflection_Types.term Prims.list, unit) @@ -3829,22 +3781,20 @@ let rec (destruct_list : (FStar_Reflection_Builtins.inspect_fv fv) = FStar_Reflection_Const.cons_qn then - Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "FStar.Tactics.Derived.fst" - (Prims.of_int (851)) (Prims.of_int (17)) - (Prims.of_int (851)) (Prims.of_int (33))) - (Prims.mk_range "FStar.Tactics.Derived.fst" - (Prims.of_int (851)) (Prims.of_int (14)) - (Prims.of_int (851)) (Prims.of_int (16))) - (Obj.magic (destruct_list a2)) - (fun uu___1 -> - FStar_Tactics_Effect.lift_div_tac - (fun uu___2 -> a1 :: uu___1))) + FStar_Tactics_Effect.tac_bind + (Prims.mk_range "FStar.Tactics.Derived.fst" + (Prims.of_int (851)) (Prims.of_int (17)) + (Prims.of_int (851)) (Prims.of_int (33))) + (Prims.mk_range "FStar.Tactics.Derived.fst" + (Prims.of_int (851)) (Prims.of_int (14)) + (Prims.of_int (851)) (Prims.of_int (16))) + (Obj.magic (destruct_list a2)) + (fun uu___1 -> + FStar_Tactics_Effect.lift_div_tac + (fun uu___2 -> a1 :: uu___1)) else - Obj.repr - (FStar_Tactics_Effect.raise - FStar_Tactics_Common.NotAListLiteral))) + FStar_Tactics_Effect.raise + FStar_Tactics_Common.NotAListLiteral)) | (FStar_Reflection_Data.Tv_FVar fv, (uu___1, FStar_Reflection_Data.Q_Implicit)::(a1, FStar_Reflection_Data.Q_Explicit):: @@ -3855,22 +3805,20 @@ let rec (destruct_list : (FStar_Reflection_Builtins.inspect_fv fv) = FStar_Reflection_Const.cons_qn then - Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "FStar.Tactics.Derived.fst" - (Prims.of_int (851)) (Prims.of_int (17)) - (Prims.of_int (851)) (Prims.of_int (33))) - (Prims.mk_range "FStar.Tactics.Derived.fst" - (Prims.of_int (851)) (Prims.of_int (14)) - (Prims.of_int (851)) (Prims.of_int (16))) - (Obj.magic (destruct_list a2)) - (fun uu___2 -> - FStar_Tactics_Effect.lift_div_tac - (fun uu___3 -> a1 :: uu___2))) + FStar_Tactics_Effect.tac_bind + (Prims.mk_range "FStar.Tactics.Derived.fst" + (Prims.of_int (851)) (Prims.of_int (17)) + (Prims.of_int (851)) (Prims.of_int (33))) + (Prims.mk_range "FStar.Tactics.Derived.fst" + (Prims.of_int (851)) (Prims.of_int (14)) + (Prims.of_int (851)) (Prims.of_int (16))) + (Obj.magic (destruct_list a2)) + (fun uu___2 -> + FStar_Tactics_Effect.lift_div_tac + (fun uu___3 -> a1 :: uu___2)) else - Obj.repr - (FStar_Tactics_Effect.raise - FStar_Tactics_Common.NotAListLiteral))) + FStar_Tactics_Effect.raise + FStar_Tactics_Common.NotAListLiteral)) | (FStar_Reflection_Data.Tv_FVar fv, uu___1) -> Obj.magic (Obj.repr @@ -3878,11 +3826,13 @@ let rec (destruct_list : (FStar_Reflection_Builtins.inspect_fv fv) = FStar_Reflection_Const.nil_qn then - FStar_Tactics_Effect.lift_div_tac - (fun uu___2 -> []) + Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___2 -> [])) else - FStar_Tactics_Effect.raise - FStar_Tactics_Common.NotAListLiteral)) + Obj.repr + (FStar_Tactics_Effect.raise + FStar_Tactics_Common.NotAListLiteral))) | uu___1 -> Obj.magic (Obj.repr @@ -3910,25 +3860,30 @@ let (get_match_body : (fun uu___1 -> (fun uu___1 -> match uu___1 with - | FStar_Pervasives_Native.None -> Obj.magic (Obj.repr (fail "")) + | FStar_Pervasives_Native.None -> Obj.magic (fail "") | FStar_Pervasives_Native.Some t -> Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "FStar.Tactics.Derived.fst" - (Prims.of_int (863)) (Prims.of_int (20)) - (Prims.of_int (863)) (Prims.of_int (39))) - (Prims.mk_range "FStar.Tactics.Derived.fst" - (Prims.of_int (863)) (Prims.of_int (14)) - (Prims.of_int (865)) (Prims.of_int (46))) - (Obj.magic (inspect_unascribe t)) + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "FStar.Tactics.Derived.fst" + (Prims.of_int (863)) (Prims.of_int (20)) + (Prims.of_int (863)) (Prims.of_int (39))) + (Prims.mk_range "FStar.Tactics.Derived.fst" + (Prims.of_int (863)) (Prims.of_int (14)) + (Prims.of_int (865)) (Prims.of_int (46))) + (Obj.magic (inspect_unascribe t)) + (fun uu___2 -> (fun uu___2 -> match uu___2 with | FStar_Reflection_Data.Tv_Match (sc, uu___3, uu___4) -> - FStar_Tactics_Effect.lift_div_tac - (fun uu___5 -> sc) - | uu___3 -> fail "Goal is not a match")))) uu___1) + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___5 -> sc))) + | uu___3 -> + Obj.magic + (Obj.repr (fail "Goal is not a match"))) + uu___2))) uu___1) let rec last : 'a . 'a Prims.list -> ('a, unit) FStar_Tactics_Effect.tac_repr = fun uu___ -> @@ -4061,17 +4016,26 @@ let (nth_binder : (Prims.of_int (896)) (Prims.of_int (2)) (Prims.of_int (898)) (Prims.of_int (15))) (if k < Prims.int_zero - then fail "not enough binders" + then + Obj.magic + (Obj.repr (fail "not enough binders")) else - FStar_Tactics_Effect.lift_div_tac - (fun uu___1 -> k)) - (fun k1 -> - match FStar_List_Tot_Base.nth bs k1 with - | FStar_Pervasives_Native.None -> - fail "not enough binders" - | FStar_Pervasives_Native.Some b -> - FStar_Tactics_Effect.lift_div_tac - (fun uu___ -> b)))) uu___))) uu___) + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___1 -> k)))) + (fun uu___ -> + (fun k1 -> + match FStar_List_Tot_Base.nth bs k1 with + | FStar_Pervasives_Native.None -> + Obj.magic + (Obj.repr (fail "not enough binders")) + | FStar_Pervasives_Native.Some b -> + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___ -> b)))) uu___))) + uu___))) uu___) exception Appears let (uu___is_Appears : Prims.exn -> Prims.bool) = fun projectee -> match projectee with | Appears -> true | uu___ -> false @@ -4113,10 +4077,15 @@ let (name_appears_in : (if (FStar_Reflection_Builtins.inspect_fv fv) = nm - then FStar_Tactics_Effect.raise Appears + then + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.raise Appears)) else - FStar_Tactics_Effect.lift_div_tac - (fun uu___3 -> ())) + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___3 -> ())))) (fun uu___2 -> FStar_Tactics_Effect.lift_div_tac (fun uu___3 -> t1))) @@ -4157,9 +4126,12 @@ let (name_appears_in : match uu___ with | Appears -> Obj.magic - (FStar_Tactics_Effect.lift_div_tac - (fun uu___1 -> true)) - | e -> Obj.magic (FStar_Tactics_Effect.raise e)) + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___1 -> true))) + | e -> + Obj.magic + (Obj.repr (FStar_Tactics_Effect.raise e))) uu___))) uu___) let rec (mk_abs : FStar_Reflection_Types.binder Prims.list -> diff --git a/ocaml/fstar-lib/generated/FStar_Tactics_Effect.ml b/ocaml/fstar-lib/generated/FStar_Tactics_Effect.ml index c20d9427b13..226dda9374b 100644 --- a/ocaml/fstar-lib/generated/FStar_Tactics_Effect.ml +++ b/ocaml/fstar-lib/generated/FStar_Tactics_Effect.ml @@ -49,9 +49,9 @@ let __proj__TAC__item__bind = tac_bind type ('a, 'wp, 'uuuuu, 'uuuuu1) lift_div_tac_wp = 'wp let lift_div_tac : 'a 'wp . (unit -> 'a) -> ('a, 'wp) tac_repr = fun f -> fun ps -> FStar_Tactics_Result.Success ((f ()), ps) -let (get : unit -> (FStar_Tactics_Types.proofstate, Obj.t) tac_repr) = +let (get : unit -> (FStar_Tactics_Types.proofstate, unit) tac_repr) = fun uu___ -> fun ps -> FStar_Tactics_Result.Success (ps, ps) -let raise : 'a . Prims.exn -> ('a, Obj.t) tac_repr = +let raise : 'a . Prims.exn -> ('a, unit) tac_repr = fun e -> fun ps -> FStar_Tactics_Result.Failed (e, ps) type ('uuuuu, 'p) with_tactic = 'p let (rewrite_with_tactic : diff --git a/ocaml/fstar-lib/generated/FStar_Tactics_Logic.ml b/ocaml/fstar-lib/generated/FStar_Tactics_Logic.ml index 453aaf00628..9848dd07536 100644 --- a/ocaml/fstar-lib/generated/FStar_Tactics_Logic.ml +++ b/ocaml/fstar-lib/generated/FStar_Tactics_Logic.ml @@ -105,10 +105,7 @@ let (split : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = (FStar_Reflection_Data.Tv_FVar (FStar_Reflection_Builtins.pack_fv ["FStar"; "Tactics"; "Logic"; "split_lem"])))) - (fun uu___1 -> - (fun uu___1 -> - Obj.magic (FStar_Tactics_Derived.fail "Could not split goal")) - uu___1) + (fun uu___1 -> FStar_Tactics_Derived.fail "Could not split goal") let (implies_intro : unit -> (FStar_Reflection_Types.binder, unit) FStar_Tactics_Effect.tac_repr) = @@ -222,9 +219,12 @@ let (pose_lemma : (Prims.of_int (126)) (Prims.of_int (5))) (match FStar_Reflection_Builtins.inspect_comp c with | FStar_Reflection_Data.C_Lemma (pre, post, uu___) -> - FStar_Tactics_Effect.lift_div_tac - (fun uu___1 -> (pre, post)) - | uu___ -> FStar_Tactics_Derived.fail "") + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___1 -> (pre, post)))) + | uu___ -> + Obj.magic (Obj.repr (FStar_Tactics_Derived.fail ""))) (fun uu___ -> (fun uu___ -> match uu___ with @@ -729,80 +729,74 @@ let rec (simplify_eq_implication : match r with | FStar_Pervasives_Native.None -> Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "Not an equality implication")) + (FStar_Tactics_Derived.fail + "Not an equality implication") | FStar_Pervasives_Native.Some (uu___1, rhs) -> Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.Logic.fst" - (Prims.of_int (160)) - (Prims.of_int (19)) - (Prims.of_int (160)) - (Prims.of_int (35))) - (Prims.mk_range - "FStar.Tactics.Logic.fst" - (Prims.of_int (161)) - (Prims.of_int (8)) - (Prims.of_int (163)) - (Prims.of_int (37))) - (Obj.magic (implies_intro ())) - (fun uu___2 -> - (fun eq_h -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.Logic.fst" - (Prims.of_int (161)) - (Prims.of_int (8)) - (Prims.of_int (161)) - (Prims.of_int (20))) - (Prims.mk_range - "FStar.Tactics.Logic.fst" - (Prims.of_int (162)) - (Prims.of_int (8)) - (Prims.of_int (163)) - (Prims.of_int (37))) - (Obj.magic - (FStar_Tactics_Builtins.rewrite - eq_h)) + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.Logic.fst" + (Prims.of_int (160)) + (Prims.of_int (19)) + (Prims.of_int (160)) + (Prims.of_int (35))) + (Prims.mk_range + "FStar.Tactics.Logic.fst" + (Prims.of_int (161)) + (Prims.of_int (8)) + (Prims.of_int (163)) + (Prims.of_int (37))) + (Obj.magic (implies_intro ())) + (fun uu___2 -> + (fun eq_h -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.Logic.fst" + (Prims.of_int (161)) + (Prims.of_int (8)) + (Prims.of_int (161)) + (Prims.of_int (20))) + (Prims.mk_range + "FStar.Tactics.Logic.fst" + (Prims.of_int (162)) + (Prims.of_int (8)) + (Prims.of_int (163)) + (Prims.of_int (37))) + (Obj.magic + (FStar_Tactics_Builtins.rewrite + eq_h)) + (fun uu___2 -> (fun uu___2 -> - (fun uu___2 -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - ( - Prims.mk_range + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "FStar.Tactics.Logic.fst" (Prims.of_int (162)) (Prims.of_int (8)) (Prims.of_int (162)) (Prims.of_int (20))) - ( - Prims.mk_range + (Prims.mk_range "FStar.Tactics.Logic.fst" (Prims.of_int (163)) (Prims.of_int (8)) (Prims.of_int (163)) (Prims.of_int (37))) + (Obj.magic ( - Obj.magic - (FStar_Tactics_Builtins.clear_top + FStar_Tactics_Builtins.clear_top ())) - ( - fun - uu___3 -> + (fun uu___3 + -> (fun uu___3 -> Obj.magic (visit simplify_eq_implication)) uu___3))) - uu___2))) - uu___2)))) uu___1))) - uu___1))) uu___1) + uu___2))) uu___2))) + uu___1))) uu___1))) uu___1) let (rewrite_all_equalities : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = fun uu___ -> visit simplify_eq_implication @@ -862,52 +856,49 @@ let rec (unfold_definition_and_simplify_eq : match r with | FStar_Pervasives_Native.None -> Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "Not an equality implication")) + (FStar_Tactics_Derived.fail + "Not an equality implication") | FStar_Pervasives_Native.Some (uu___2, rhs) -> Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.Logic.fst" - (Prims.of_int (180)) - (Prims.of_int (23)) - (Prims.of_int (180)) - (Prims.of_int (39))) - (Prims.mk_range - "FStar.Tactics.Logic.fst" - (Prims.of_int (181)) - (Prims.of_int (12)) - (Prims.of_int (183)) - (Prims.of_int (66))) - (Obj.magic - (implies_intro ())) - (fun uu___3 -> - (fun eq_h -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.Logic.fst" - (Prims.of_int (181)) - (Prims.of_int (12)) - (Prims.of_int (181)) - (Prims.of_int (24))) - (Prims.mk_range - "FStar.Tactics.Logic.fst" - (Prims.of_int (182)) - (Prims.of_int (12)) - (Prims.of_int (183)) - (Prims.of_int (66))) - (Obj.magic - (FStar_Tactics_Builtins.rewrite - eq_h)) + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.Logic.fst" + (Prims.of_int (180)) + (Prims.of_int (23)) + (Prims.of_int (180)) + (Prims.of_int (39))) + (Prims.mk_range + "FStar.Tactics.Logic.fst" + (Prims.of_int (181)) + (Prims.of_int (12)) + (Prims.of_int (183)) + (Prims.of_int (66))) + (Obj.magic + (implies_intro ())) + (fun uu___3 -> + (fun eq_h -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.Logic.fst" + (Prims.of_int (181)) + (Prims.of_int (12)) + (Prims.of_int (181)) + (Prims.of_int (24))) + (Prims.mk_range + "FStar.Tactics.Logic.fst" + (Prims.of_int (182)) + (Prims.of_int (12)) + (Prims.of_int (183)) + (Prims.of_int (66))) + (Obj.magic + (FStar_Tactics_Builtins.rewrite + eq_h)) + (fun uu___3 -> (fun uu___3 -> - (fun - uu___3 -> - Obj.magic + Obj.magic (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.Logic.fst" @@ -935,8 +926,8 @@ let rec (unfold_definition_and_simplify_eq : unfold_definition_and_simplify_eq tm))) uu___4))) - uu___3))) - uu___3)))) uu___2)))) + uu___3))) + uu___3))) uu___2)))) uu___))) uu___) let (unsquash : FStar_Reflection_Types.term -> @@ -1247,10 +1238,7 @@ let (instantiate : (fa, FStar_Reflection_Data.Q_Explicit)))), (x, FStar_Reflection_Data.Q_Explicit))))) (fun uu___1 -> - (fun uu___1 -> - Obj.magic - (FStar_Tactics_Derived.fail "could not instantiate")) - uu___1)) + FStar_Tactics_Derived.fail "could not instantiate")) let (instantiate_as : FStar_Reflection_Types.term -> FStar_Reflection_Types.term -> @@ -1373,12 +1361,19 @@ let rec (sk_binder' : (fun uu___4 -> uu___3 <> Prims.int_one)))) (fun uu___3 -> - if uu___3 - then - FStar_Tactics_Derived.fail "no" - else - FStar_Tactics_Effect.lift_div_tac - (fun uu___5 -> ())))) + (fun uu___3 -> + if uu___3 + then + Obj.magic + (Obj.repr + (FStar_Tactics_Derived.fail + "no")) + else + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___5 -> ())))) + uu___3))) (fun uu___3 -> (fun uu___3 -> Obj.magic @@ -1573,7 +1568,5 @@ let (using_lemma : "lem3_fa"]))), (t, FStar_Reflection_Data.Q_Explicit))))) (fun uu___2 -> - (fun uu___2 -> - Obj.magic - (FStar_Tactics_Derived.fail - "using_lemma: failed to instantiate")) uu___2))) \ No newline at end of file + FStar_Tactics_Derived.fail + "using_lemma: failed to instantiate"))) \ No newline at end of file diff --git a/ocaml/fstar-lib/generated/FStar_Tactics_PatternMatching.ml b/ocaml/fstar-lib/generated/FStar_Tactics_PatternMatching.ml index 976c932a28e..de9f29a87f5 100644 --- a/ocaml/fstar-lib/generated/FStar_Tactics_PatternMatching.ml +++ b/ocaml/fstar-lib/generated/FStar_Tactics_PatternMatching.ml @@ -28,79 +28,73 @@ let (fetch_eq_side : | FStar_Reflection_Data.Tv_App (squash, (g1, uu___2)) -> Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.PatternMatching.fst" - (Prims.of_int (66)) (Prims.of_int (11)) - (Prims.of_int (66)) (Prims.of_int (25))) - (Prims.mk_range - "FStar.Tactics.PatternMatching.fst" - (Prims.of_int (66)) (Prims.of_int (4)) - (Prims.of_int (87)) (Prims.of_int (51))) - (Obj.magic - (FStar_Tactics_Builtins.inspect squash)) + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.PatternMatching.fst" + (Prims.of_int (66)) (Prims.of_int (11)) + (Prims.of_int (66)) (Prims.of_int (25))) + (Prims.mk_range + "FStar.Tactics.PatternMatching.fst" + (Prims.of_int (66)) (Prims.of_int (4)) + (Prims.of_int (87)) (Prims.of_int (51))) + (Obj.magic + (FStar_Tactics_Builtins.inspect squash)) + (fun uu___3 -> (fun uu___3 -> - (fun uu___3 -> - match uu___3 with - | FStar_Reflection_Data.Tv_UInst - (squash1, uu___4) -> - Obj.magic - (Obj.repr - (if - (FStar_Reflection_Derived.fv_to_string - squash1) - = - (FStar_Reflection_Derived.flatten_name - FStar_Reflection_Const.squash_qn) - then - Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.PatternMatching.fst" - (Prims.of_int (70)) - (Prims.of_int (16)) - (Prims.of_int (70)) - (Prims.of_int (25))) - (Prims.mk_range - "FStar.Tactics.PatternMatching.fst" - (Prims.of_int (70)) - (Prims.of_int (9)) - (Prims.of_int (85)) - (Prims.of_int (48))) - (Obj.magic - (FStar_Tactics_Builtins.inspect - g1)) - (fun uu___5 -> - (fun uu___5 -> - match uu___5 - with - | FStar_Reflection_Data.Tv_App - (eq_type_x, - (y, - uu___6)) - -> - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.PatternMatching.fst" - (Prims.of_int (72)) - (Prims.of_int (19)) - (Prims.of_int (72)) - (Prims.of_int (36))) - (Prims.mk_range - "FStar.Tactics.PatternMatching.fst" - (Prims.of_int (72)) - (Prims.of_int (12)) - (Prims.of_int (84)) - (Prims.of_int (39))) - (Obj.magic - (FStar_Tactics_Builtins.inspect + match uu___3 with + | FStar_Reflection_Data.Tv_UInst + (squash1, uu___4) -> + if + (FStar_Reflection_Derived.fv_to_string + squash1) + = + (FStar_Reflection_Derived.flatten_name + FStar_Reflection_Const.squash_qn) + then + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.PatternMatching.fst" + (Prims.of_int (70)) + (Prims.of_int (16)) + (Prims.of_int (70)) + (Prims.of_int (25))) + (Prims.mk_range + "FStar.Tactics.PatternMatching.fst" + (Prims.of_int (70)) + (Prims.of_int (9)) + (Prims.of_int (85)) + (Prims.of_int (48))) + (Obj.magic + (FStar_Tactics_Builtins.inspect + g1)) + (fun uu___5 -> + (fun uu___5 -> + match uu___5 with + | FStar_Reflection_Data.Tv_App + (eq_type_x, + (y, uu___6)) + -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.PatternMatching.fst" + (Prims.of_int (72)) + (Prims.of_int (19)) + (Prims.of_int (72)) + (Prims.of_int (36))) + (Prims.mk_range + "FStar.Tactics.PatternMatching.fst" + (Prims.of_int (72)) + (Prims.of_int (12)) + (Prims.of_int (84)) + (Prims.of_int (39))) + (Obj.magic + (FStar_Tactics_Builtins.inspect eq_type_x)) - (fun - uu___7 -> - (fun + (fun uu___7 + -> + (fun uu___7 -> match uu___7 with @@ -111,7 +105,6 @@ let (fetch_eq_side : uu___8)) -> Obj.magic - (Obj.repr (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.PatternMatching.fst" @@ -141,7 +134,6 @@ let (fetch_eq_side : uu___10)) -> Obj.magic - (Obj.repr (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.PatternMatching.fst" @@ -161,6 +153,9 @@ let (fetch_eq_side : (fun uu___11 -> + (fun + uu___11 + -> match uu___11 with | @@ -168,123 +163,126 @@ let (fetch_eq_side : (eq1, uu___12) -> - if + Obj.magic + (Obj.repr + (if (FStar_Reflection_Derived.fv_to_string eq1) = (FStar_Reflection_Derived.flatten_name FStar_Reflection_Const.eq2_qn) then - FStar_Tactics_Effect.lift_div_tac + Obj.repr + (FStar_Tactics_Effect.lift_div_tac (fun uu___13 - -> (x, y)) + -> (x, y))) else - FStar_Tactics_Derived.fail - "not an equality" + Obj.repr + (FStar_Tactics_Derived.fail + "not an equality"))) | FStar_Reflection_Data.Tv_FVar eq1 -> - if + Obj.magic + (Obj.repr + (if (FStar_Reflection_Derived.fv_to_string eq1) = (FStar_Reflection_Derived.flatten_name FStar_Reflection_Const.eq2_qn) then - FStar_Tactics_Effect.lift_div_tac + Obj.repr + (FStar_Tactics_Effect.lift_div_tac (fun uu___12 - -> (x, y)) + -> (x, y))) else - FStar_Tactics_Derived.fail - "not an equality" + Obj.repr + (FStar_Tactics_Derived.fail + "not an equality"))) | uu___12 -> - FStar_Tactics_Derived.fail + Obj.magic + (Obj.repr + (FStar_Tactics_Derived.fail "not an app2 of fvar: "))) + uu___11)) | uu___10 -> Obj.magic - (Obj.repr (FStar_Tactics_Derived.fail - "not an app3"))) - uu___9))) + "not an app3")) + uu___9)) | uu___8 -> Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "not an app2"))) - uu___7))) - | uu___6 -> - Obj.magic - (Obj.repr (FStar_Tactics_Derived.fail - "not an app under squash"))) - uu___5)) - else - Obj.repr - (FStar_Tactics_Derived.fail - "not a squash"))) - | FStar_Reflection_Data.Tv_FVar - squash1 -> - Obj.magic - (Obj.repr - (if - (FStar_Reflection_Derived.fv_to_string - squash1) - = - (FStar_Reflection_Derived.flatten_name - FStar_Reflection_Const.squash_qn) - then - Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.PatternMatching.fst" - (Prims.of_int (70)) - (Prims.of_int (16)) - (Prims.of_int (70)) - (Prims.of_int (25))) - (Prims.mk_range - "FStar.Tactics.PatternMatching.fst" - (Prims.of_int (70)) - (Prims.of_int (9)) - (Prims.of_int (85)) - (Prims.of_int (48))) - (Obj.magic - (FStar_Tactics_Builtins.inspect - g1)) - (fun uu___4 -> - (fun uu___4 -> - match uu___4 - with - | FStar_Reflection_Data.Tv_App - (eq_type_x, - (y, - uu___5)) - -> - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.PatternMatching.fst" - (Prims.of_int (72)) - (Prims.of_int (19)) - (Prims.of_int (72)) - (Prims.of_int (36))) - (Prims.mk_range - "FStar.Tactics.PatternMatching.fst" - (Prims.of_int (72)) - (Prims.of_int (12)) - (Prims.of_int (84)) - (Prims.of_int (39))) - (Obj.magic - (FStar_Tactics_Builtins.inspect + "not an app2")) + uu___7)) + | uu___6 -> + Obj.magic + (FStar_Tactics_Derived.fail + "not an app under squash")) + uu___5)) + else + Obj.magic + (FStar_Tactics_Derived.fail + "not a squash") + | FStar_Reflection_Data.Tv_FVar squash1 + -> + if + (FStar_Reflection_Derived.fv_to_string + squash1) + = + (FStar_Reflection_Derived.flatten_name + FStar_Reflection_Const.squash_qn) + then + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.PatternMatching.fst" + (Prims.of_int (70)) + (Prims.of_int (16)) + (Prims.of_int (70)) + (Prims.of_int (25))) + (Prims.mk_range + "FStar.Tactics.PatternMatching.fst" + (Prims.of_int (70)) + (Prims.of_int (9)) + (Prims.of_int (85)) + (Prims.of_int (48))) + (Obj.magic + (FStar_Tactics_Builtins.inspect + g1)) + (fun uu___4 -> + (fun uu___4 -> + match uu___4 with + | FStar_Reflection_Data.Tv_App + (eq_type_x, + (y, uu___5)) + -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.PatternMatching.fst" + (Prims.of_int (72)) + (Prims.of_int (19)) + (Prims.of_int (72)) + (Prims.of_int (36))) + (Prims.mk_range + "FStar.Tactics.PatternMatching.fst" + (Prims.of_int (72)) + (Prims.of_int (12)) + (Prims.of_int (84)) + (Prims.of_int (39))) + (Obj.magic + (FStar_Tactics_Builtins.inspect eq_type_x)) - (fun - uu___6 -> - (fun + (fun uu___6 + -> + (fun uu___6 -> match uu___6 with @@ -295,7 +293,6 @@ let (fetch_eq_side : uu___7)) -> Obj.magic - (Obj.repr (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.PatternMatching.fst" @@ -325,7 +322,6 @@ let (fetch_eq_side : uu___9)) -> Obj.magic - (Obj.repr (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.PatternMatching.fst" @@ -345,6 +341,9 @@ let (fetch_eq_side : (fun uu___10 -> + (fun + uu___10 + -> match uu___10 with | @@ -352,76 +351,81 @@ let (fetch_eq_side : (eq1, uu___11) -> - if + Obj.magic + (Obj.repr + (if (FStar_Reflection_Derived.fv_to_string eq1) = (FStar_Reflection_Derived.flatten_name FStar_Reflection_Const.eq2_qn) then - FStar_Tactics_Effect.lift_div_tac + Obj.repr + (FStar_Tactics_Effect.lift_div_tac (fun uu___12 - -> (x, y)) + -> (x, y))) else - FStar_Tactics_Derived.fail - "not an equality" + Obj.repr + (FStar_Tactics_Derived.fail + "not an equality"))) | FStar_Reflection_Data.Tv_FVar eq1 -> - if + Obj.magic + (Obj.repr + (if (FStar_Reflection_Derived.fv_to_string eq1) = (FStar_Reflection_Derived.flatten_name FStar_Reflection_Const.eq2_qn) then - FStar_Tactics_Effect.lift_div_tac + Obj.repr + (FStar_Tactics_Effect.lift_div_tac (fun uu___11 - -> (x, y)) + -> (x, y))) else - FStar_Tactics_Derived.fail - "not an equality" + Obj.repr + (FStar_Tactics_Derived.fail + "not an equality"))) | uu___11 -> - FStar_Tactics_Derived.fail + Obj.magic + (Obj.repr + (FStar_Tactics_Derived.fail "not an app2 of fvar: "))) + uu___10)) | uu___9 -> Obj.magic - (Obj.repr (FStar_Tactics_Derived.fail - "not an app3"))) - uu___8))) + "not an app3")) + uu___8)) | uu___7 -> Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "not an app2"))) - uu___6))) - | uu___5 -> - Obj.magic - (Obj.repr (FStar_Tactics_Derived.fail - "not an app under squash"))) - uu___4)) - else - Obj.repr - (FStar_Tactics_Derived.fail - "not a squash"))) - | uu___4 -> - Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "not an app of fvar at top level"))) - uu___3))) + "not an app2")) + uu___6)) + | uu___5 -> + Obj.magic + (FStar_Tactics_Derived.fail + "not an app under squash")) + uu___4)) + else + Obj.magic + (FStar_Tactics_Derived.fail + "not a squash") + | uu___4 -> + Obj.magic + (FStar_Tactics_Derived.fail + "not an app of fvar at top level")) + uu___3)) | uu___2 -> Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "not an app at top level"))) uu___1))) - uu___1) + (FStar_Tactics_Derived.fail + "not an app at top level")) uu___1))) uu___1) let mustfail : 'a . (unit -> ('a, unit) FStar_Tactics_Effect.tac_repr) -> @@ -437,11 +441,15 @@ let mustfail : (Prims.of_int (130)) (Prims.of_int (4)) (Prims.of_int (132)) (Prims.of_int (16))) (Obj.magic (FStar_Tactics_Derived.trytac t)) (fun uu___ -> - match uu___ with - | FStar_Pervasives_Native.Some uu___1 -> - FStar_Tactics_Derived.fail message - | FStar_Pervasives_Native.None -> - FStar_Tactics_Effect.lift_div_tac (fun uu___1 -> ())) + (fun uu___ -> + match uu___ with + | FStar_Pervasives_Native.Some uu___1 -> + Obj.magic (Obj.repr (FStar_Tactics_Derived.fail message)) + | FStar_Pervasives_Native.None -> + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac (fun uu___1 -> ())))) + uu___) let (implies_intro' : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = fun uu___ -> FStar_Tactics_Effect.tac_bind @@ -1037,8 +1045,10 @@ let lift_exn_tac : (Prims.of_int (268)) (Prims.of_int (18)) (Prims.of_int (268)) (Prims.of_int (61))) (Obj.magic (string_of_match_exception ex)) - (fun uu___ -> FStar_Tactics_Derived.fail uu___)))) - uu___1 uu___ + (fun uu___ -> + (fun uu___ -> + Obj.magic (FStar_Tactics_Derived.fail uu___)) + uu___)))) uu___1 uu___ let lift_exn_tactic : 'a 'b . ('a -> 'b match_res) -> 'a -> ('b, unit) FStar_Tactics_Effect.tac_repr @@ -1063,8 +1073,10 @@ let lift_exn_tactic : (Prims.of_int (273)) (Prims.of_int (18)) (Prims.of_int (273)) (Prims.of_int (61))) (Obj.magic (string_of_match_exception ex)) - (fun uu___ -> FStar_Tactics_Derived.fail uu___)))) - uu___1 uu___ + (fun uu___ -> + (fun uu___ -> + Obj.magic (FStar_Tactics_Derived.fail uu___)) + uu___)))) uu___1 uu___ type bindings = (varname * FStar_Reflection_Types.term) Prims.list let (string_of_bindings : bindings -> (Prims.string, unit) FStar_Tactics_Effect.tac_repr) = @@ -1484,8 +1496,10 @@ let (match_term : (Prims.of_int (338)) (Prims.of_int (20)) (Prims.of_int (338)) (Prims.of_int (63))) (Obj.magic (string_of_match_exception ex)) - (fun uu___1 -> FStar_Tactics_Derived.fail uu___1)))) - uu___) + (fun uu___1 -> + (fun uu___1 -> + Obj.magic (FStar_Tactics_Derived.fail uu___1)) + uu___1)))) uu___) let debug : 'uuuuu . 'uuuuu -> (unit, unit) FStar_Tactics_Effect.tac_repr = fun uu___ -> (fun msg -> @@ -1649,9 +1663,13 @@ let assoc_varname_fail : match FStar_List_Tot_Base.assoc key ls with | FStar_Pervasives_Native.None -> Obj.magic - (FStar_Tactics_Derived.fail (Prims.strcat "Not found: " key)) + (Obj.repr + (FStar_Tactics_Derived.fail + (Prims.strcat "Not found: " key))) | FStar_Pervasives_Native.Some x -> - Obj.magic (FStar_Tactics_Effect.lift_div_tac (fun uu___ -> x))) + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac (fun uu___ -> x)))) uu___1 uu___ let ms_locate_hyp : 'a . @@ -1694,123 +1712,98 @@ let rec solve_mp_for_single_hyp : (matching_solution -> ('a, unit) FStar_Tactics_Effect.tac_repr) -> matching_solution -> ('a, unit) FStar_Tactics_Effect.tac_repr = - fun uu___4 -> - fun uu___3 -> - fun uu___2 -> - fun uu___1 -> - fun uu___ -> - (fun name -> - fun pat -> - fun hypotheses -> - fun body -> - fun part_sol -> - match hypotheses with - | [] -> - Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "No matching hypothesis")) - | h::hs -> - Obj.magic - (Obj.repr - (FStar_Tactics_Derived.or_else - (fun uu___ -> - FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.PatternMatching.fst" - (Prims.of_int (448)) - (Prims.of_int (15)) - (Prims.of_int (448)) - (Prims.of_int (73))) - (Prims.mk_range - "FStar.Tactics.PatternMatching.fst" - (Prims.of_int (448)) - (Prims.of_int (9)) - (Prims.of_int (453)) - (Prims.of_int (73))) - (Obj.magic - (interp_pattern_aux pat - part_sol.ms_vars - (FStar_Reflection_Derived.type_of_binder - h))) - (fun uu___1 -> - (fun uu___1 -> - match uu___1 with - | Failure ex -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.PatternMatching.fst" - (Prims.of_int (450)) - (Prims.of_int (16)) - (Prims.of_int (450)) - (Prims.of_int (74))) - (Prims.mk_range - "FStar.Tactics.PatternMatching.fst" - (Prims.of_int (450)) - (Prims.of_int (11)) - (Prims.of_int (450)) - (Prims.of_int (74))) - (Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.PatternMatching.fst" - (Prims.of_int (450)) - (Prims.of_int (43)) - (Prims.of_int (450)) - (Prims.of_int (73))) - (Prims.mk_range - "prims.fst" - (Prims.of_int (606)) - (Prims.of_int (19)) - (Prims.of_int (606)) - (Prims.of_int (31))) - (Obj.magic - (string_of_match_exception - ex)) - (fun uu___2 -> - FStar_Tactics_Effect.lift_div_tac - (fun uu___3 - -> - Prims.strcat - "Failed to match hyp: " - uu___2)))) - (fun uu___2 -> - FStar_Tactics_Derived.fail - uu___2)) - | Success bindings1 -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.PatternMatching.fst" - (Prims.of_int (452)) - (Prims.of_int (35)) - (Prims.of_int (452)) - (Prims.of_int (37))) - (Prims.mk_range - "FStar.Tactics.PatternMatching.fst" - (Prims.of_int (453)) - (Prims.of_int (11)) - (Prims.of_int (453)) - (Prims.of_int (73))) - (FStar_Tactics_Effect.lift_div_tac - (fun uu___2 -> - (name, h) :: - (part_sol.ms_hyps))) - (fun uu___2 -> - (fun ms_hyps -> - Obj.magic - (body - { - ms_vars = - bindings1; - ms_hyps - })) uu___2))) - uu___1)) - (fun uu___ -> - solve_mp_for_single_hyp name pat hs - body part_sol)))) uu___4 uu___3 - uu___2 uu___1 uu___ + fun name -> + fun pat -> + fun hypotheses -> + fun body -> + fun part_sol -> + match hypotheses with + | [] -> FStar_Tactics_Derived.fail "No matching hypothesis" + | h::hs -> + FStar_Tactics_Derived.or_else + (fun uu___ -> + FStar_Tactics_Effect.tac_bind + (Prims.mk_range "FStar.Tactics.PatternMatching.fst" + (Prims.of_int (448)) (Prims.of_int (15)) + (Prims.of_int (448)) (Prims.of_int (73))) + (Prims.mk_range "FStar.Tactics.PatternMatching.fst" + (Prims.of_int (448)) (Prims.of_int (9)) + (Prims.of_int (453)) (Prims.of_int (73))) + (Obj.magic + (interp_pattern_aux pat part_sol.ms_vars + (FStar_Reflection_Derived.type_of_binder h))) + (fun uu___1 -> + (fun uu___1 -> + match uu___1 with + | Failure ex -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.PatternMatching.fst" + (Prims.of_int (450)) + (Prims.of_int (16)) + (Prims.of_int (450)) + (Prims.of_int (74))) + (Prims.mk_range + "FStar.Tactics.PatternMatching.fst" + (Prims.of_int (450)) + (Prims.of_int (11)) + (Prims.of_int (450)) + (Prims.of_int (74))) + (Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.PatternMatching.fst" + (Prims.of_int (450)) + (Prims.of_int (43)) + (Prims.of_int (450)) + (Prims.of_int (73))) + (Prims.mk_range "prims.fst" + (Prims.of_int (606)) + (Prims.of_int (19)) + (Prims.of_int (606)) + (Prims.of_int (31))) + (Obj.magic + (string_of_match_exception ex)) + (fun uu___2 -> + FStar_Tactics_Effect.lift_div_tac + (fun uu___3 -> + Prims.strcat + "Failed to match hyp: " + uu___2)))) + (fun uu___2 -> + (fun uu___2 -> + Obj.magic + (FStar_Tactics_Derived.fail + uu___2)) uu___2)) + | Success bindings1 -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.PatternMatching.fst" + (Prims.of_int (452)) + (Prims.of_int (35)) + (Prims.of_int (452)) + (Prims.of_int (37))) + (Prims.mk_range + "FStar.Tactics.PatternMatching.fst" + (Prims.of_int (453)) + (Prims.of_int (11)) + (Prims.of_int (453)) + (Prims.of_int (73))) + (FStar_Tactics_Effect.lift_div_tac + (fun uu___2 -> (name, h) :: + (part_sol.ms_hyps))) + (fun uu___2 -> + (fun ms_hyps -> + Obj.magic + (body + { + ms_vars = bindings1; + ms_hyps + })) uu___2))) uu___1)) + (fun uu___ -> + solve_mp_for_single_hyp name pat hs body part_sol) let rec solve_mp_for_hyps : 'a . (varname * pattern) Prims.list -> @@ -1905,8 +1898,10 @@ let solve_mp : "Failed to match goal: " uu___1)))) (fun uu___1 -> - FStar_Tactics_Derived.fail - uu___1))) + (fun uu___1 -> + Obj.magic + (FStar_Tactics_Derived.fail + uu___1)) uu___1))) | Success bindings1 -> Obj.magic (Obj.repr @@ -2047,8 +2042,10 @@ let (pattern_of_term : (Prims.of_int (532)) (Prims.of_int (20)) (Prims.of_int (532)) (Prims.of_int (63))) (Obj.magic (string_of_match_exception ex)) - (fun uu___1 -> FStar_Tactics_Derived.fail uu___1)))) - uu___) + (fun uu___1 -> + (fun uu___1 -> + Obj.magic (FStar_Tactics_Derived.fail uu___1)) + uu___1)))) uu___) type 'a hyp = FStar_Reflection_Types.binder type 'a pm_goal = unit let (hyp_qn : Prims.string) = "FStar.Tactics.PatternMatching.hyp" @@ -2208,6 +2205,8 @@ let (classify_abspat_binder : ( fun uu___2 -> + (fun + uu___2 -> match uu___2 with | @@ -2215,25 +2214,32 @@ let (classify_abspat_binder : ((uu___3, goal_typ)::[]) -> - FStar_Tactics_Effect.lift_div_tac + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac (fun uu___4 -> (ABKGoal, - goal_typ)) + goal_typ)))) | Success uu___3 -> - FStar_Tactics_Derived.fail - "classifiy_abspat_binder: impossible (2)" + Obj.magic + (Obj.repr + (FStar_Tactics_Derived.fail + "classifiy_abspat_binder: impossible (2)")) | Failure uu___3 -> - FStar_Tactics_Effect.lift_div_tac + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac (fun uu___4 -> ((ABKVar typ), - typ)))))) + typ))))) + uu___2)))) uu___))) uu___))) uu___))) uu___))) uu___) let rec (binders_and_body_of_abs : diff --git a/ocaml/fstar-lib/generated/FStar_Tactics_Simplifier.ml b/ocaml/fstar-lib/generated/FStar_Tactics_Simplifier.ml index e6d04f4ca30..d3708421bb0 100644 --- a/ocaml/fstar-lib/generated/FStar_Tactics_Simplifier.ml +++ b/ocaml/fstar-lib/generated/FStar_Tactics_Simplifier.ml @@ -144,55 +144,45 @@ let (inhabit : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = match FStar_Reflection_Builtins.inspect_ln t with | FStar_Reflection_Data.Tv_FVar fv -> Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "FStar.Tactics.Simplifier.fst" - (Prims.of_int (200)) (Prims.of_int (17)) - (Prims.of_int (200)) (Prims.of_int (30))) - (Prims.mk_range "FStar.Tactics.Simplifier.fst" - (Prims.of_int (201)) (Prims.of_int (13)) - (Prims.of_int (204)) (Prims.of_int (20))) - (FStar_Tactics_Effect.lift_div_tac - (fun uu___1 -> - FStar_Reflection_Builtins.inspect_fv fv)) + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "FStar.Tactics.Simplifier.fst" + (Prims.of_int (200)) (Prims.of_int (17)) + (Prims.of_int (200)) (Prims.of_int (30))) + (Prims.mk_range "FStar.Tactics.Simplifier.fst" + (Prims.of_int (201)) (Prims.of_int (13)) + (Prims.of_int (204)) (Prims.of_int (20))) + (FStar_Tactics_Effect.lift_div_tac (fun uu___1 -> - (fun qn -> - if qn = FStar_Reflection_Const.int_lid - then - Obj.magic - (Obj.repr - (FStar_Tactics_Derived.exact - (FStar_Reflection_Builtins.pack_ln - (FStar_Reflection_Data.Tv_Const - (FStar_Reflection_Data.C_Int - (Prims.of_int (42))))))) - else - Obj.magic - (Obj.repr - (if qn = FStar_Reflection_Const.bool_lid - then - Obj.repr - (FStar_Tactics_Derived.exact - (FStar_Reflection_Builtins.pack_ln - (FStar_Reflection_Data.Tv_Const - FStar_Reflection_Data.C_True))) - else - Obj.repr - (if - qn = - FStar_Reflection_Const.unit_lid - then - Obj.repr - (FStar_Tactics_Derived.exact - (FStar_Reflection_Builtins.pack_ln - (FStar_Reflection_Data.Tv_Const - FStar_Reflection_Data.C_Unit))) - else - Obj.repr - (FStar_Tactics_Derived.fail ""))))) - uu___1))) - | uu___1 -> Obj.magic (Obj.repr (FStar_Tactics_Derived.fail ""))) - uu___1) + FStar_Reflection_Builtins.inspect_fv fv)) + (fun uu___1 -> + (fun qn -> + if qn = FStar_Reflection_Const.int_lid + then + Obj.magic + (FStar_Tactics_Derived.exact + (FStar_Reflection_Builtins.pack_ln + (FStar_Reflection_Data.Tv_Const + (FStar_Reflection_Data.C_Int + (Prims.of_int (42)))))) + else + if qn = FStar_Reflection_Const.bool_lid + then + Obj.magic + (FStar_Tactics_Derived.exact + (FStar_Reflection_Builtins.pack_ln + (FStar_Reflection_Data.Tv_Const + FStar_Reflection_Data.C_True))) + else + if qn = FStar_Reflection_Const.unit_lid + then + Obj.magic + (FStar_Tactics_Derived.exact + (FStar_Reflection_Builtins.pack_ln + (FStar_Reflection_Data.Tv_Const + FStar_Reflection_Data.C_Unit))) + else Obj.magic (FStar_Tactics_Derived.fail "")) + uu___1)) + | uu___1 -> Obj.magic (FStar_Tactics_Derived.fail "")) uu___1) let rec (simplify_point : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = fun uu___ -> @@ -249,31 +239,29 @@ let rec (simplify_point : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) | FStar_Reflection_Formula.Iff (l, r) -> Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.Simplifier.fst" - (Prims.of_int (217)) - (Prims.of_int (20)) - (Prims.of_int (217)) - (Prims.of_int (38))) - (Prims.mk_range - "FStar.Tactics.Simplifier.fst" - (Prims.of_int (217)) - (Prims.of_int (14)) - (Prims.of_int (264)) - (Prims.of_int (22))) - (Obj.magic - (FStar_Reflection_Formula.term_as_formula' - l)) + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.Simplifier.fst" + (Prims.of_int (217)) + (Prims.of_int (20)) + (Prims.of_int (217)) + (Prims.of_int (38))) + (Prims.mk_range + "FStar.Tactics.Simplifier.fst" + (Prims.of_int (217)) + (Prims.of_int (14)) + (Prims.of_int (264)) + (Prims.of_int (22))) + (Obj.magic + (FStar_Reflection_Formula.term_as_formula' + l)) + (fun uu___3 -> (fun uu___3 -> - (fun uu___3 -> - match uu___3 - with - | FStar_Reflection_Formula.And - (p, q) -> - Obj.magic - (FStar_Tactics_Effect.tac_bind + match uu___3 with + | FStar_Reflection_Formula.And + (p, q) -> + Obj.magic + (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.Simplifier.fst" (Prims.of_int (219)) @@ -413,10 +401,10 @@ let rec (simplify_point : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) uu___8))) uu___6))) uu___4)) - | FStar_Reflection_Formula.Or - (p, q) -> - Obj.magic - (FStar_Tactics_Effect.tac_bind + | FStar_Reflection_Formula.Or + (p, q) -> + Obj.magic + (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.Simplifier.fst" (Prims.of_int (226)) @@ -556,10 +544,10 @@ let rec (simplify_point : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) uu___8))) uu___6))) uu___4)) - | FStar_Reflection_Formula.Implies - (p, q) -> - Obj.magic - (FStar_Tactics_Effect.tac_bind + | FStar_Reflection_Formula.Implies + (p, q) -> + Obj.magic + (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.Simplifier.fst" (Prims.of_int (233)) @@ -662,10 +650,10 @@ let rec (simplify_point : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) uu___8))) uu___6))) uu___4)) - | FStar_Reflection_Formula.Forall - (b, p) -> - Obj.magic - (FStar_Tactics_Effect.tac_bind + | FStar_Reflection_Formula.Forall + (b, p) -> + Obj.magic + (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.Simplifier.fst" (Prims.of_int (239)) @@ -760,10 +748,10 @@ let rec (simplify_point : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) (tiff ())) uu___6))) uu___4)) - | FStar_Reflection_Formula.Exists - (b, p) -> - Obj.magic - (FStar_Tactics_Effect.tac_bind + | FStar_Reflection_Formula.Exists + (b, p) -> + Obj.magic + (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.Simplifier.fst" (Prims.of_int (244)) @@ -858,10 +846,10 @@ let rec (simplify_point : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) (tiff ())) uu___6))) uu___4)) - | FStar_Reflection_Formula.Not - p -> - Obj.magic - (FStar_Tactics_Effect.tac_bind + | FStar_Reflection_Formula.Not + p -> + Obj.magic + (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.Simplifier.fst" (Prims.of_int (249)) @@ -930,10 +918,10 @@ let rec (simplify_point : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) (tiff ())) uu___6))) uu___4)) - | FStar_Reflection_Formula.Iff - (p, q) -> - Obj.magic - (FStar_Tactics_Effect.tac_bind + | FStar_Reflection_Formula.Iff + (p, q) -> + Obj.magic + (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.Simplifier.fst" (Prims.of_int (256)) @@ -1116,15 +1104,14 @@ let rec (simplify_point : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) ())) uu___5))) uu___4)) - | uu___4 -> - Obj.magic - (tiff ())) - uu___3))) + | uu___4 -> + Obj.magic + (tiff ())) + uu___3)) | uu___3 -> Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "simplify_point: failed precondition: goal should be `g <==> ?u`"))) + (FStar_Tactics_Derived.fail + "simplify_point: failed precondition: goal should be `g <==> ?u`")) uu___3))) uu___3))) uu___2))) uu___1) and (recurse : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = @@ -1182,33 +1169,31 @@ and (recurse : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = | FStar_Reflection_Formula.Iff (l, r) -> Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.Simplifier.fst" - (Prims.of_int (275)) - (Prims.of_int (20)) - (Prims.of_int (275)) - (Prims.of_int (38))) - (Prims.mk_range - "FStar.Tactics.Simplifier.fst" - (Prims.of_int (275)) - (Prims.of_int (14)) - (Prims.of_int (302)) - (Prims.of_int (22))) - (Obj.magic - (FStar_Reflection_Formula.term_as_formula' - l)) + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.Simplifier.fst" + (Prims.of_int (275)) + (Prims.of_int (20)) + (Prims.of_int (275)) + (Prims.of_int (38))) + (Prims.mk_range + "FStar.Tactics.Simplifier.fst" + (Prims.of_int (275)) + (Prims.of_int (14)) + (Prims.of_int (302)) + (Prims.of_int (22))) + (Obj.magic + (FStar_Reflection_Formula.term_as_formula' + l)) + (fun uu___3 -> (fun uu___3 -> - (fun uu___3 -> - match uu___3 - with - | FStar_Reflection_Formula.And - (uu___4, - uu___5) - -> - Obj.magic - (FStar_Tactics_Derived.seq + match uu___3 with + | FStar_Reflection_Formula.And + (uu___4, + uu___5) + -> + Obj.magic + (FStar_Tactics_Derived.seq (fun uu___6 -> FStar_Tactics_Derived.apply_lemma @@ -1220,12 +1205,12 @@ and (recurse : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = "Simplifier"; "and_cong"])))) simplify_point) - | FStar_Reflection_Formula.Or - (uu___4, - uu___5) - -> - Obj.magic - (FStar_Tactics_Derived.seq + | FStar_Reflection_Formula.Or + (uu___4, + uu___5) + -> + Obj.magic + (FStar_Tactics_Derived.seq (fun uu___6 -> FStar_Tactics_Derived.apply_lemma @@ -1237,12 +1222,12 @@ and (recurse : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = "Simplifier"; "or_cong"])))) simplify_point) - | FStar_Reflection_Formula.Implies - (uu___4, - uu___5) - -> - Obj.magic - (FStar_Tactics_Derived.seq + | FStar_Reflection_Formula.Implies + (uu___4, + uu___5) + -> + Obj.magic + (FStar_Tactics_Derived.seq (fun uu___6 -> FStar_Tactics_Derived.apply_lemma @@ -1254,12 +1239,12 @@ and (recurse : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = "Simplifier"; "imp_cong"])))) simplify_point) - | FStar_Reflection_Formula.Forall - (uu___4, - uu___5) - -> - Obj.magic - (FStar_Tactics_Effect.tac_bind + | FStar_Reflection_Formula.Forall + (uu___4, + uu___5) + -> + Obj.magic + (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.Simplifier.fst" (Prims.of_int (286)) @@ -1311,12 +1296,12 @@ and (recurse : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = ())) uu___7))) uu___6)) - | FStar_Reflection_Formula.Exists - (uu___4, - uu___5) - -> - Obj.magic - (FStar_Tactics_Effect.tac_bind + | FStar_Reflection_Formula.Exists + (uu___4, + uu___5) + -> + Obj.magic + (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.Simplifier.fst" (Prims.of_int (291)) @@ -1368,10 +1353,10 @@ and (recurse : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = ())) uu___7))) uu___6)) - | FStar_Reflection_Formula.Not - uu___4 -> - Obj.magic - (FStar_Tactics_Effect.tac_bind + | FStar_Reflection_Formula.Not + uu___4 -> + Obj.magic + (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.Simplifier.fst" (Prims.of_int (296)) @@ -1401,12 +1386,12 @@ and (recurse : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = (simplify_point ())) uu___5)) - | FStar_Reflection_Formula.Iff - (uu___4, - uu___5) - -> - Obj.magic - (FStar_Tactics_Derived.seq + | FStar_Reflection_Formula.Iff + (uu___4, + uu___5) + -> + Obj.magic + (FStar_Tactics_Derived.seq (fun uu___6 -> FStar_Tactics_Derived.apply_lemma @@ -1418,15 +1403,14 @@ and (recurse : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = "Simplifier"; "iff_cong"])))) simplify_point) - | uu___4 -> - Obj.magic - (tiff ())) - uu___3))) + | uu___4 -> + Obj.magic + (tiff ())) + uu___3)) | uu___3 -> Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "recurse: failed precondition: goal should be `g <==> ?u`"))) + (FStar_Tactics_Derived.fail + "recurse: failed precondition: goal should be `g <==> ?u`")) uu___3))) uu___3))) uu___2))) uu___1) let (simplify : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = diff --git a/ocaml/fstar-lib/generated/FStar_Tactics_SyntaxHelpers.ml b/ocaml/fstar-lib/generated/FStar_Tactics_SyntaxHelpers.ml index 805016faf18..6f6a6fa5759 100644 --- a/ocaml/fstar-lib/generated/FStar_Tactics_SyntaxHelpers.ml +++ b/ocaml/fstar-lib/generated/FStar_Tactics_SyntaxHelpers.ml @@ -123,11 +123,7 @@ let (collect_abs : match uu___ with | (bs, t') -> ((FStar_List_Tot_Base.rev bs), t'))) let fail : 'a . Prims.string -> ('a, unit) FStar_Tactics_Effect.tac_repr = - fun uu___ -> - (fun m -> - Obj.magic - (FStar_Tactics_Effect.raise (FStar_Tactics_Common.TacticFailure m))) - uu___ + fun m -> FStar_Tactics_Effect.raise (FStar_Tactics_Common.TacticFailure m) let rec (mk_arr : FStar_Reflection_Types.binder Prims.list -> FStar_Reflection_Types.comp -> diff --git a/ocaml/fstar-lib/generated/FStar_Tactics_Typeclasses.ml b/ocaml/fstar-lib/generated/FStar_Tactics_Typeclasses.ml index b296a249006..df399efc6e2 100644 --- a/ocaml/fstar-lib/generated/FStar_Tactics_Typeclasses.ml +++ b/ocaml/fstar-lib/generated/FStar_Tactics_Typeclasses.ml @@ -4,18 +4,13 @@ let rec first : ('a -> ('b, unit) FStar_Tactics_Effect.tac_repr) -> 'a Prims.list -> ('b, unit) FStar_Tactics_Effect.tac_repr = - fun uu___1 -> - fun uu___ -> - (fun f -> - fun l -> - match l with - | [] -> - Obj.magic (Obj.repr (FStar_Tactics_Derived.fail "no cands")) - | x::xs -> - Obj.magic - (Obj.repr - (FStar_Tactics_Derived.or_else (fun uu___ -> f x) - (fun uu___ -> first f xs)))) uu___1 uu___ + fun f -> + fun l -> + match l with + | [] -> FStar_Tactics_Derived.fail "no cands" + | x::xs -> + FStar_Tactics_Derived.or_else (fun uu___ -> f x) + (fun uu___ -> first f xs) let rec (tcresolve' : FStar_Reflection_Types.term Prims.list -> Prims.int -> (unit, unit) FStar_Tactics_Effect.tac_repr) @@ -28,8 +23,10 @@ let rec (tcresolve' : (Prims.mk_range "FStar.Tactics.Typeclasses.fst" (Prims.of_int (45)) (Prims.of_int (4)) (Prims.of_int (50)) (Prims.of_int (137))) (if fuel <= Prims.int_zero - then FStar_Tactics_Derived.fail "out of fuel" - else FStar_Tactics_Effect.lift_div_tac (fun uu___1 -> ())) + then Obj.magic (Obj.repr (FStar_Tactics_Derived.fail "out of fuel")) + else + Obj.magic + (Obj.repr (FStar_Tactics_Effect.lift_div_tac (fun uu___1 -> ())))) (fun uu___ -> (fun uu___ -> Obj.magic @@ -75,10 +72,15 @@ let rec (tcresolve' : (FStar_Reflection_Builtins.term_eq g) seen then - FStar_Tactics_Derived.fail "loop" + Obj.magic + (Obj.repr + (FStar_Tactics_Derived.fail + "loop")) else - FStar_Tactics_Effect.lift_div_tac - (fun uu___3 -> ())) + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___3 -> ())))) (fun uu___2 -> (fun uu___2 -> Obj.magic @@ -153,7 +155,11 @@ let rec (tcresolve' : uu___5)))) (fun uu___5 -> - FStar_Tactics_Derived.fail + (fun + uu___5 -> + Obj.magic + (FStar_Tactics_Derived.fail + uu___5)) uu___5))))) uu___3))) uu___2))) uu___2))) uu___1))) uu___) @@ -354,15 +360,11 @@ let (tcresolve : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = (fun uu___2 -> match () with | () -> tcresolve' [] (Prims.of_int (16))) (fun uu___2 -> - (fun uu___2 -> - match uu___2 with - | FStar_Tactics_Common.TacticFailure s -> - Obj.magic - (FStar_Tactics_Derived.fail - (Prims.strcat "Typeclass resolution failed: " - s)) - | e -> Obj.magic (FStar_Tactics_Effect.raise e)) - uu___2))) uu___1) + match uu___2 with + | FStar_Tactics_Common.TacticFailure s -> + FStar_Tactics_Derived.fail + (Prims.strcat "Typeclass resolution failed: " s) + | e -> FStar_Tactics_Effect.raise e))) uu___1) let _ = FStar_Tactics_Native.register_tactic "FStar.Tactics.Typeclasses.tcresolve" (Prims.of_int (2)) @@ -1098,14 +1100,11 @@ let (mk_class : FStar_Pervasives_Native.None -> Obj.magic - (Obj.repr (FStar_Tactics_Derived.fail - "mk_class: proj not found?")) + "mk_class: proj not found?") | FStar_Pervasives_Native.Some se1 -> - Obj.magic - (Obj.repr (match FStar_Reflection_Builtins.inspect_sigelt se1 @@ -1114,7 +1113,7 @@ let (mk_class : FStar_Reflection_Data.Sg_Let (uu___10, lbs) -> - Obj.repr + Obj.magic (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.Typeclasses.fst" @@ -1155,9 +1154,9 @@ let (mk_class : | uu___10 -> - Obj.repr + Obj.magic (FStar_Tactics_Derived.fail - "mk_class: proj not Sg_Let?")))) + "mk_class: proj not Sg_Let?"))) uu___9))) (fun uu___9 -> @@ -1242,14 +1241,12 @@ let (mk_class : | [] -> Obj.magic - (Obj.repr (FStar_Tactics_Derived.fail - "mk_class: impossible, no binders")) + "mk_class: impossible, no binders") | b1::bs' -> Obj.magic - (Obj.repr (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.Typeclasses.fst" @@ -1314,7 +1311,7 @@ let (mk_class : :: bs')) cod2)) uu___12))) - uu___11))))) + uu___11)))) uu___10))) uu___9))) (fun diff --git a/ocaml/fstar-lib/generated/Steel_Effect_Common.ml b/ocaml/fstar-lib/generated/Steel_Effect_Common.ml index cdd2331bbcf..afb14f087b1 100644 --- a/ocaml/fstar-lib/generated/Steel_Effect_Common.ml +++ b/ocaml/fstar-lib/generated/Steel_Effect_Common.ml @@ -804,15 +804,26 @@ let (name_appears_in : (Prims.of_int (730)) (Prims.of_int (13))) (Obj.magic (FStar_Tactics_Builtins.inspect t1)) (fun uu___1 -> - match uu___1 with - | FStar_Reflection_Data.Tv_FVar fv -> - if (FStar_Reflection_Builtins.inspect_fv fv) = nm - then FStar_Tactics_Effect.raise Appears - else - FStar_Tactics_Effect.lift_div_tac - (fun uu___3 -> ()) - | t2 -> - FStar_Tactics_Effect.lift_div_tac (fun uu___2 -> ())))) + (fun uu___1 -> + match uu___1 with + | FStar_Reflection_Data.Tv_FVar fv -> + Obj.magic + (Obj.repr + (if + (FStar_Reflection_Builtins.inspect_fv fv) + = nm + then + Obj.repr + (FStar_Tactics_Effect.raise Appears) + else + Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___3 -> ())))) + | t2 -> + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___2 -> ())))) uu___1))) (fun uu___ -> (fun ff -> Obj.magic @@ -847,9 +858,12 @@ let (name_appears_in : match uu___ with | Appears -> Obj.magic - (FStar_Tactics_Effect.lift_div_tac - (fun uu___1 -> true)) - | e -> Obj.magic (FStar_Tactics_Effect.raise e)) + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___1 -> true))) + | e -> + Obj.magic + (Obj.repr (FStar_Tactics_Effect.raise e))) uu___))) uu___) let (term_appears_in : FStar_Reflection_Types.term -> @@ -1298,13 +1312,17 @@ let rec (try_candidates : (select t am) (select hd am))) (fun uu___3 -> - if uu___3 - then - FStar_Tactics_Effect.raise - Success - else - FStar_Tactics_Effect.raise - Failed)) + (fun uu___3 -> + if uu___3 + then + Obj.magic + (FStar_Tactics_Effect.raise + Success) + else + Obj.magic + (FStar_Tactics_Effect.raise + Failed)) + uu___3)) (fun uu___2 -> (fun uu___2 -> Obj.magic @@ -1366,8 +1384,9 @@ let rec (remove_from_list : (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (863)) (Prims.of_int (72)) (Prims.of_int (863)) (Prims.of_int (74))) - (FStar_Tactics_Derived.fail - "atom in remove_from_list not found: should not happen") + (Obj.magic + (FStar_Tactics_Derived.fail + "atom in remove_from_list not found: should not happen")) (fun uu___ -> FStar_Tactics_Effect.lift_div_tac (fun uu___1 -> [])))) @@ -1729,10 +1748,16 @@ let rec (try_unifying_remaining : (FStar_Tactics_Derived.unify u (select hd am))) (fun uu___1 -> - if uu___1 - then - FStar_Tactics_Effect.raise Success - else FStar_Tactics_Effect.raise Failed)) + (fun uu___1 -> + if uu___1 + then + Obj.magic + (FStar_Tactics_Effect.raise + Success) + else + Obj.magic + (FStar_Tactics_Effect.raise + Failed)) uu___1)) (fun uu___ -> match uu___ with | Success -> try_unifying_remaining tl u am @@ -1773,8 +1798,11 @@ let rec (try_unifying_remaining : "could not find candidate for scrutinee " uu___2)))) (fun uu___2 -> - FStar_Tactics_Derived.fail uu___2))))) - uu___2 uu___1 uu___ + (fun uu___2 -> + Obj.magic + (FStar_Tactics_Derived.fail + uu___2)) uu___2))))) uu___2 + uu___1 uu___ let (is_smt_binder : FStar_Reflection_Types.binder -> (Prims.bool, unit) FStar_Tactics_Effect.tac_repr) @@ -2057,11 +2085,13 @@ let rec (new_args_for_smt_attrs : with | FStar_Reflection_Data.C_Total ty2 -> - FStar_Tactics_Effect.lift_div_tac + Obj.magic ( - fun + Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___1 -> - ty2) + ty2))) | FStar_Reflection_Data.C_Eff (uu___1, eff_name, @@ -2069,22 +2099,30 @@ let rec (new_args_for_smt_attrs : uu___2, uu___3) -> - if + Obj.magic + ( + Obj.repr + (if eff_name = ["Prims"; "Tot"] - then - FStar_Tactics_Effect.lift_div_tac + then + Obj.repr + (FStar_Tactics_Effect.lift_div_tac (fun uu___4 -> - ty2) - else - FStar_Tactics_Derived.fail - "computation type not supported in definition of slprops" + ty2)) + else + Obj.repr + (FStar_Tactics_Derived.fail + "computation type not supported in definition of slprops"))) | uu___1 -> - FStar_Tactics_Derived.fail - "computation type not supported in definition of slprops") + Obj.magic + ( + Obj.repr + (FStar_Tactics_Derived.fail + "computation type not supported in definition of slprops"))) (fun uu___1 -> (fun ty2 -> Obj.magic @@ -2322,7 +2360,9 @@ let fail_atoms : (fun uu___1 -> Prims.strcat "could not find a solution for unifying\n" uu___)))) - (fun uu___ -> FStar_Tactics_Derived.fail uu___) + (fun uu___ -> + (fun uu___ -> Obj.magic (FStar_Tactics_Derived.fail uu___)) + uu___) let rec (equivalent_lists_fallback : Prims.nat -> atom Prims.list -> @@ -2431,7 +2471,11 @@ let rec (equivalent_lists_fallback : "could not find candidates for " uu___2)))) (fun uu___2 -> - FStar_Tactics_Derived.fail + (fun uu___2 + -> + Obj.magic + (FStar_Tactics_Derived.fail + uu___2)) uu___2)))) uu___)) | uu___ -> @@ -2473,8 +2517,10 @@ let rec (equivalent_lists_fallback : "could not find candidates for " uu___1)))) (fun uu___1 -> - FStar_Tactics_Derived.fail - uu___1)))) + (fun uu___1 -> + Obj.magic + (FStar_Tactics_Derived.fail + uu___1)) uu___1)))) | uu___ -> Obj.magic (Obj.repr @@ -2721,7 +2767,11 @@ let rec (equivalent_lists' : uu___2)))) (fun uu___2 -> - FStar_Tactics_Derived.fail + (fun + uu___2 -> + Obj.magic + (FStar_Tactics_Derived.fail + uu___2)) uu___2)))) uu___)) | uu___ -> @@ -2763,8 +2813,10 @@ let rec (equivalent_lists' : "could not find candidates for " uu___1)))) (fun uu___1 -> - FStar_Tactics_Derived.fail - uu___1)))) + (fun uu___1 -> + Obj.magic + (FStar_Tactics_Derived.fail + uu___1)) uu___1)))) | uu___ -> Obj.magic (Obj.repr @@ -3119,13 +3171,18 @@ let rec (unifies_with_all_uvars : t hd_t)) (fun uu___3 -> + (fun + uu___3 -> if uu___3 then - FStar_Tactics_Effect.raise - Success + Obj.magic + (FStar_Tactics_Effect.raise + Success) else - FStar_Tactics_Effect.raise + Obj.magic + (FStar_Tactics_Effect.raise Failed)) + uu___3)) (fun uu___2 -> (fun @@ -4286,93 +4343,96 @@ let rec (unify_pr_with_true : | (hd, tl) -> if is_and hd then - Obj.magic - (Obj.repr - (match tl with - | (pr_l, uu___1)::(pr_r, uu___2)::[] -> - Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (1795)) (Prims.of_int (6)) - (Prims.of_int (1795)) (Prims.of_int (29))) - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (1796)) (Prims.of_int (6)) - (Prims.of_int (1796)) (Prims.of_int (29))) - (Obj.magic (unify_pr_with_true pr_l)) - (fun uu___3 -> - (fun uu___3 -> - Obj.magic (unify_pr_with_true pr_r)) - uu___3)) - | uu___1 -> - Obj.repr - (FStar_Tactics_Derived.fail - "unify_pr_with_true: ill-formed /\\"))) + (match tl with + | (pr_l, uu___1)::(pr_r, uu___2)::[] -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (1795)) (Prims.of_int (6)) + (Prims.of_int (1795)) (Prims.of_int (29))) + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (1796)) (Prims.of_int (6)) + (Prims.of_int (1796)) (Prims.of_int (29))) + (Obj.magic (unify_pr_with_true pr_l)) + (fun uu___3 -> + (fun uu___3 -> + Obj.magic (unify_pr_with_true pr_r)) uu___3)) + | uu___1 -> + Obj.magic + (FStar_Tactics_Derived.fail + "unify_pr_with_true: ill-formed /\\")) else Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (1799)) (Prims.of_int (10)) - (Prims.of_int (1799)) (Prims.of_int (30))) - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (1799)) (Prims.of_int (4)) - (Prims.of_int (1811)) (Prims.of_int (73))) - (Obj.magic - (FStar_Tactics_Derived.inspect_unascribe hd)) + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (1799)) (Prims.of_int (10)) + (Prims.of_int (1799)) (Prims.of_int (30))) + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (1799)) (Prims.of_int (4)) + (Prims.of_int (1811)) (Prims.of_int (73))) + (Obj.magic + (FStar_Tactics_Derived.inspect_unascribe hd)) + (fun uu___2 -> (fun uu___2 -> - (fun uu___2 -> - match uu___2 with - | FStar_Reflection_Data.Tv_Uvar - (uu___3, uu___4) -> - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (1801)) - (Prims.of_int (9)) - (Prims.of_int (1801)) - (Prims.of_int (27))) - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (1801)) - (Prims.of_int (6)) - (Prims.of_int (1805)) - (Prims.of_int (9))) - (Obj.magic - (FStar_Tactics_Derived.unify - pr - (FStar_Reflection_Builtins.pack_ln - (FStar_Reflection_Data.Tv_FVar - (FStar_Reflection_Builtins.pack_fv - ["Steel"; - "Effect"; - "Common"; - "true_p"]))))) + match uu___2 with + | FStar_Reflection_Data.Tv_Uvar (uu___3, uu___4) + -> + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (1801)) + (Prims.of_int (9)) + (Prims.of_int (1801)) + (Prims.of_int (27))) + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (1801)) + (Prims.of_int (6)) + (Prims.of_int (1805)) + (Prims.of_int (9))) + (Obj.magic + (FStar_Tactics_Derived.unify pr + (FStar_Reflection_Builtins.pack_ln + (FStar_Reflection_Data.Tv_FVar + (FStar_Reflection_Builtins.pack_fv + ["Steel"; + "Effect"; + "Common"; + "true_p"]))))) + (fun uu___5 -> (fun uu___5 -> if uu___5 then - FStar_Tactics_Effect.lift_div_tac - (fun uu___6 -> ()) + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___6 -> ()))) else - FStar_Tactics_Derived.fail - "unify_pr_with_true: could not unify SMT prop with True"))) - | uu___3 -> - Obj.magic - (Obj.repr - (if - (FStar_List_Tot_Base.length - (FStar_Reflection_Builtins.free_uvars - pr)) - = Prims.int_zero - then - FStar_Tactics_Effect.lift_div_tac - (fun uu___4 -> ()) - else - FStar_Tactics_Effect.raise + Obj.magic + (Obj.repr + (FStar_Tactics_Derived.fail + "unify_pr_with_true: could not unify SMT prop with True"))) + uu___5))) + | uu___3 -> + Obj.magic + (Obj.repr + (if + (FStar_List_Tot_Base.length + (FStar_Reflection_Builtins.free_uvars + pr)) + = Prims.int_zero + then + Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___4 -> ())) + else + Obj.repr + (FStar_Tactics_Effect.raise (Postpone - "unify_pr_with_true: some uvars are still there")))) - uu___2)))) uu___) + "unify_pr_with_true: some uvars are still there"))))) + uu___2))) uu___) let rec (set_abduction_variable_term : FStar_Reflection_Types.term -> (FStar_Reflection_Types.term, unit) FStar_Tactics_Effect.tac_repr) @@ -4391,112 +4451,27 @@ let rec (set_abduction_variable_term : | (hd, tl) -> if is_and hd then - Obj.magic - (Obj.repr - (match tl with - | (pr_l, FStar_Reflection_Data.Q_Explicit)::(pr_r, - FStar_Reflection_Data.Q_Explicit)::[] - -> - Obj.repr - (if - (FStar_List_Tot_Base.length - (FStar_Reflection_Builtins.free_uvars - pr_r)) - = Prims.int_zero - then - Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (1838)) - (Prims.of_int (18)) - (Prims.of_int (1838)) - (Prims.of_int (50))) - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (1839)) - (Prims.of_int (8)) - (Prims.of_int (1839)) - (Prims.of_int (53))) - (Obj.magic - (set_abduction_variable_term pr_l)) - (fun arg -> - FStar_Tactics_Effect.lift_div_tac - (fun uu___1 -> - FStar_Reflection_Derived.mk_app - (FStar_Reflection_Builtins.pack_ln - (FStar_Reflection_Data.Tv_FVar - (FStar_Reflection_Builtins.pack_fv - ["Steel"; - "Effect"; - "Common"; - "elim_and_l_squash"]))) - [(arg, - FStar_Reflection_Data.Q_Explicit)]))) - else - Obj.repr - (if - (FStar_List_Tot_Base.length - (FStar_Reflection_Builtins.free_uvars - pr_l)) - = Prims.int_zero - then - Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (1842)) - (Prims.of_int (18)) - (Prims.of_int (1842)) - (Prims.of_int (50))) - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (1843)) - (Prims.of_int (8)) - (Prims.of_int (1843)) - (Prims.of_int (53))) - (Obj.magic - (set_abduction_variable_term - pr_r)) - (fun arg -> - FStar_Tactics_Effect.lift_div_tac - (fun uu___2 -> - FStar_Reflection_Derived.mk_app - (FStar_Reflection_Builtins.pack_ln - (FStar_Reflection_Data.Tv_FVar - (FStar_Reflection_Builtins.pack_fv - ["Steel"; - "Effect"; - "Common"; - "elim_and_r_squash"]))) - [(arg, - FStar_Reflection_Data.Q_Explicit)]))) - else - Obj.repr - (FStar_Tactics_Effect.raise - (Postpone - "set_abduction_variable_term: there are still uvars on both sides of l_and")))) - | uu___1 -> - Obj.repr - (FStar_Tactics_Derived.fail - "set_abduction_variable: ill-formed /\\"))) - else - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (1831)) (Prims.of_int (6)) - (Prims.of_int (1831)) (Prims.of_int (8))) - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (1849)) (Prims.of_int (4)) - (Prims.of_int (1852)) (Prims.of_int (54))) - (Obj.magic (FStar_Tactics_Builtins.inspect hd)) - (fun uu___2 -> - match uu___2 with - | FStar_Reflection_Data.Tv_Uvar (uu___3, uu___4) - -> + (match tl with + | (pr_l, FStar_Reflection_Data.Q_Explicit)::(pr_r, + FStar_Reflection_Data.Q_Explicit)::[] + -> + if + (FStar_List_Tot_Base.length + (FStar_Reflection_Builtins.free_uvars pr_r)) + = Prims.int_zero + then + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (1838)) (Prims.of_int (18)) + (Prims.of_int (1838)) (Prims.of_int (50))) + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (1839)) (Prims.of_int (8)) + (Prims.of_int (1839)) (Prims.of_int (53))) + (Obj.magic (set_abduction_variable_term pr_l)) + (fun arg -> FStar_Tactics_Effect.lift_div_tac - (fun uu___5 -> + (fun uu___1 -> FStar_Reflection_Derived.mk_app (FStar_Reflection_Builtins.pack_ln (FStar_Reflection_Data.Tv_FVar @@ -4504,15 +4479,83 @@ let rec (set_abduction_variable_term : ["Steel"; "Effect"; "Common"; - "_return_squash"]))) - [((FStar_Reflection_Builtins.pack_ln - (FStar_Reflection_Data.Tv_Const - FStar_Reflection_Data.C_Unit)), - FStar_Reflection_Data.Q_Explicit)]) + "elim_and_l_squash"]))) + [(arg, + FStar_Reflection_Data.Q_Explicit)]))) + else + if + (FStar_List_Tot_Base.length + (FStar_Reflection_Builtins.free_uvars pr_l)) + = Prims.int_zero + then + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (1842)) (Prims.of_int (18)) + (Prims.of_int (1842)) (Prims.of_int (50))) + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (1843)) (Prims.of_int (8)) + (Prims.of_int (1843)) (Prims.of_int (53))) + (Obj.magic (set_abduction_variable_term pr_r)) + (fun arg -> + FStar_Tactics_Effect.lift_div_tac + (fun uu___2 -> + FStar_Reflection_Derived.mk_app + (FStar_Reflection_Builtins.pack_ln + (FStar_Reflection_Data.Tv_FVar + (FStar_Reflection_Builtins.pack_fv + ["Steel"; + "Effect"; + "Common"; + "elim_and_r_squash"]))) + [(arg, + FStar_Reflection_Data.Q_Explicit)]))) + else + Obj.magic + (FStar_Tactics_Effect.raise + (Postpone + "set_abduction_variable_term: there are still uvars on both sides of l_and")) + | uu___1 -> + Obj.magic + (FStar_Tactics_Derived.fail + "set_abduction_variable: ill-formed /\\")) + else + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (1831)) (Prims.of_int (6)) + (Prims.of_int (1831)) (Prims.of_int (8))) + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (1849)) (Prims.of_int (4)) + (Prims.of_int (1852)) (Prims.of_int (54))) + (Obj.magic (FStar_Tactics_Builtins.inspect hd)) + (fun uu___2 -> + (fun uu___2 -> + match uu___2 with + | FStar_Reflection_Data.Tv_Uvar (uu___3, uu___4) + -> + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___5 -> + FStar_Reflection_Derived.mk_app + (FStar_Reflection_Builtins.pack_ln + (FStar_Reflection_Data.Tv_FVar + (FStar_Reflection_Builtins.pack_fv + ["Steel"; + "Effect"; + "Common"; + "_return_squash"]))) + [((FStar_Reflection_Builtins.pack_ln + (FStar_Reflection_Data.Tv_Const + FStar_Reflection_Data.C_Unit)), + FStar_Reflection_Data.Q_Explicit)]))) | uu___3 -> - FStar_Tactics_Derived.fail - "set_abduction_variable: cannot unify")))) - uu___) + Obj.magic + (Obj.repr + (FStar_Tactics_Derived.fail + "set_abduction_variable: cannot unify"))) + uu___2))) uu___) let (set_abduction_variable : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = fun uu___ -> @@ -4538,68 +4581,62 @@ let (set_abduction_variable : match uu___1 with | FStar_Reflection_Data.Tv_Arrow (b, uu___2) -> Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (1858)) - (Prims.of_int (18)) - (Prims.of_int (1858)) - (Prims.of_int (34))) - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (1858)) - (Prims.of_int (4)) - (Prims.of_int (1861)) - (Prims.of_int (42))) - (FStar_Tactics_Effect.lift_div_tac - (fun uu___3 -> - FStar_Reflection_Builtins.inspect_binder - b)) + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (1858)) (Prims.of_int (18)) + (Prims.of_int (1858)) (Prims.of_int (34))) + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (1858)) (Prims.of_int (4)) + (Prims.of_int (1861)) (Prims.of_int (42))) + (FStar_Tactics_Effect.lift_div_tac (fun uu___3 -> - (fun uu___3 -> - match uu___3 with - | (bv, uu___4) -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (1859)) - (Prims.of_int (13)) - (Prims.of_int (1859)) - (Prims.of_int (26))) - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (1860)) - (Prims.of_int (4)) - (Prims.of_int (1861)) - (Prims.of_int (42))) - (FStar_Tactics_Effect.lift_div_tac - (fun uu___5 -> - FStar_Reflection_Builtins.inspect_bv - bv)) + FStar_Reflection_Builtins.inspect_binder + b)) + (fun uu___3 -> + (fun uu___3 -> + match uu___3 with + | (bv, uu___4) -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (1859)) + (Prims.of_int (13)) + (Prims.of_int (1859)) + (Prims.of_int (26))) + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (1860)) + (Prims.of_int (4)) + (Prims.of_int (1861)) + (Prims.of_int (42))) + (FStar_Tactics_Effect.lift_div_tac (fun uu___5 -> - (fun bv1 -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (1860)) - (Prims.of_int (13)) - (Prims.of_int (1860)) - (Prims.of_int (23))) - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (1861)) - (Prims.of_int (4)) - (Prims.of_int (1861)) - (Prims.of_int (42))) - (FStar_Tactics_Effect.lift_div_tac - (fun uu___5 - -> - bv1.FStar_Reflection_Data.bv_sort)) + FStar_Reflection_Builtins.inspect_bv + bv)) + (fun uu___5 -> + (fun bv1 -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (1860)) + (Prims.of_int (13)) + (Prims.of_int (1860)) + (Prims.of_int (23))) + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (1861)) + (Prims.of_int (4)) + (Prims.of_int (1861)) + (Prims.of_int (42))) + (FStar_Tactics_Effect.lift_div_tac (fun uu___5 -> - (fun pr -> - Obj.magic - (FStar_Tactics_Effect.tac_bind + bv1.FStar_Reflection_Data.bv_sort)) + (fun uu___5 -> + (fun pr -> + Obj.magic + (FStar_Tactics_Effect.tac_bind (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (1861)) @@ -4623,13 +4660,12 @@ let (set_abduction_variable : (FStar_Tactics_Derived.exact uu___5)) uu___5))) - uu___5))) - uu___5))) uu___3))) + uu___5))) + uu___5))) uu___3)) | uu___2 -> Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "Not an arrow goal"))) uu___1))) uu___1) + (FStar_Tactics_Derived.fail "Not an arrow goal")) + uu___1))) uu___1) let (canon_l_r : Prims.bool -> FStar_Reflection_Types.term -> @@ -4814,11 +4850,15 @@ let (canon_l_r : (flatten r2_raw) am2)) + (fun + uu___4 -> (fun res -> - FStar_Tactics_Effect.raise + Obj.magic + (FStar_Tactics_Effect.raise (Result res))) + uu___4)) (fun uu___3 -> (fun @@ -4829,21 +4869,24 @@ let (canon_l_r : FStar_Tactics_Common.TacticFailure m1 -> Obj.magic + (Obj.repr (FStar_Tactics_Derived.fail - m1) + m1)) | Result res -> Obj.magic + (Obj.repr (FStar_Tactics_Effect.lift_div_tac (fun uu___4 -> - res)) + res))) | uu___4 -> Obj.magic + (Obj.repr (FStar_Tactics_Derived.fail - "uncaught exception in equivalent_lists")) + "uncaught exception in equivalent_lists"))) uu___3))) (fun uu___3 -> @@ -6534,25 +6577,19 @@ let (canon_l_r : (fun uu___11 -> - (fun - uu___11 - -> match uu___11 with | FStar_Tactics_Common.TacticFailure msg -> - Obj.magic - (FStar_Tactics_Derived.fail + FStar_Tactics_Derived.fail (Prims.strcat "Cannot unify pr with true: " - msg)) + msg) | e -> - Obj.magic - (FStar_Tactics_Effect.raise - e)) - uu___11))) + FStar_Tactics_Effect.raise + e))) | l -> Obj.magic @@ -6795,40 +6832,34 @@ let (canon_monoid : (match rel_xy with | (rel_xy1, uu___2)::[] -> Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2090)) - (Prims.of_int (21)) - (Prims.of_int (2090)) - (Prims.of_int (43))) - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2090)) - (Prims.of_int (7)) - (Prims.of_int (2100)) - (Prims.of_int (8))) - (FStar_Tactics_Effect.lift_div_tac - (fun uu___3 - -> - FStar_Reflection_Derived_Lemmas.collect_app_ref - rel_xy1)) + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2090)) + (Prims.of_int (21)) + (Prims.of_int (2090)) + (Prims.of_int (43))) + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2090)) + (Prims.of_int (7)) + (Prims.of_int (2100)) + (Prims.of_int (8))) + (FStar_Tactics_Effect.lift_div_tac (fun uu___3 -> - (fun uu___3 - -> - match uu___3 - with - | - (rel, xy) + FStar_Reflection_Derived_Lemmas.collect_app_ref + rel_xy1)) + (fun uu___3 -> + (fun uu___3 -> + match uu___3 + with + | (rel, xy) -> if (FStar_List_Tot_Base.length xy) >= (Prims.of_int (2)) then - Obj.magic - (Obj.repr (match ((FStar_List_Tot_Base.index xy @@ -6847,7 +6878,7 @@ let (canon_monoid : (rhs, FStar_Reflection_Data.Q_Explicit)) -> - Obj.repr + Obj.magic (canon_l_r use_smt carrier_t @@ -6857,20 +6888,18 @@ let (canon_monoid : rhs) | uu___4 -> - Obj.repr + Obj.magic (FStar_Tactics_Derived.fail - "Goal should have been an application of a binary relation to 2 explicit arguments"))) + "Goal should have been an application of a binary relation to 2 explicit arguments")) else Obj.magic - (Obj.repr (FStar_Tactics_Derived.fail - "Goal should have been an application of a binary relation to n implicit and 2 explicit arguments"))) - uu___3))) + "Goal should have been an application of a binary relation to n implicit and 2 explicit arguments")) + uu___3)) | uu___2 -> Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "Goal should be squash applied to a binary relation")))) + (FStar_Tactics_Derived.fail + "Goal should be squash applied to a binary relation"))) uu___1))) uu___1))) uu___) let (canon' : Prims.bool -> @@ -7736,15 +7765,13 @@ let rec (extract_contexts : FStar_Pervasives_Native.None -> Obj.magic - (Obj.repr (FStar_Tactics_Derived.fail - "no context on the right either")) + "no context on the right either") | FStar_Pervasives_Native.Some g -> Obj.magic - (Obj.repr - (g ()))) + (g ())) uu___5))))))) uu___1))) uu___1)) | uu___1 -> @@ -7991,9 +8018,15 @@ let (open_existentials : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) "Common"; "solve_can_be_split_lookup"]))) e) then - FStar_Tactics_Derived.fail - "Tactic disabled: no available lemmas in context" - else FStar_Tactics_Effect.lift_div_tac (fun uu___2 -> ())) + Obj.magic + (Obj.repr + (FStar_Tactics_Derived.fail + "Tactic disabled: no available lemmas in context")) + else + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___2 -> ())))) (fun uu___1 -> (fun uu___1 -> Obj.magic @@ -8036,44 +8069,40 @@ let (open_existentials : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) FStar_Reflection_Data.Q_Explicit)::[]) -> Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2309)) - (Prims.of_int (22)) - (Prims.of_int (2309)) - (Prims.of_int (36))) - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2309)) - (Prims.of_int (9)) - (Prims.of_int (2325)) - (Prims.of_int (60))) - (FStar_Tactics_Effect.lift_div_tac - (fun uu___4 -> - FStar_Reflection_Derived.collect_app - t1)) + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2309)) + (Prims.of_int (22)) + (Prims.of_int (2309)) + (Prims.of_int (36))) + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2309)) + (Prims.of_int (9)) + (Prims.of_int (2325)) + (Prims.of_int (60))) + (FStar_Tactics_Effect.lift_div_tac (fun uu___4 -> - (fun uu___4 -> - match uu___4 - with - | (hd, tl) -> - if - FStar_Reflection_Derived.is_fvar + FStar_Reflection_Derived.collect_app + t1)) + (fun uu___4 -> + (fun uu___4 -> + match uu___4 with + | (hd, tl) -> + if + FStar_Reflection_Derived.is_fvar hd "Steel.Effect.Common.can_be_split" - then - Obj.magic - (Obj.repr - (match tl - with - | - uu___5:: + then + (match tl + with + | + uu___5:: (rhs, FStar_Reflection_Data.Q_Explicit)::[] -> - Obj.repr + Obj.magic (FStar_Tactics_Effect.tac_bind (Prims.mk_range "Steel.Effect.Common.fsti" @@ -8100,14 +8129,12 @@ let (open_existentials : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) FStar_Pervasives_Native.None -> Obj.magic - (Obj.repr (FStar_Tactics_Derived.fail - "open_existentials: no context found")) + "open_existentials: no context found") | FStar_Pervasives_Native.Some f -> Obj.magic - (Obj.repr (FStar_Tactics_Effect.tac_bind (Prims.mk_range "Steel.Effect.Common.fsti" @@ -8205,24 +8232,23 @@ let (open_existentials : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) uu___10))) uu___9))) uu___8))) - uu___7)))) + uu___7))) uu___6)) - | - uu___5 -> - Obj.repr - (FStar_Tactics_Derived.fail - "open_existentials: ill-formed can_be_split"))) - else + | + uu___5 -> Obj.magic - (Obj.repr (FStar_Tactics_Derived.fail - "open_existentials: not a can_be_split goal"))) - uu___4))) + "open_existentials: ill-formed can_be_split")) + else + Obj.magic + ( + FStar_Tactics_Derived.fail + "open_existentials: not a can_be_split goal")) + uu___4)) | uu___3 -> Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "open_existentials: not a squash goal"))) + (FStar_Tactics_Derived.fail + "open_existentials: not a squash goal")) uu___3))) uu___2))) uu___1))) uu___1) let (try_open_existentials : @@ -8541,104 +8567,85 @@ let (solve_can_be_split_dep : FStar_Reflection_Data.argv Prims.list -> (Prims.bool, unit) FStar_Tactics_Effect.tac_repr) = - fun uu___ -> - (fun args -> - match args with - | (p, uu___)::(t1, uu___1)::(t2, uu___2)::[] -> - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2388)) (Prims.of_int (17)) - (Prims.of_int (2388)) (Prims.of_int (36))) - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2389)) (Prims.of_int (6)) - (Prims.of_int (2420)) (Prims.of_int (18))) - (Obj.magic (slterm_nbr_uvars t1)) - (fun uu___3 -> - (fun lnbr -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2389)) (Prims.of_int (17)) - (Prims.of_int (2389)) (Prims.of_int (36))) - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2390)) (Prims.of_int (6)) - (Prims.of_int (2420)) (Prims.of_int (18))) - (Obj.magic (slterm_nbr_uvars t2)) - (fun uu___3 -> - (fun rnbr -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2391)) - (Prims.of_int (8)) - (Prims.of_int (2393)) - (Prims.of_int (18))) - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2390)) - (Prims.of_int (6)) - (Prims.of_int (2420)) - (Prims.of_int (18))) - (if (lnbr + rnbr) <= Prims.int_one - then - Obj.magic - (Obj.repr (unfold_guard ())) - else - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___4 -> false)))) - (fun uu___3 -> - (fun uu___3 -> - if uu___3 - then - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2396)) - (Prims.of_int (8)) - (Prims.of_int (2416)) - (Prims.of_int (36))) - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2418)) - (Prims.of_int (8)) - (Prims.of_int (2418)) - (Prims.of_int (12))) - (Obj.magic - (FStar_Tactics_Derived.focus - (fun uu___4 -> - FStar_Tactics_Effect.tac_bind - ( - Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2397)) - (Prims.of_int (23)) - (Prims.of_int (2397)) - (Prims.of_int (39))) - ( - Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2398)) - (Prims.of_int (10)) - (Prims.of_int (2416)) - (Prims.of_int (35))) - ( - Obj.magic - (FStar_Tactics_Logic.implies_intro - ())) - ( - fun - uu___5 -> - (fun - p_bind -> - Obj.magic - (FStar_Tactics_Effect.tac_bind + fun args -> + match args with + | (p, uu___)::(t1, uu___1)::(t2, uu___2)::[] -> + FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2388)) + (Prims.of_int (17)) (Prims.of_int (2388)) (Prims.of_int (36))) + (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2389)) + (Prims.of_int (6)) (Prims.of_int (2420)) (Prims.of_int (18))) + (Obj.magic (slterm_nbr_uvars t1)) + (fun uu___3 -> + (fun lnbr -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2389)) (Prims.of_int (17)) + (Prims.of_int (2389)) (Prims.of_int (36))) + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2390)) (Prims.of_int (6)) + (Prims.of_int (2420)) (Prims.of_int (18))) + (Obj.magic (slterm_nbr_uvars t2)) + (fun uu___3 -> + (fun rnbr -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2391)) (Prims.of_int (8)) + (Prims.of_int (2393)) (Prims.of_int (18))) + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2390)) (Prims.of_int (6)) + (Prims.of_int (2420)) (Prims.of_int (18))) + (if (lnbr + rnbr) <= Prims.int_one + then Obj.magic (Obj.repr (unfold_guard ())) + else + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___4 -> false)))) + (fun uu___3 -> + (fun uu___3 -> + if uu___3 + then + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2396)) + (Prims.of_int (8)) + (Prims.of_int (2416)) + (Prims.of_int (36))) + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2418)) + (Prims.of_int (8)) + (Prims.of_int (2418)) + (Prims.of_int (12))) + (Obj.magic + (FStar_Tactics_Derived.focus + (fun uu___4 -> + FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2397)) + (Prims.of_int (23)) + (Prims.of_int (2397)) + (Prims.of_int (39))) + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2398)) + (Prims.of_int (10)) + (Prims.of_int (2416)) + (Prims.of_int (35))) + (Obj.magic + (FStar_Tactics_Logic.implies_intro + ())) + (fun uu___5 -> + (fun p_bind -> + Obj.magic + (FStar_Tactics_Effect.tac_bind (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2398)) @@ -8733,13 +8740,17 @@ let (solve_can_be_split_dep : Prims.op_Negation b then - FStar_Tactics_Derived.fail - "could not unify SMT prop with True" + Obj.magic + (Obj.repr + (FStar_Tactics_Derived.fail + "could not unify SMT prop with True")) else - FStar_Tactics_Effect.lift_div_tac + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac (fun uu___9 -> - ())) + ())))) (fun uu___8 -> (fun @@ -8901,144 +8912,121 @@ let (solve_can_be_split_dep : uu___8)))) uu___6))) uu___5))) - uu___5)))) - (fun uu___4 -> - FStar_Tactics_Effect.lift_div_tac - (fun uu___5 -> - true)))) - else - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___5 -> false)))) - uu___3))) uu___3))) uu___3))) - | uu___ -> - Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail "ill-formed can_be_split_dep"))) - uu___ + uu___5)))) + (fun uu___4 -> + FStar_Tactics_Effect.lift_div_tac + (fun uu___5 -> true)))) + else + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___5 -> false)))) + uu___3))) uu___3))) uu___3) + | uu___ -> FStar_Tactics_Derived.fail "ill-formed can_be_split_dep" let (solve_can_be_split_forall : FStar_Reflection_Data.argv Prims.list -> (Prims.bool, unit) FStar_Tactics_Effect.tac_repr) = - fun uu___ -> - (fun args -> - match args with - | uu___::(t1, uu___1)::(t2, uu___2)::[] -> - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2432)) (Prims.of_int (17)) - (Prims.of_int (2432)) (Prims.of_int (36))) - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2433)) (Prims.of_int (6)) - (Prims.of_int (2458)) (Prims.of_int (18))) - (Obj.magic (slterm_nbr_uvars t1)) - (fun uu___3 -> - (fun lnbr -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2433)) (Prims.of_int (17)) - (Prims.of_int (2433)) (Prims.of_int (36))) - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2434)) (Prims.of_int (6)) - (Prims.of_int (2458)) (Prims.of_int (18))) - (Obj.magic (slterm_nbr_uvars t2)) - (fun uu___3 -> - (fun rnbr -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2435)) - (Prims.of_int (8)) - (Prims.of_int (2437)) - (Prims.of_int (18))) - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2434)) - (Prims.of_int (6)) - (Prims.of_int (2458)) - (Prims.of_int (18))) - (if (lnbr + rnbr) <= Prims.int_one - then - Obj.magic - (Obj.repr (unfold_guard ())) - else - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___4 -> false)))) - (fun uu___3 -> - (fun uu___3 -> - if uu___3 - then - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2440)) - (Prims.of_int (8)) - (Prims.of_int (2456)) - (Prims.of_int (46))) - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2457)) - (Prims.of_int (8)) - (Prims.of_int (2457)) - (Prims.of_int (12))) - (Obj.magic - (FStar_Tactics_Derived.focus - (fun uu___4 -> - FStar_Tactics_Effect.tac_bind - ( - Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2441)) - (Prims.of_int (10)) - (Prims.of_int (2441)) - (Prims.of_int (33))) - ( - Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2442)) - (Prims.of_int (10)) - (Prims.of_int (2456)) - (Prims.of_int (45))) - ( - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range + fun args -> + match args with + | uu___::(t1, uu___1)::(t2, uu___2)::[] -> + FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2432)) + (Prims.of_int (17)) (Prims.of_int (2432)) (Prims.of_int (36))) + (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2433)) + (Prims.of_int (6)) (Prims.of_int (2458)) (Prims.of_int (18))) + (Obj.magic (slterm_nbr_uvars t1)) + (fun uu___3 -> + (fun lnbr -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2433)) (Prims.of_int (17)) + (Prims.of_int (2433)) (Prims.of_int (36))) + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2434)) (Prims.of_int (6)) + (Prims.of_int (2458)) (Prims.of_int (18))) + (Obj.magic (slterm_nbr_uvars t2)) + (fun uu___3 -> + (fun rnbr -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2435)) (Prims.of_int (8)) + (Prims.of_int (2437)) (Prims.of_int (18))) + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2434)) (Prims.of_int (6)) + (Prims.of_int (2458)) (Prims.of_int (18))) + (if (lnbr + rnbr) <= Prims.int_one + then Obj.magic (Obj.repr (unfold_guard ())) + else + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___4 -> false)))) + (fun uu___3 -> + (fun uu___3 -> + if uu___3 + then + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2440)) + (Prims.of_int (8)) + (Prims.of_int (2456)) + (Prims.of_int (46))) + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2457)) + (Prims.of_int (8)) + (Prims.of_int (2457)) + (Prims.of_int (12))) + (Obj.magic + (FStar_Tactics_Derived.focus + (fun uu___4 -> + FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2441)) + (Prims.of_int (10)) + (Prims.of_int (2441)) + (Prims.of_int (33))) + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2442)) + (Prims.of_int (10)) + (Prims.of_int (2456)) + (Prims.of_int (45))) + (Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2441)) (Prims.of_int (17)) (Prims.of_int (2441)) (Prims.of_int (33))) - (Prims.mk_range + (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2441)) (Prims.of_int (10)) (Prims.of_int (2441)) (Prims.of_int (33))) - (Obj.magic - (FStar_Tactics_Logic.forall_intro + (Obj.magic + ( + FStar_Tactics_Logic.forall_intro ())) - (fun - uu___5 -> + (fun uu___5 + -> FStar_Tactics_Effect.lift_div_tac (fun uu___6 -> ())))) - ( - fun - uu___5 -> - (fun - uu___5 -> - Obj.magic - (FStar_Tactics_Effect.tac_bind + (fun uu___5 -> + (fun uu___5 -> + Obj.magic + (FStar_Tactics_Effect.tac_bind (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2442)) @@ -9237,23 +9225,19 @@ let (solve_can_be_split_forall : uu___9)))) uu___7))) uu___6))) - uu___5)))) - (fun uu___4 -> - FStar_Tactics_Effect.lift_div_tac - (fun uu___5 -> - true)))) - else - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___5 -> false)))) - uu___3))) uu___3))) uu___3))) - | uu___ -> - Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "Ill-formed can_be_split_forall, should not happen"))) - uu___ + uu___5)))) + (fun uu___4 -> + FStar_Tactics_Effect.lift_div_tac + (fun uu___5 -> true)))) + else + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___5 -> false)))) + uu___3))) uu___3))) uu___3) + | uu___ -> + FStar_Tactics_Derived.fail + "Ill-formed can_be_split_forall, should not happen" let (extract_cbs_forall_dep_contexts : FStar_Reflection_Types.term -> ((unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) @@ -9314,9 +9298,15 @@ let (open_existentials_forall_dep : "solve_can_be_split_forall_dep_lookup"]))) e) then - FStar_Tactics_Derived.fail - "Tactic disabled: no available lemmas in context" - else FStar_Tactics_Effect.lift_div_tac (fun uu___2 -> ())) + Obj.magic + (Obj.repr + (FStar_Tactics_Derived.fail + "Tactic disabled: no available lemmas in context")) + else + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___2 -> ())))) (fun uu___1 -> (fun uu___1 -> Obj.magic @@ -9364,44 +9354,40 @@ let (open_existentials_forall_dep : FStar_Reflection_Data.Q_Explicit)::[]) -> Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2491)) - (Prims.of_int (17)) - (Prims.of_int (2491)) - (Prims.of_int (31))) - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2491)) - (Prims.of_int (4)) - (Prims.of_int (2513)) - (Prims.of_int (78))) - (FStar_Tactics_Effect.lift_div_tac - (fun uu___4 -> - FStar_Reflection_Derived.collect_app - t1)) + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2491)) + (Prims.of_int (17)) + (Prims.of_int (2491)) + (Prims.of_int (31))) + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2491)) + (Prims.of_int (4)) + (Prims.of_int (2513)) + (Prims.of_int (78))) + (FStar_Tactics_Effect.lift_div_tac (fun uu___4 -> - (fun uu___4 -> - match uu___4 - with - | (hd, tl) -> - if - FStar_Reflection_Derived.is_fvar + FStar_Reflection_Derived.collect_app + t1)) + (fun uu___4 -> + (fun uu___4 -> + match uu___4 with + | (hd, tl) -> + if + FStar_Reflection_Derived.is_fvar hd "Steel.Effect.Common.can_be_split_forall_dep" - then - Obj.magic - (Obj.repr - (match tl - with - | - uu___5::uu___6:: + then + (match tl + with + | + uu___5::uu___6:: (rhs, FStar_Reflection_Data.Q_Explicit)::[] -> - Obj.repr + Obj.magic (FStar_Tactics_Effect.tac_bind (Prims.mk_range "Steel.Effect.Common.fsti" @@ -9429,7 +9415,6 @@ let (open_existentials_forall_dep : (uu___8, body) -> Obj.magic - (Obj.repr (FStar_Tactics_Effect.tac_bind (Prims.mk_range "Steel.Effect.Common.fsti" @@ -9456,14 +9441,12 @@ let (open_existentials_forall_dep : FStar_Pervasives_Native.None -> Obj.magic - (Obj.repr (FStar_Tactics_Derived.fail - "open_existentials_forall_dep: no candidate")) + "open_existentials_forall_dep: no candidate") | FStar_Pervasives_Native.Some f -> Obj.magic - (Obj.repr (FStar_Tactics_Effect.tac_bind (Prims.mk_range "Steel.Effect.Common.fsti" @@ -9649,22 +9632,21 @@ let (open_existentials_forall_dep : uu___13))) uu___12))) uu___11))) - uu___10)))) - uu___9))) + uu___10))) + uu___9)) | uu___8 -> Obj.magic - (Obj.repr (FStar_Tactics_Derived.fail - "open_existentials_forall_dep : not an abstraction"))) + "open_existentials_forall_dep : not an abstraction")) uu___7)) - | - (uu___5, + | + (uu___5, FStar_Reflection_Data.Q_Implicit)::uu___6::uu___7:: (rhs, FStar_Reflection_Data.Q_Explicit)::[] -> - Obj.repr + Obj.magic (FStar_Tactics_Effect.tac_bind (Prims.mk_range "Steel.Effect.Common.fsti" @@ -9692,7 +9674,6 @@ let (open_existentials_forall_dep : (uu___9, body) -> Obj.magic - (Obj.repr (FStar_Tactics_Effect.tac_bind (Prims.mk_range "Steel.Effect.Common.fsti" @@ -9721,14 +9702,12 @@ let (open_existentials_forall_dep : FStar_Pervasives_Native.None -> Obj.magic - (Obj.repr (FStar_Tactics_Derived.fail - "open_existentials_forall_dep: no candidate")) + "open_existentials_forall_dep: no candidate") | FStar_Pervasives_Native.Some f -> Obj.magic - (Obj.repr (FStar_Tactics_Effect.tac_bind (Prims.mk_range "Steel.Effect.Common.fsti" @@ -9914,31 +9893,29 @@ let (open_existentials_forall_dep : uu___14))) uu___13))) uu___12))) - uu___11)))) - uu___10))) + uu___11))) + uu___10)) | uu___9 -> Obj.magic - (Obj.repr (FStar_Tactics_Derived.fail - "open_existentials_forall_dep : not an abstraction"))) + "open_existentials_forall_dep : not an abstraction")) uu___8)) - | - uu___5 -> - Obj.repr - (FStar_Tactics_Derived.fail - "open_existentials_forall_dep : wrong number of arguments to can_be_split_forall_dep"))) - else + | + uu___5 -> Obj.magic - (Obj.repr (FStar_Tactics_Derived.fail - "open_existentials_forall_dep : not a can_be_split_forall_dep goal"))) - uu___4))) + "open_existentials_forall_dep : wrong number of arguments to can_be_split_forall_dep")) + else + Obj.magic + ( + FStar_Tactics_Derived.fail + "open_existentials_forall_dep : not a can_be_split_forall_dep goal")) + uu___4)) | uu___3 -> Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "open_existentials_forall_dep : not a squash/auto_squash goal"))) + (FStar_Tactics_Derived.fail + "open_existentials_forall_dep : not a squash/auto_squash goal")) uu___3))) uu___2))) uu___1))) uu___1) let (try_open_existentials_forall_dep : @@ -9969,100 +9946,90 @@ let rec (solve_can_be_split_forall_dep : FStar_Reflection_Data.argv Prims.list -> (Prims.bool, unit) FStar_Tactics_Effect.tac_repr) = - fun uu___ -> - (fun args -> - match args with - | uu___::(pr, uu___1)::(t1, uu___2)::(t2, uu___3)::[] -> - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2530)) (Prims.of_int (17)) - (Prims.of_int (2530)) (Prims.of_int (36))) - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2531)) (Prims.of_int (6)) - (Prims.of_int (2575)) (Prims.of_int (18))) - (Obj.magic (slterm_nbr_uvars t1)) - (fun uu___4 -> - (fun lnbr -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2531)) (Prims.of_int (17)) - (Prims.of_int (2531)) (Prims.of_int (36))) - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2532)) (Prims.of_int (6)) - (Prims.of_int (2575)) (Prims.of_int (18))) - (Obj.magic (slterm_nbr_uvars t2)) - (fun uu___4 -> - (fun rnbr -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2533)) - (Prims.of_int (8)) - (Prims.of_int (2535)) - (Prims.of_int (18))) - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2532)) - (Prims.of_int (6)) - (Prims.of_int (2575)) - (Prims.of_int (18))) - (if (lnbr + rnbr) <= Prims.int_one - then - Obj.magic - (Obj.repr (unfold_guard ())) - else - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___5 -> false)))) - (fun uu___4 -> - (fun uu___4 -> - if uu___4 - then - Obj.magic - (Obj.repr - (FStar_Tactics_Derived.try_with - (fun uu___5 -> - match () with - | () -> - FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2539)) - (Prims.of_int (9)) - (Prims.of_int (2563)) - (Prims.of_int (37))) - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2565)) - (Prims.of_int (9)) - (Prims.of_int (2565)) - (Prims.of_int (13))) - (Obj.magic - (FStar_Tactics_Derived.focus - (fun - uu___6 -> - FStar_Tactics_Effect.tac_bind - (Prims.mk_range + fun args -> + match args with + | uu___::(pr, uu___1)::(t1, uu___2)::(t2, uu___3)::[] -> + FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2530)) + (Prims.of_int (17)) (Prims.of_int (2530)) (Prims.of_int (36))) + (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2531)) + (Prims.of_int (6)) (Prims.of_int (2575)) (Prims.of_int (18))) + (Obj.magic (slterm_nbr_uvars t1)) + (fun uu___4 -> + (fun lnbr -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2531)) (Prims.of_int (17)) + (Prims.of_int (2531)) (Prims.of_int (36))) + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2532)) (Prims.of_int (6)) + (Prims.of_int (2575)) (Prims.of_int (18))) + (Obj.magic (slterm_nbr_uvars t2)) + (fun uu___4 -> + (fun rnbr -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2533)) (Prims.of_int (8)) + (Prims.of_int (2535)) (Prims.of_int (18))) + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2532)) (Prims.of_int (6)) + (Prims.of_int (2575)) (Prims.of_int (18))) + (if (lnbr + rnbr) <= Prims.int_one + then Obj.magic (Obj.repr (unfold_guard ())) + else + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___5 -> false)))) + (fun uu___4 -> + (fun uu___4 -> + if uu___4 + then + Obj.magic + (Obj.repr + (FStar_Tactics_Derived.try_with + (fun uu___5 -> + match () with + | () -> + FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2539)) + (Prims.of_int (9)) + (Prims.of_int (2563)) + (Prims.of_int (37))) + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2565)) + (Prims.of_int (9)) + (Prims.of_int (2565)) + (Prims.of_int (13))) + (Obj.magic + (FStar_Tactics_Derived.focus + (fun uu___6 -> + FStar_Tactics_Effect.tac_bind + ( + Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2540)) (Prims.of_int (10)) (Prims.of_int (2540)) (Prims.of_int (17))) - (Prims.mk_range + ( + Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2541)) (Prims.of_int (10)) (Prims.of_int (2563)) (Prims.of_int (36))) - (Obj.magic + ( + Obj.magic (FStar_Tactics_Builtins.norm [])) - (fun + ( + fun uu___7 -> (fun uu___7 -> @@ -10323,13 +10290,17 @@ let rec (solve_can_be_split_forall_dep : Prims.op_Negation b then - FStar_Tactics_Derived.fail - "could not unify SMT prop with True" + Obj.magic + (Obj.repr + (FStar_Tactics_Derived.fail + "could not unify SMT prop with True")) else - FStar_Tactics_Effect.lift_div_tac + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac (fun uu___12 - -> ())) + -> ())))) (fun uu___11 -> @@ -10506,186 +10477,158 @@ let rec (solve_can_be_split_forall_dep : uu___8))) uu___8))) uu___7)))) + (fun uu___6 -> + FStar_Tactics_Effect.lift_div_tac + (fun uu___7 -> + true))) + (fun uu___5 -> + (fun uu___5 -> + match uu___5 with + | Postpone msg -> + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac (fun uu___6 - -> - FStar_Tactics_Effect.lift_div_tac - (fun - uu___7 -> - true))) - (fun uu___5 -> - (fun uu___5 -> - match uu___5 - with - | Postpone msg - -> - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun - uu___6 -> - false))) - | FStar_Tactics_Common.TacticFailure - msg -> - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range + -> false))) + | FStar_Tactics_Common.TacticFailure + msg -> + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2570)) (Prims.of_int (22)) (Prims.of_int (2570)) (Prims.of_int (57))) - (Prims.mk_range + (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2571)) (Prims.of_int (9)) (Prims.of_int (2573)) (Prims.of_int (22))) - (Obj.magic + (Obj.magic (try_open_existentials_forall_dep ())) - (fun - uu___6 -> + (fun uu___6 + -> (fun opened -> if opened then Obj.magic - (Obj.repr (solve_can_be_split_forall_dep - args)) + args) else Obj.magic - (Obj.repr (FStar_Tactics_Derived.fail - msg))) + msg)) uu___6))) - | uu___6 -> - Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "Unexpected exception in framing tactic"))) - uu___5))) - else - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___6 -> false)))) - uu___4))) uu___4))) uu___4))) - | uu___ -> - Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "Ill-formed can_be_split_forall_dep, should not happen"))) - uu___ + | uu___6 -> + Obj.magic + (Obj.repr + (FStar_Tactics_Derived.fail + "Unexpected exception in framing tactic"))) + uu___5))) + else + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___6 -> false)))) + uu___4))) uu___4))) uu___4) + | uu___ -> + FStar_Tactics_Derived.fail + "Ill-formed can_be_split_forall_dep, should not happen" let (solve_equiv_forall : FStar_Reflection_Data.argv Prims.list -> (Prims.bool, unit) FStar_Tactics_Effect.tac_repr) = - fun uu___ -> - (fun args -> - match args with - | uu___::(t1, uu___1)::(t2, uu___2)::[] -> - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2583)) (Prims.of_int (17)) - (Prims.of_int (2583)) (Prims.of_int (36))) - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2584)) (Prims.of_int (6)) - (Prims.of_int (2612)) (Prims.of_int (18))) - (Obj.magic (slterm_nbr_uvars t1)) - (fun uu___3 -> - (fun lnbr -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2584)) (Prims.of_int (17)) - (Prims.of_int (2584)) (Prims.of_int (36))) - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2585)) (Prims.of_int (6)) - (Prims.of_int (2612)) (Prims.of_int (18))) - (Obj.magic (slterm_nbr_uvars t2)) - (fun uu___3 -> - (fun rnbr -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2586)) - (Prims.of_int (8)) - (Prims.of_int (2588)) - (Prims.of_int (18))) - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2585)) - (Prims.of_int (6)) - (Prims.of_int (2612)) - (Prims.of_int (18))) - (if (lnbr + rnbr) <= Prims.int_one - then - Obj.magic - (Obj.repr (unfold_guard ())) - else - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___4 -> false)))) - (fun uu___3 -> - (fun uu___3 -> - if uu___3 - then - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2591)) - (Prims.of_int (8)) - (Prims.of_int (2610)) - (Prims.of_int (62))) - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2611)) - (Prims.of_int (8)) - (Prims.of_int (2611)) - (Prims.of_int (12))) - (Obj.magic - (FStar_Tactics_Derived.focus - (fun uu___4 -> - FStar_Tactics_Effect.tac_bind - ( - Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2591)) - (Prims.of_int (24)) - (Prims.of_int (2591)) - (Prims.of_int (56))) - ( - Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2592)) - (Prims.of_int (22)) - (Prims.of_int (2610)) - (Prims.of_int (61))) + fun args -> + match args with + | uu___::(t1, uu___1)::(t2, uu___2)::[] -> + FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2583)) + (Prims.of_int (17)) (Prims.of_int (2583)) (Prims.of_int (36))) + (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2584)) + (Prims.of_int (6)) (Prims.of_int (2612)) (Prims.of_int (18))) + (Obj.magic (slterm_nbr_uvars t1)) + (fun uu___3 -> + (fun lnbr -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2584)) (Prims.of_int (17)) + (Prims.of_int (2584)) (Prims.of_int (36))) + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2585)) (Prims.of_int (6)) + (Prims.of_int (2612)) (Prims.of_int (18))) + (Obj.magic (slterm_nbr_uvars t2)) + (fun uu___3 -> + (fun rnbr -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2586)) (Prims.of_int (8)) + (Prims.of_int (2588)) (Prims.of_int (18))) + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2585)) (Prims.of_int (6)) + (Prims.of_int (2612)) (Prims.of_int (18))) + (if (lnbr + rnbr) <= Prims.int_one + then Obj.magic (Obj.repr (unfold_guard ())) + else + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___4 -> false)))) + (fun uu___3 -> + (fun uu___3 -> + if uu___3 + then + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2591)) + (Prims.of_int (8)) + (Prims.of_int (2610)) + (Prims.of_int (62))) + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2611)) + (Prims.of_int (8)) + (Prims.of_int (2611)) + (Prims.of_int (12))) + (Obj.magic + (FStar_Tactics_Derived.focus + (fun uu___4 -> + FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2591)) + (Prims.of_int (24)) + (Prims.of_int (2591)) + (Prims.of_int (56))) + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2592)) + (Prims.of_int (22)) + (Prims.of_int (2610)) + (Prims.of_int (61))) + (Obj.magic + (FStar_Tactics_Derived.apply_lemma + (FStar_Reflection_Builtins.pack_ln ( - Obj.magic - (FStar_Tactics_Derived.apply_lemma - (FStar_Reflection_Builtins.pack_ln - (FStar_Reflection_Data.Tv_FVar + FStar_Reflection_Data.Tv_FVar (FStar_Reflection_Builtins.pack_fv ["Steel"; "Effect"; "Common"; "equiv_forall_elim"]))))) - ( - fun - uu___5 -> - (fun - uu___5 -> - Obj.magic - (FStar_Tactics_Effect.tac_bind + (fun uu___5 -> + (fun uu___5 -> + Obj.magic + (FStar_Tactics_Effect.tac_bind (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2592)) @@ -10937,134 +10880,114 @@ let (solve_equiv_forall : uu___9))) uu___8)))) uu___6))) - uu___5)))) - (fun uu___4 -> - FStar_Tactics_Effect.lift_div_tac - (fun uu___5 -> - true)))) - else - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___5 -> false)))) - uu___3))) uu___3))) uu___3))) - | uu___ -> - Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "Ill-formed equiv_forall, should not happen"))) uu___ + uu___5)))) + (fun uu___4 -> + FStar_Tactics_Effect.lift_div_tac + (fun uu___5 -> true)))) + else + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___5 -> false)))) + uu___3))) uu___3))) uu___3) + | uu___ -> + FStar_Tactics_Derived.fail + "Ill-formed equiv_forall, should not happen" let (solve_equiv : FStar_Reflection_Data.argv Prims.list -> (Prims.bool, unit) FStar_Tactics_Effect.tac_repr) = - fun uu___ -> - (fun args -> - match args with - | (t1, uu___)::(t2, uu___1)::[] -> - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2620)) (Prims.of_int (17)) - (Prims.of_int (2620)) (Prims.of_int (36))) - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2621)) (Prims.of_int (6)) - (Prims.of_int (2645)) (Prims.of_int (18))) - (Obj.magic (slterm_nbr_uvars t1)) - (fun uu___2 -> - (fun lnbr -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2621)) (Prims.of_int (17)) - (Prims.of_int (2621)) (Prims.of_int (36))) - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2622)) (Prims.of_int (6)) - (Prims.of_int (2645)) (Prims.of_int (18))) - (Obj.magic (slterm_nbr_uvars t2)) - (fun uu___2 -> - (fun rnbr -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2623)) - (Prims.of_int (8)) - (Prims.of_int (2625)) - (Prims.of_int (18))) - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2622)) - (Prims.of_int (6)) - (Prims.of_int (2645)) - (Prims.of_int (18))) - (if (lnbr + rnbr) <= Prims.int_one - then - Obj.magic - (Obj.repr (unfold_guard ())) - else - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___3 -> false)))) - (fun uu___2 -> - (fun uu___2 -> - if uu___2 - then - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2628)) - (Prims.of_int (8)) - (Prims.of_int (2642)) - (Prims.of_int (48))) - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2643)) - (Prims.of_int (8)) - (Prims.of_int (2643)) - (Prims.of_int (12))) - (Obj.magic - (FStar_Tactics_Derived.focus - (fun uu___3 -> - FStar_Tactics_Derived.or_else - ( - fun - uu___4 -> - FStar_Tactics_Derived.apply_lemma - (FStar_Reflection_Builtins.pack_ln - (FStar_Reflection_Data.Tv_FVar + fun args -> + match args with + | (t1, uu___)::(t2, uu___1)::[] -> + FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2620)) + (Prims.of_int (17)) (Prims.of_int (2620)) (Prims.of_int (36))) + (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2621)) + (Prims.of_int (6)) (Prims.of_int (2645)) (Prims.of_int (18))) + (Obj.magic (slterm_nbr_uvars t1)) + (fun uu___2 -> + (fun lnbr -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2621)) (Prims.of_int (17)) + (Prims.of_int (2621)) (Prims.of_int (36))) + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2622)) (Prims.of_int (6)) + (Prims.of_int (2645)) (Prims.of_int (18))) + (Obj.magic (slterm_nbr_uvars t2)) + (fun uu___2 -> + (fun rnbr -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2623)) (Prims.of_int (8)) + (Prims.of_int (2625)) (Prims.of_int (18))) + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2622)) (Prims.of_int (6)) + (Prims.of_int (2645)) (Prims.of_int (18))) + (if (lnbr + rnbr) <= Prims.int_one + then Obj.magic (Obj.repr (unfold_guard ())) + else + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___3 -> false)))) + (fun uu___2 -> + (fun uu___2 -> + if uu___2 + then + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2628)) + (Prims.of_int (8)) + (Prims.of_int (2642)) + (Prims.of_int (48))) + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2643)) + (Prims.of_int (8)) + (Prims.of_int (2643)) + (Prims.of_int (12))) + (Obj.magic + (FStar_Tactics_Derived.focus + (fun uu___3 -> + FStar_Tactics_Derived.or_else + (fun uu___4 -> + FStar_Tactics_Derived.apply_lemma + (FStar_Reflection_Builtins.pack_ln + (FStar_Reflection_Data.Tv_FVar (FStar_Reflection_Builtins.pack_fv ["Steel"; "Effect"; "Common"; "equiv_refl"])))) - ( - fun - uu___4 -> - FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2632)) - (Prims.of_int (14)) - (Prims.of_int (2632)) - (Prims.of_int (68))) - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2633)) - (Prims.of_int (14)) - (Prims.of_int (2642)) - (Prims.of_int (46))) - (if - (lnbr <> + (fun uu___4 -> + FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2632)) + (Prims.of_int (14)) + (Prims.of_int (2632)) + (Prims.of_int (68))) + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2633)) + (Prims.of_int (14)) + (Prims.of_int (2642)) + (Prims.of_int (46))) + (if + (lnbr <> Prims.int_zero) && (rnbr = Prims.int_zero) - then - Obj.magic + then + Obj.magic (Obj.repr (FStar_Tactics_Derived.apply_lemma (FStar_Reflection_Builtins.pack_ln @@ -11074,16 +10997,16 @@ let (solve_equiv : "Effect"; "Common"; "equiv_sym"]))))) - else - Obj.magic + else + Obj.magic (Obj.repr (FStar_Tactics_Effect.lift_div_tac (fun uu___6 -> ())))) - (fun - uu___5 -> - (fun + (fun uu___5 + -> + (fun uu___5 -> Obj.magic (FStar_Tactics_Effect.tac_bind @@ -11146,150 +11069,127 @@ let (solve_equiv : "FStar.Pervasives.Native.snd"]; FStar_Pervasives.delta_attr ["Steel.Effect.Common.__reduce__"]; - FStar_Pervasives.primops; - FStar_Pervasives.iota; - FStar_Pervasives.zeta])) - (fun - uu___7 -> - (fun - uu___7 -> - Obj.magic - (canon' - false - (FStar_Reflection_Builtins.pack_ln - (FStar_Reflection_Data.Tv_FVar - (FStar_Reflection_Builtins.pack_fv - ["Steel"; - "Effect"; - "Common"; - "true_p"]))) - (FStar_Reflection_Builtins.pack_ln - (FStar_Reflection_Data.Tv_FVar - (FStar_Reflection_Builtins.pack_fv - ["Steel"; - "Effect"; - "Common"; - "true_p"]))))) - uu___7))) - uu___6))) - uu___5))))) - (fun uu___3 -> - FStar_Tactics_Effect.lift_div_tac - (fun uu___4 -> - true)))) - else - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___4 -> false)))) - uu___2))) uu___2))) uu___2))) - | uu___ -> - Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "Ill-formed equiv, should not happen"))) uu___ -let (solve_can_be_split_post : - FStar_Reflection_Data.argv Prims.list -> - (Prims.bool, unit) FStar_Tactics_Effect.tac_repr) - = - fun uu___ -> - (fun args -> - match args with - | uu___::uu___1::(t1, uu___2)::(t2, uu___3)::[] -> - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2653)) (Prims.of_int (17)) - (Prims.of_int (2653)) (Prims.of_int (36))) - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2654)) (Prims.of_int (6)) - (Prims.of_int (2685)) (Prims.of_int (18))) - (Obj.magic (slterm_nbr_uvars t1)) - (fun uu___4 -> - (fun lnbr -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2654)) (Prims.of_int (17)) - (Prims.of_int (2654)) (Prims.of_int (36))) - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2655)) (Prims.of_int (6)) - (Prims.of_int (2685)) (Prims.of_int (18))) - (Obj.magic (slterm_nbr_uvars t2)) - (fun uu___4 -> - (fun rnbr -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2656)) - (Prims.of_int (8)) - (Prims.of_int (2658)) - (Prims.of_int (18))) - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2655)) - (Prims.of_int (6)) - (Prims.of_int (2685)) - (Prims.of_int (18))) - (if (lnbr + rnbr) <= Prims.int_one - then - Obj.magic - (Obj.repr (unfold_guard ())) - else - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___5 -> false)))) - (fun uu___4 -> - (fun uu___4 -> - if uu___4 - then - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2661)) - (Prims.of_int (8)) - (Prims.of_int (2683)) - (Prims.of_int (62))) - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2684)) - (Prims.of_int (8)) - (Prims.of_int (2684)) - (Prims.of_int (12))) - (Obj.magic - (FStar_Tactics_Derived.focus - (fun uu___5 -> - FStar_Tactics_Effect.tac_bind - ( - Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2661)) - (Prims.of_int (24)) - (Prims.of_int (2661)) - (Prims.of_int (30))) - ( - Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2662)) - (Prims.of_int (22)) - (Prims.of_int (2683)) - (Prims.of_int (61))) - ( - Obj.magic - (FStar_Tactics_Builtins.norm - [])) - ( - fun - uu___6 -> + FStar_Pervasives.primops; + FStar_Pervasives.iota; + FStar_Pervasives.zeta])) (fun - uu___6 -> + uu___7 -> + (fun + uu___7 -> Obj.magic - (FStar_Tactics_Effect.tac_bind + (canon' + false + (FStar_Reflection_Builtins.pack_ln + (FStar_Reflection_Data.Tv_FVar + (FStar_Reflection_Builtins.pack_fv + ["Steel"; + "Effect"; + "Common"; + "true_p"]))) + (FStar_Reflection_Builtins.pack_ln + (FStar_Reflection_Data.Tv_FVar + (FStar_Reflection_Builtins.pack_fv + ["Steel"; + "Effect"; + "Common"; + "true_p"]))))) + uu___7))) + uu___6))) + uu___5))))) + (fun uu___3 -> + FStar_Tactics_Effect.lift_div_tac + (fun uu___4 -> true)))) + else + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___4 -> false)))) + uu___2))) uu___2))) uu___2) + | uu___ -> + FStar_Tactics_Derived.fail "Ill-formed equiv, should not happen" +let (solve_can_be_split_post : + FStar_Reflection_Data.argv Prims.list -> + (Prims.bool, unit) FStar_Tactics_Effect.tac_repr) + = + fun args -> + match args with + | uu___::uu___1::(t1, uu___2)::(t2, uu___3)::[] -> + FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2653)) + (Prims.of_int (17)) (Prims.of_int (2653)) (Prims.of_int (36))) + (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2654)) + (Prims.of_int (6)) (Prims.of_int (2685)) (Prims.of_int (18))) + (Obj.magic (slterm_nbr_uvars t1)) + (fun uu___4 -> + (fun lnbr -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2654)) (Prims.of_int (17)) + (Prims.of_int (2654)) (Prims.of_int (36))) + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2655)) (Prims.of_int (6)) + (Prims.of_int (2685)) (Prims.of_int (18))) + (Obj.magic (slterm_nbr_uvars t2)) + (fun uu___4 -> + (fun rnbr -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2656)) (Prims.of_int (8)) + (Prims.of_int (2658)) (Prims.of_int (18))) + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2655)) (Prims.of_int (6)) + (Prims.of_int (2685)) (Prims.of_int (18))) + (if (lnbr + rnbr) <= Prims.int_one + then Obj.magic (Obj.repr (unfold_guard ())) + else + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___5 -> false)))) + (fun uu___4 -> + (fun uu___4 -> + if uu___4 + then + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2661)) + (Prims.of_int (8)) + (Prims.of_int (2683)) + (Prims.of_int (62))) + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2684)) + (Prims.of_int (8)) + (Prims.of_int (2684)) + (Prims.of_int (12))) + (Obj.magic + (FStar_Tactics_Derived.focus + (fun uu___5 -> + FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2661)) + (Prims.of_int (24)) + (Prims.of_int (2661)) + (Prims.of_int (30))) + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2662)) + (Prims.of_int (22)) + (Prims.of_int (2683)) + (Prims.of_int (61))) + (Obj.magic + (FStar_Tactics_Builtins.norm + [])) + (fun uu___6 -> + (fun uu___6 -> + Obj.magic + (FStar_Tactics_Effect.tac_bind (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2662)) @@ -11638,22 +11538,17 @@ let (solve_can_be_split_post : uu___8))) uu___7))) uu___7))) - uu___6)))) - (fun uu___5 -> - FStar_Tactics_Effect.lift_div_tac - (fun uu___6 -> - true)))) - else - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___6 -> false)))) - uu___4))) uu___4))) uu___4))) - | uu___ -> - Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail "ill-formed can_be_split_post"))) - uu___ + uu___6)))) + (fun uu___5 -> + FStar_Tactics_Effect.lift_div_tac + (fun uu___6 -> true)))) + else + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___6 -> false)))) + uu___4))) uu___4))) uu___4) + | uu___ -> FStar_Tactics_Derived.fail "ill-formed can_be_split_post" let (is_return_eq : FStar_Reflection_Types.term -> FStar_Reflection_Types.term -> @@ -12029,118 +11924,116 @@ let (goal_to_equiv : match f with | FStar_Reflection_Formula.App (hd0, t1) -> Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2746)) - (Prims.of_int (6)) - (Prims.of_int (2747)) - (Prims.of_int (70))) - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2748)) - (Prims.of_int (6)) - (Prims.of_int (2768)) - (Prims.of_int (51))) - (if - Prims.op_Negation - (FStar_Reflection_Derived.is_fvar hd0 - "Prims.squash") - then - FStar_Tactics_Derived.fail - (Prims.strcat loc - " unexpected non-squash goal in goal_to_equiv") - else - FStar_Tactics_Effect.lift_div_tac - (fun uu___1 -> ())) + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2746)) (Prims.of_int (6)) + (Prims.of_int (2747)) (Prims.of_int (70))) + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2748)) (Prims.of_int (6)) + (Prims.of_int (2768)) (Prims.of_int (51))) + (if + Prims.op_Negation + (FStar_Reflection_Derived.is_fvar hd0 + "Prims.squash") + then + Obj.magic + (Obj.repr + (FStar_Tactics_Derived.fail + (Prims.strcat loc + " unexpected non-squash goal in goal_to_equiv"))) + else + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___1 -> ())))) + (fun uu___ -> (fun uu___ -> - (fun uu___ -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2748)) - (Prims.of_int (21)) - (Prims.of_int (2748)) - (Prims.of_int (34))) - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2748)) - (Prims.of_int (6)) - (Prims.of_int (2768)) - (Prims.of_int (51))) - (FStar_Tactics_Effect.lift_div_tac - (fun uu___1 -> - FStar_Reflection_Derived.collect_app - t1)) + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2748)) + (Prims.of_int (21)) + (Prims.of_int (2748)) + (Prims.of_int (34))) + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2748)) + (Prims.of_int (6)) + (Prims.of_int (2768)) + (Prims.of_int (51))) + (FStar_Tactics_Effect.lift_div_tac (fun uu___1 -> - (fun uu___1 -> - match uu___1 with - | (hd, args) -> - if - FStar_Reflection_Derived.is_fvar - hd - "Steel.Effect.Common.can_be_split" - then - Obj.magic - (Obj.repr - (FStar_Tactics_Derived.apply_lemma - (FStar_Reflection_Builtins.pack_ln - (FStar_Reflection_Data.Tv_FVar - (FStar_Reflection_Builtins.pack_fv + FStar_Reflection_Derived.collect_app + t1)) + (fun uu___1 -> + (fun uu___1 -> + match uu___1 with + | (hd, args) -> + if + FStar_Reflection_Derived.is_fvar + hd + "Steel.Effect.Common.can_be_split" + then + Obj.magic + (FStar_Tactics_Derived.apply_lemma + (FStar_Reflection_Builtins.pack_ln + (FStar_Reflection_Data.Tv_FVar + (FStar_Reflection_Builtins.pack_fv ["Steel"; "Effect"; "Common"; - "equiv_can_be_split"]))))) - else - Obj.magic - (Obj.repr - (if - FStar_Reflection_Derived.is_fvar - hd - "Steel.Effect.Common.can_be_split_forall" - then - Obj.repr + "equiv_can_be_split"])))) + else + if + FStar_Reflection_Derived.is_fvar + hd + "Steel.Effect.Common.can_be_split_forall" + then + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2752)) + (Prims.of_int (8)) + (Prims.of_int (2752)) + (Prims.of_int (32))) + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2753)) + (Prims.of_int (8)) + (Prims.of_int (2753)) + (Prims.of_int (41))) + (Obj.magic + (FStar_Tactics_Effect.tac_bind ( - FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2752)) - (Prims.of_int (8)) - (Prims.of_int (2752)) - (Prims.of_int (32))) - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2753)) - (Prims.of_int (8)) - (Prims.of_int (2753)) - (Prims.of_int (41))) - (Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range + Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2752)) (Prims.of_int (15)) (Prims.of_int (2752)) (Prims.of_int (32))) - (Prims.mk_range + ( + Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2752)) (Prims.of_int (8)) (Prims.of_int (2752)) (Prims.of_int (32))) - (Obj.magic + ( + Obj.magic (FStar_Tactics_Logic.forall_intro ())) - (fun + ( + fun uu___3 -> FStar_Tactics_Effect.lift_div_tac (fun uu___4 -> ())))) - (fun - uu___3 -> - (fun - uu___3 -> + (fun uu___3 -> + (fun uu___3 + -> Obj.magic (FStar_Tactics_Derived.apply_lemma (FStar_Reflection_Builtins.pack_ln @@ -12150,31 +12043,29 @@ let (goal_to_equiv : "Effect"; "Common"; "equiv_can_be_split"]))))) - uu___3)) - else - Obj.repr - ( - if - FStar_Reflection_Derived.is_fvar - hd - "Steel.Effect.Common.equiv_forall" - then - Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2755)) - (Prims.of_int (8)) - (Prims.of_int (2755)) - (Prims.of_int (40))) - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2756)) - (Prims.of_int (8)) - (Prims.of_int (2756)) - (Prims.of_int (32))) - (Obj.magic - (FStar_Tactics_Derived.apply_lemma + uu___3)) + else + if + FStar_Reflection_Derived.is_fvar + hd + "Steel.Effect.Common.equiv_forall" + then + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2755)) + (Prims.of_int (8)) + (Prims.of_int (2755)) + (Prims.of_int (40))) + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2756)) + (Prims.of_int (8)) + (Prims.of_int (2756)) + (Prims.of_int (32))) + (Obj.magic + (FStar_Tactics_Derived.apply_lemma (FStar_Reflection_Builtins.pack_ln (FStar_Reflection_Data.Tv_FVar (FStar_Reflection_Builtins.pack_fv @@ -12182,9 +12073,9 @@ let (goal_to_equiv : "Effect"; "Common"; "equiv_forall_elim"]))))) - (fun - uu___4 -> - (fun + (fun uu___4 + -> + (fun uu___4 -> Obj.magic (FStar_Tactics_Effect.tac_bind @@ -12210,28 +12101,27 @@ let (goal_to_equiv : uu___6 -> ())))) uu___4)) - else - Obj.repr - (if - FStar_Reflection_Derived.is_fvar - hd - "Steel.Effect.Common.can_be_split_post" - then - Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range + else + if + FStar_Reflection_Derived.is_fvar + hd + "Steel.Effect.Common.can_be_split_post" + then + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2758)) (Prims.of_int (8)) (Prims.of_int (2758)) (Prims.of_int (45))) - (Prims.mk_range + (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2759)) (Prims.of_int (8)) (Prims.of_int (2761)) (Prims.of_int (32))) - (Obj.magic + (Obj.magic (FStar_Tactics_Derived.apply_lemma (FStar_Reflection_Builtins.pack_ln (FStar_Reflection_Data.Tv_FVar @@ -12240,8 +12130,8 @@ let (goal_to_equiv : "Effect"; "Common"; "can_be_split_post_elim"]))))) - (fun - uu___5 -> + (fun uu___5 + -> (fun uu___5 -> Obj.magic @@ -12332,39 +12222,41 @@ let (goal_to_equiv : uu___7))) uu___6))) uu___5)) - else - Obj.repr - (if - FStar_Reflection_Derived.is_fvar - hd - "Steel.Effect.Common.can_be_split_dep" - then - FStar_Tactics_Derived.fail - (Prims.strcat + else + if + FStar_Reflection_Derived.is_fvar + hd + "Steel.Effect.Common.can_be_split_dep" + then + Obj.magic + (FStar_Tactics_Derived.fail + ( + Prims.strcat "can_be_split_dep not supported in " - loc) - else - if - FStar_Reflection_Derived.is_fvar - hd - "Steel.Effect.Common.can_be_split_forall_dep" - then - FStar_Tactics_Derived.fail + loc)) + else + if + FStar_Reflection_Derived.is_fvar + hd + "Steel.Effect.Common.can_be_split_forall_dep" + then + Obj.magic + (FStar_Tactics_Derived.fail (Prims.strcat "can_be_split_forall_dep not supported in " - loc) - else - FStar_Tactics_Derived.fail + loc)) + else + Obj.magic + (FStar_Tactics_Derived.fail (Prims.strcat loc - " goal in unexpected position"))))))) - uu___1))) uu___))) + " goal in unexpected position"))) + uu___1))) uu___)) | uu___ -> Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - (Prims.strcat loc " unexpected goal")))) - uu___))) uu___) + (FStar_Tactics_Derived.fail + (Prims.strcat loc " unexpected goal"))) uu___))) + uu___) let rec term_dict_assoc : 'a . FStar_Reflection_Types.term -> @@ -13419,13 +13311,11 @@ let rec (resolve_tac : if uu___2 then Obj.magic - (Obj.repr - (resolve_tac dict)) + (resolve_tac dict) else Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "Could not make progress, no solvable goal found"))) + (FStar_Tactics_Derived.fail + "Could not make progress, no solvable goal found")) uu___2))) uu___2))) uu___1)))) uu___) let rec (pick_next_logical : @@ -13714,16 +13604,14 @@ let (is_true : match uu___1 with | FStar_Reflection_Formula.True_ -> Obj.magic - (Obj.repr - (FStar_Tactics_Derived.exact - (FStar_Reflection_Builtins.pack_ln - (FStar_Reflection_Data.Tv_Const - FStar_Reflection_Data.C_Unit)))) + (FStar_Tactics_Derived.exact + (FStar_Reflection_Builtins.pack_ln + (FStar_Reflection_Data.Tv_Const + FStar_Reflection_Data.C_Unit))) | uu___2 -> Obj.magic - (Obj.repr - (FStar_Tactics_Effect.raise - FStar_Tactics_Derived.Goal_not_trivial))) uu___1) + (FStar_Tactics_Effect.raise + FStar_Tactics_Derived.Goal_not_trivial)) uu___1) let rec (solve_maybe_emps : Prims.nat -> (unit, unit) FStar_Tactics_Effect.tac_repr) = fun uu___ -> @@ -14641,13 +14529,11 @@ let (ite_soundness_tac : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) match uu___4 with | [] -> Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "should not happen")) + (FStar_Tactics_Derived.fail + "should not happen") | uu___5::tl -> Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind + (FStar_Tactics_Effect.tac_bind (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (3153)) @@ -14865,7 +14751,7 @@ let (ite_soundness_tac : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) uu___9))) uu___8))) uu___7))) - uu___6)))) + uu___6))) uu___4))) uu___4))) uu___3))) uu___2))) uu___1) let (vc_norm : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = diff --git a/ocaml/fstar-lib/generated/Steel_ST_GenElim_Base.ml b/ocaml/fstar-lib/generated/Steel_ST_GenElim_Base.ml index 2ca503ea1dc..ad134920e48 100644 --- a/ocaml/fstar-lib/generated/Steel_ST_GenElim_Base.ml +++ b/ocaml/fstar-lib/generated/Steel_ST_GenElim_Base.ml @@ -240,28 +240,24 @@ let rec (solve_gen_unit_elim : (t2, FStar_Reflection_Data.Q_Explicit)::[] -> - Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.ST.GenElim.Base.fsti" - (Prims.of_int (335)) - (Prims.of_int (20)) - (Prims.of_int (335)) - (Prims.of_int (42))) - (Prims.mk_range - "Steel.ST.GenElim.Base.fsti" - (Prims.of_int (336)) - (Prims.of_int (10)) - (Prims.of_int (337)) - (Prims.of_int (68))) - (Obj.magic - ( - solve_gen_unit_elim + FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.ST.GenElim.Base.fsti" + (Prims.of_int (335)) + (Prims.of_int (20)) + (Prims.of_int (335)) + (Prims.of_int (42))) + (Prims.mk_range + "Steel.ST.GenElim.Base.fsti" + (Prims.of_int (336)) + (Prims.of_int (10)) + (Prims.of_int (337)) + (Prims.of_int (68))) + (Obj.magic + (solve_gen_unit_elim t1)) - (fun uu___5 - -> - (fun t1' - -> + (fun uu___5 -> + (fun t1' -> Obj.magic (FStar_Tactics_Effect.tac_bind (Prims.mk_range @@ -298,11 +294,10 @@ let rec (solve_gen_unit_elim : FStar_Reflection_Data.Q_Explicit); (t2', FStar_Reflection_Data.Q_Explicit)])))) - uu___5)) + uu___5) | uu___5 -> - Obj.repr - (FStar_Tactics_Derived.fail - "ill-formed star"))) + FStar_Tactics_Derived.fail + "ill-formed star")) else Obj.magic (Obj.repr @@ -434,19 +429,25 @@ let rec (solve_gen_elim : (body, FStar_Reflection_Data.Q_Explicit)::[] -> - FStar_Tactics_Effect.lift_div_tac - (fun uu___3 -> - ([(ty, - FStar_Reflection_Data.Q_Implicit)], - body)) + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___3 -> + ([(ty, + FStar_Reflection_Data.Q_Implicit)], + body)))) | (body, FStar_Reflection_Data.Q_Explicit)::[] -> - FStar_Tactics_Effect.lift_div_tac - (fun uu___3 -> ([], body)) + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___3 -> ([], body)))) | uu___3 -> - FStar_Tactics_Derived.fail - "ill-formed exists_") + Obj.magic + (Obj.repr + (FStar_Tactics_Derived.fail + "ill-formed exists_"))) (fun uu___3 -> (fun uu___3 -> match uu___3 with @@ -804,37 +805,35 @@ let rec (solve_gen_elim : (tr, FStar_Reflection_Data.Q_Explicit)::[] -> - Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.ST.GenElim.Base.fsti" - (Prims.of_int (383)) - (Prims.of_int (15)) - (Prims.of_int (383)) - (Prims.of_int (42))) - (Prims.mk_range - "Steel.ST.GenElim.Base.fsti" - (Prims.of_int (383)) - (Prims.of_int (12)) - (Prims.of_int (396)) - (Prims.of_int (72))) - (Obj.magic - (term_has_head - tl - (FStar_Reflection_Builtins.pack_ln + FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.ST.GenElim.Base.fsti" + (Prims.of_int (383)) + (Prims.of_int (15)) + (Prims.of_int (383)) + (Prims.of_int (42))) + (Prims.mk_range + "Steel.ST.GenElim.Base.fsti" + (Prims.of_int (383)) + (Prims.of_int (12)) + (Prims.of_int (396)) + (Prims.of_int (72))) + (Obj.magic + (term_has_head tl + (FStar_Reflection_Builtins.pack_ln + (FStar_Reflection_Data.Tv_FVar ( - FStar_Reflection_Data.Tv_FVar - (FStar_Reflection_Builtins.pack_fv + FStar_Reflection_Builtins.pack_fv ["Steel"; "ST"; "Util"; "exists_"]))))) + (fun uu___5 -> (fun uu___5 -> - (fun uu___5 -> - if uu___5 - then - Obj.magic - (FStar_Tactics_Effect.tac_bind + if uu___5 + then + Obj.magic + (FStar_Tactics_Effect.tac_bind (Prims.mk_range "Steel.ST.GenElim.Base.fsti" (Prims.of_int (385)) @@ -959,9 +958,9 @@ let rec (solve_gen_elim : FStar_Reflection_Data.Q_Explicit)])))) uu___6))) uu___6)) - else - Obj.magic - (FStar_Tactics_Effect.tac_bind + else + Obj.magic + (FStar_Tactics_Effect.tac_bind (Prims.mk_range "Steel.ST.GenElim.Base.fsti" (Prims.of_int (394)) @@ -1018,11 +1017,10 @@ let rec (solve_gen_elim : (tr', FStar_Reflection_Data.Q_Explicit)])))) uu___7))) - uu___5)) + uu___5) | uu___5 -> - Obj.repr - (FStar_Tactics_Derived.fail - "ill-formed star"))) + FStar_Tactics_Derived.fail + "ill-formed star")) else Obj.magic (Obj.repr @@ -1784,70 +1782,78 @@ let (solve_gen_elim_prop : FStar_Tactics_Effect.lift_div_tac (fun uu___3 -> Prims.op_Negation uu___2)))) (fun uu___2 -> - if uu___2 - then - FStar_Tactics_Derived.fail - "not a squash goal" - else - FStar_Tactics_Effect.lift_div_tac - (fun uu___4 -> ())))) + (fun uu___2 -> + if uu___2 + then + Obj.magic + (Obj.repr + (FStar_Tactics_Derived.fail + "not a squash goal")) + else + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___4 -> ())))) uu___2))) (fun uu___2 -> (fun uu___2 -> match tl with | (body1, FStar_Reflection_Data.Q_Explicit)::[] -> Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.ST.GenElim.Base.fsti" - (Prims.of_int (571)) - (Prims.of_int (21)) - (Prims.of_int (571)) - (Prims.of_int (40))) - (Prims.mk_range - "Steel.ST.GenElim.Base.fsti" - (Prims.of_int (571)) - (Prims.of_int (4)) - (Prims.of_int (603)) - (Prims.of_int (7))) - (FStar_Tactics_Effect.lift_div_tac - (fun uu___3 -> - FStar_Reflection_Derived.collect_app - body1)) + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.ST.GenElim.Base.fsti" + (Prims.of_int (571)) + (Prims.of_int (21)) + (Prims.of_int (571)) + (Prims.of_int (40))) + (Prims.mk_range + "Steel.ST.GenElim.Base.fsti" + (Prims.of_int (571)) + (Prims.of_int (4)) + (Prims.of_int (603)) + (Prims.of_int (7))) + (FStar_Tactics_Effect.lift_div_tac (fun uu___3 -> - (fun uu___3 -> - match uu___3 with - | (hd1, tl1) -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.ST.GenElim.Base.fsti" - (Prims.of_int (572)) - (Prims.of_int (4)) - (Prims.of_int (573)) - (Prims.of_int (42))) - (Prims.mk_range - "Steel.ST.GenElim.Base.fsti" - (Prims.of_int (574)) - (Prims.of_int (10)) - (Prims.of_int (602)) - (Prims.of_int (44))) - (if - Prims.op_Negation - (is_fvar hd1 - "Steel.ST.GenElim.Base.gen_elim_prop") - then - FStar_Tactics_Derived.fail - "not a gen_elim_prop goal" - else - FStar_Tactics_Effect.lift_div_tac - (fun uu___5 -> ())) + FStar_Reflection_Derived.collect_app + body1)) + (fun uu___3 -> + (fun uu___3 -> + match uu___3 with + | (hd1, tl1) -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.ST.GenElim.Base.fsti" + (Prims.of_int (572)) + (Prims.of_int (4)) + (Prims.of_int (573)) + (Prims.of_int (42))) + (Prims.mk_range + "Steel.ST.GenElim.Base.fsti" + (Prims.of_int (574)) + (Prims.of_int (10)) + (Prims.of_int (602)) + (Prims.of_int (44))) + (if + Prims.op_Negation + (is_fvar hd1 + "Steel.ST.GenElim.Base.gen_elim_prop") + then + Obj.magic + (Obj.repr + (FStar_Tactics_Derived.fail + "not a gen_elim_prop goal")) + else + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___5 -> + ())))) + (fun uu___4 -> (fun uu___4 -> - (fun uu___4 -> - match FStar_List_Tot_Base.filter - ( - fun - uu___5 -> + match FStar_List_Tot_Base.filter + (fun uu___5 + -> match uu___5 with | @@ -1855,37 +1861,35 @@ let (solve_gen_elim_prop : x) -> FStar_Reflection_Data.uu___is_Q_Explicit x) tl1 - with - | (enable_nondep_opt_tm, - uu___5):: - (p, uu___6):: - (a, uu___7):: - (q, uu___8):: - (post, - uu___9)::[] - -> - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range + with + | (enable_nondep_opt_tm, + uu___5):: + (p, uu___6):: + (a, uu___7):: + (q, uu___8):: + (post, uu___9)::[] + -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.ST.GenElim.Base.fsti" (Prims.of_int (576)) (Prims.of_int (30)) (Prims.of_int (576)) (Prims.of_int (74))) - (Prims.mk_range + (Prims.mk_range "Steel.ST.GenElim.Base.fsti" (Prims.of_int (577)) (Prims.of_int (6)) (Prims.of_int (601)) (Prims.of_int (44))) - (Obj.magic + (Obj.magic (FStar_Tactics_Builtins.term_eq_old enable_nondep_opt_tm (FStar_Reflection_Builtins.pack_ln (FStar_Reflection_Data.Tv_Const FStar_Reflection_Data.C_True)))) - (fun + (fun uu___10 -> (fun @@ -2412,19 +2416,16 @@ let (solve_gen_elim_prop : uu___10))) uu___10))) uu___10))) - uu___10))) - | uu___5 -> - Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "ill formed gen_elim_prop"))) - uu___4))) uu___3))) + uu___10)) + | uu___5 -> + Obj.magic + (FStar_Tactics_Derived.fail + "ill formed gen_elim_prop")) + uu___4))) uu___3)) | uu___3 -> Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "ill-formed squash"))) uu___2))) - uu___1) + (FStar_Tactics_Derived.fail + "ill-formed squash")) uu___2))) uu___1) let (solve_gen_elim_prop_placeholder : unit -> (Prims.bool, unit) FStar_Tactics_Effect.tac_repr) = fun uu___ -> @@ -2476,70 +2477,78 @@ let (solve_gen_elim_prop_placeholder : FStar_Tactics_Effect.lift_div_tac (fun uu___3 -> Prims.op_Negation uu___2)))) (fun uu___2 -> - if uu___2 - then - FStar_Tactics_Derived.fail - "not a squash goal" - else - FStar_Tactics_Effect.lift_div_tac - (fun uu___4 -> ())))) + (fun uu___2 -> + if uu___2 + then + Obj.magic + (Obj.repr + (FStar_Tactics_Derived.fail + "not a squash goal")) + else + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___4 -> ())))) uu___2))) (fun uu___2 -> (fun uu___2 -> match tl with | (body1, FStar_Reflection_Data.Q_Explicit)::[] -> Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.ST.GenElim.Base.fsti" - (Prims.of_int (615)) - (Prims.of_int (21)) - (Prims.of_int (615)) - (Prims.of_int (40))) - (Prims.mk_range - "Steel.ST.GenElim.Base.fsti" - (Prims.of_int (615)) - (Prims.of_int (4)) - (Prims.of_int (643)) - (Prims.of_int (7))) - (FStar_Tactics_Effect.lift_div_tac - (fun uu___3 -> - FStar_Reflection_Derived.collect_app - body1)) + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.ST.GenElim.Base.fsti" + (Prims.of_int (615)) + (Prims.of_int (21)) + (Prims.of_int (615)) + (Prims.of_int (40))) + (Prims.mk_range + "Steel.ST.GenElim.Base.fsti" + (Prims.of_int (615)) + (Prims.of_int (4)) + (Prims.of_int (643)) + (Prims.of_int (7))) + (FStar_Tactics_Effect.lift_div_tac (fun uu___3 -> - (fun uu___3 -> - match uu___3 with - | (hd1, tl1) -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.ST.GenElim.Base.fsti" - (Prims.of_int (616)) - (Prims.of_int (4)) - (Prims.of_int (617)) - (Prims.of_int (54))) - (Prims.mk_range - "Steel.ST.GenElim.Base.fsti" - (Prims.of_int (618)) - (Prims.of_int (10)) - (Prims.of_int (642)) - (Prims.of_int (56))) - (if - Prims.op_Negation - (is_fvar hd1 - "Steel.ST.GenElim.Base.gen_elim_prop_placeholder") - then - FStar_Tactics_Derived.fail - "not a gen_elim_prop_placeholder goal" - else - FStar_Tactics_Effect.lift_div_tac - (fun uu___5 -> ())) + FStar_Reflection_Derived.collect_app + body1)) + (fun uu___3 -> + (fun uu___3 -> + match uu___3 with + | (hd1, tl1) -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.ST.GenElim.Base.fsti" + (Prims.of_int (616)) + (Prims.of_int (4)) + (Prims.of_int (617)) + (Prims.of_int (54))) + (Prims.mk_range + "Steel.ST.GenElim.Base.fsti" + (Prims.of_int (618)) + (Prims.of_int (10)) + (Prims.of_int (642)) + (Prims.of_int (56))) + (if + Prims.op_Negation + (is_fvar hd1 + "Steel.ST.GenElim.Base.gen_elim_prop_placeholder") + then + Obj.magic + (Obj.repr + (FStar_Tactics_Derived.fail + "not a gen_elim_prop_placeholder goal")) + else + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___5 -> + ())))) + (fun uu___4 -> (fun uu___4 -> - (fun uu___4 -> - match FStar_List_Tot_Base.filter - ( - fun - uu___5 -> + match FStar_List_Tot_Base.filter + (fun uu___5 + -> match uu___5 with | @@ -2547,31 +2556,29 @@ let (solve_gen_elim_prop_placeholder : x) -> FStar_Reflection_Data.uu___is_Q_Explicit x) tl1 - with - | (enable_nondep_opt_tm, - uu___5):: - (p, uu___6):: - (a, uu___7):: - (q, uu___8):: - (post, - uu___9)::[] - -> - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range + with + | (enable_nondep_opt_tm, + uu___5):: + (p, uu___6):: + (a, uu___7):: + (q, uu___8):: + (post, uu___9)::[] + -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.ST.GenElim.Base.fsti" (Prims.of_int (620)) (Prims.of_int (6)) (Prims.of_int (621)) (Prims.of_int (47))) - (Prims.mk_range + (Prims.mk_range "Steel.ST.GenElim.Base.fsti" (Prims.of_int (622)) (Prims.of_int (6)) (Prims.of_int (641)) (Prims.of_int (10))) - (Obj.magic + (Obj.magic (FStar_Tactics_Effect.tac_bind (Prims.mk_range "Steel.ST.GenElim.Base.fsti" @@ -2615,17 +2622,25 @@ let (solve_gen_elim_prop_placeholder : (fun uu___10 -> + (fun + uu___10 + -> if uu___10 then - FStar_Tactics_Derived.fail - "pre-resource not solved yet" + Obj.magic + (Obj.repr + (FStar_Tactics_Derived.fail + "pre-resource not solved yet")) else - FStar_Tactics_Effect.lift_div_tac + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac (fun uu___12 -> ())))) - (fun + uu___10))) + (fun uu___10 -> (fun @@ -2791,13 +2806,17 @@ let (solve_gen_elim_prop_placeholder : && post_is_uvar) then - FStar_Tactics_Derived.fail - "gen_elim_prop_placeholder is already solved" + Obj.magic + (Obj.repr + (FStar_Tactics_Derived.fail + "gen_elim_prop_placeholder is already solved")) else - FStar_Tactics_Effect.lift_div_tac + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac (fun uu___12 - -> ())) + -> ())))) (fun uu___11 -> @@ -3196,19 +3215,16 @@ let (solve_gen_elim_prop_placeholder : uu___11))) uu___11))) uu___11))) - uu___10))) - | uu___5 -> - Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "ill-formed gen_elim_prop_placeholder"))) - uu___4))) uu___3))) + uu___10)) + | uu___5 -> + Obj.magic + (FStar_Tactics_Derived.fail + "ill-formed gen_elim_prop_placeholder")) + uu___4))) uu___3)) | uu___3 -> Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "ill-formed squash"))) uu___2))) - uu___1) + (FStar_Tactics_Derived.fail + "ill-formed squash")) uu___2))) uu___1) let (init_resolve_tac : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = fun uu___ -> Steel_Effect_Common.init_resolve_tac' From ff030ad48d84fd5446d086b65f82ff4a90e10c31 Mon Sep 17 00:00:00 2001 From: Aseem Rastogi Date: Mon, 20 Mar 2023 16:16:31 +0530 Subject: [PATCH 19/48] fixing an off by 1 error --- ocaml/fstar-lib/FStar_Parser_ParseIt.ml | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/ocaml/fstar-lib/FStar_Parser_ParseIt.ml b/ocaml/fstar-lib/FStar_Parser_ParseIt.ml index 9bc10ea9f5f..9a91d7fc285 100644 --- a/ocaml/fstar-lib/FStar_Parser_ParseIt.ml +++ b/ocaml/fstar-lib/FStar_Parser_ParseIt.ml @@ -147,7 +147,7 @@ let parse fn = let contents_at = let lines = U.splitlines contents in fun (start_pos:Lexing.position) (end_pos:Lexing.position) -> - let suffix = FStar_Compiler_Util.nth_tail (Z.of_int (start_pos.pos_lnum - 1)) lines in + let suffix = FStar_Compiler_Util.nth_tail (Z.of_int start_pos.pos_lnum) lines in let text, _ = FStar_Compiler_Util.first_N (Z.of_int (end_pos.pos_lnum - start_pos.pos_lnum)) suffix in let range = range_of_positions start_pos end_pos in { range; From 64eb234539de816eb48f80d377b417a7b769f695 Mon Sep 17 00:00:00 2001 From: Aseem Rastogi Date: Mon, 20 Mar 2023 16:16:34 +0530 Subject: [PATCH 20/48] snap --- .../FStar_InteractiveHelpers_Base.ml | 4 +- .../FStar_InteractiveHelpers_Effectful.ml | 21 +- .../FStar_InteractiveHelpers_ExploreTerm.ml | 265 +- .../FStar_InteractiveHelpers_PostProcess.ml | 109 +- ocaml/fstar-lib/generated/FStar_Tactics_BV.ml | 8 +- .../generated/FStar_Tactics_Canon.ml | 8 +- .../FStar_Tactics_CanonCommMonoid.ml | 68 +- .../FStar_Tactics_CanonCommMonoidSimple.ml | 188 +- ...tar_Tactics_CanonCommMonoidSimple_Equiv.ml | 99 +- .../FStar_Tactics_CanonCommSemiring.ml | 56 +- .../generated/FStar_Tactics_CanonMonoid.ml | 148 +- .../generated/FStar_Tactics_Derived.ml | 1062 +++---- .../generated/FStar_Tactics_Effect.ml | 4 +- .../generated/FStar_Tactics_Logic.ml | 225 +- .../FStar_Tactics_PatternMatching.ml | 678 +++-- .../generated/FStar_Tactics_Simplifier.ml | 324 +- .../generated/FStar_Tactics_SyntaxHelpers.ml | 6 +- .../generated/FStar_Tactics_Typeclasses.ml | 73 +- .../generated/Steel_Effect_Common.ml | 2686 +++++++++-------- .../generated/Steel_ST_GenElim_Base.ml | 498 ++- 20 files changed, 3388 insertions(+), 3142 deletions(-) diff --git a/ocaml/fstar-lib/generated/FStar_InteractiveHelpers_Base.ml b/ocaml/fstar-lib/generated/FStar_InteractiveHelpers_Base.ml index bea6412242a..e8b1b42a385 100644 --- a/ocaml/fstar-lib/generated/FStar_InteractiveHelpers_Base.ml +++ b/ocaml/fstar-lib/generated/FStar_InteractiveHelpers_Base.ml @@ -946,7 +946,9 @@ let (__proj__MetaAnalysis__item__uu___ : Prims.exn -> Prims.string) = fun projectee -> match projectee with | MetaAnalysis uu___ -> uu___ let mfail : 'uuuuu . Prims.string -> ('uuuuu, unit) FStar_Tactics_Effect.tac_repr = - fun str -> FStar_Tactics_Effect.raise (MetaAnalysis str) + fun uu___ -> + (fun str -> Obj.magic (FStar_Tactics_Effect.raise (MetaAnalysis str))) + uu___ let (print_dbg : Prims.bool -> Prims.string -> (unit, unit) FStar_Tactics_Effect.tac_repr) = fun uu___1 -> diff --git a/ocaml/fstar-lib/generated/FStar_InteractiveHelpers_Effectful.ml b/ocaml/fstar-lib/generated/FStar_InteractiveHelpers_Effectful.ml index d113c2ed518..b6e37006c5d 100644 --- a/ocaml/fstar-lib/generated/FStar_InteractiveHelpers_Effectful.ml +++ b/ocaml/fstar-lib/generated/FStar_InteractiveHelpers_Effectful.ml @@ -1431,13 +1431,20 @@ let (compute_eterm_info : parameters)))) uu___2))) uu___2)) (fun uu___1 -> - match uu___1 with - | FStar_Tactics_Common.TacticFailure msg -> - FStar_InteractiveHelpers_Base.mfail - (Prims.strcat - "compute_eterm_info: failure: '" - (Prims.strcat msg "'")) - | e1 -> FStar_Tactics_Effect.raise e1))) uu___) + (fun uu___1 -> + match uu___1 with + | FStar_Tactics_Common.TacticFailure msg -> + Obj.magic + (Obj.repr + (FStar_InteractiveHelpers_Base.mfail + (Prims.strcat + "compute_eterm_info: failure: '" + (Prims.strcat msg "'")))) + | e1 -> + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.raise e1))) + uu___1))) uu___) let (has_refinement : FStar_InteractiveHelpers_ExploreTerm.type_info -> Prims.bool) = fun ty -> diff --git a/ocaml/fstar-lib/generated/FStar_InteractiveHelpers_ExploreTerm.ml b/ocaml/fstar-lib/generated/FStar_InteractiveHelpers_ExploreTerm.ml index 7428ba9e9f6..fa2b8a6733f 100644 --- a/ocaml/fstar-lib/generated/FStar_InteractiveHelpers_ExploreTerm.ml +++ b/ocaml/fstar-lib/generated/FStar_InteractiveHelpers_ExploreTerm.ml @@ -1556,13 +1556,21 @@ let rec (inst_comp : match () with | () -> inst_comp_once e c t) (fun uu___ -> - match uu___ with - | FStar_InteractiveHelpers_Base.MetaAnalysis - msg -> - FStar_InteractiveHelpers_Base.mfail - (Prims.strcat "inst_comp: error: " - msg) - | err -> FStar_Tactics_Effect.raise err))) + (fun uu___ -> + match uu___ with + | FStar_InteractiveHelpers_Base.MetaAnalysis + msg -> + Obj.magic + (Obj.repr + (FStar_InteractiveHelpers_Base.mfail + (Prims.strcat + "inst_comp: error: " + msg))) + | err -> + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.raise + err))) uu___))) (fun uu___ -> (fun c' -> Obj.magic (inst_comp e c' tl')) uu___)))) uu___2 uu___1 uu___ @@ -1661,54 +1669,68 @@ let (_abs_update_typ : "_abs_update_typ: inconsistent state")) uu___1))) uu___1)) (fun uu___ -> - match uu___ with - | FStar_InteractiveHelpers_Base.MetaAnalysis msg -> - FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.InteractiveHelpers.ExploreTerm.fst" - (Prims.of_int (303)) (Prims.of_int (10)) - (Prims.of_int (303)) (Prims.of_int (93))) - (Prims.mk_range - "FStar.InteractiveHelpers.ExploreTerm.fst" - (Prims.of_int (303)) (Prims.of_int (4)) - (Prims.of_int (303)) (Prims.of_int (93))) - (Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.InteractiveHelpers.ExploreTerm.fst" - (Prims.of_int (303)) (Prims.of_int (61)) - (Prims.of_int (303)) (Prims.of_int (92))) - (Prims.mk_range "prims.fst" (Prims.of_int (606)) - (Prims.of_int (19)) (Prims.of_int (606)) - (Prims.of_int (31))) - (Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.InteractiveHelpers.ExploreTerm.fst" - (Prims.of_int (303)) (Prims.of_int (61)) - (Prims.of_int (303)) (Prims.of_int (78))) - (Prims.mk_range "prims.fst" - (Prims.of_int (606)) (Prims.of_int (19)) - (Prims.of_int (606)) (Prims.of_int (31))) - (Obj.magic - (FStar_Tactics_Builtins.term_to_string ty)) + (fun uu___ -> + match uu___ with + | FStar_InteractiveHelpers_Base.MetaAnalysis msg -> + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.InteractiveHelpers.ExploreTerm.fst" + (Prims.of_int (303)) (Prims.of_int (10)) + (Prims.of_int (303)) (Prims.of_int (93))) + (Prims.mk_range + "FStar.InteractiveHelpers.ExploreTerm.fst" + (Prims.of_int (303)) (Prims.of_int (4)) + (Prims.of_int (303)) (Prims.of_int (93))) + (Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.InteractiveHelpers.ExploreTerm.fst" + (Prims.of_int (303)) + (Prims.of_int (61)) + (Prims.of_int (303)) + (Prims.of_int (92))) + (Prims.mk_range "prims.fst" + (Prims.of_int (606)) + (Prims.of_int (19)) + (Prims.of_int (606)) + (Prims.of_int (31))) + (Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.InteractiveHelpers.ExploreTerm.fst" + (Prims.of_int (303)) + (Prims.of_int (61)) + (Prims.of_int (303)) + (Prims.of_int (78))) + (Prims.mk_range "prims.fst" + (Prims.of_int (606)) + (Prims.of_int (19)) + (Prims.of_int (606)) + (Prims.of_int (31))) + (Obj.magic + (FStar_Tactics_Builtins.term_to_string + ty)) + (fun uu___1 -> + FStar_Tactics_Effect.lift_div_tac + (fun uu___2 -> + Prims.strcat uu___1 + (Prims.strcat ":\n" msg))))) + (fun uu___1 -> + FStar_Tactics_Effect.lift_div_tac + (fun uu___2 -> + Prims.strcat + "_abs_update_typ: could not find an arrow in: " + uu___1)))) + (fun uu___1 -> (fun uu___1 -> - FStar_Tactics_Effect.lift_div_tac - (fun uu___2 -> - Prims.strcat uu___1 - (Prims.strcat ":\n" msg))))) - (fun uu___1 -> - FStar_Tactics_Effect.lift_div_tac - (fun uu___2 -> - Prims.strcat - "_abs_update_typ: could not find an arrow in: " - uu___1)))) - (fun uu___1 -> - (fun uu___1 -> - Obj.magic - (FStar_InteractiveHelpers_Base.mfail uu___1)) - uu___1) - | err -> FStar_Tactics_Effect.raise err) + Obj.magic + (FStar_InteractiveHelpers_Base.mfail + uu___1)) uu___1))) + | err -> + Obj.magic (Obj.repr (FStar_Tactics_Effect.raise err))) + uu___) let (abs_update_typ_or_comp : FStar_Reflection_Types.binder -> typ_or_comp -> @@ -1780,14 +1802,12 @@ let (abs_update_opt_typ_or_comp : | FStar_InteractiveHelpers_Base.MetaAnalysis msg -> Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___1 -> - FStar_Pervasives_Native.None))) + (FStar_Tactics_Effect.lift_div_tac + (fun uu___1 -> + FStar_Pervasives_Native.None)) | err -> Obj.magic - (Obj.repr - (FStar_Tactics_Effect.raise err))) + (FStar_Tactics_Effect.raise err)) uu___)))) uu___2 uu___1 uu___ let rec (_flush_typ_or_comp_comp : Prims.bool -> @@ -2315,62 +2335,77 @@ let (flush_typ_or_comp : uu___1) | TC_Comp (c, pl, n) -> flush_comp pl n c)) (fun uu___ -> - match uu___ with - | FStar_InteractiveHelpers_Base.MetaAnalysis msg -> - FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.InteractiveHelpers.ExploreTerm.fst" - (Prims.of_int (379)) (Prims.of_int (15)) - (Prims.of_int (379)) (Prims.of_int (90))) - (Prims.mk_range - "FStar.InteractiveHelpers.ExploreTerm.fst" - (Prims.of_int (379)) (Prims.of_int (9)) - (Prims.of_int (379)) (Prims.of_int (90))) - (Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.InteractiveHelpers.ExploreTerm.fst" - (Prims.of_int (379)) - (Prims.of_int (50)) - (Prims.of_int (379)) - (Prims.of_int (89))) - (Prims.mk_range "prims.fst" - (Prims.of_int (606)) - (Prims.of_int (19)) - (Prims.of_int (606)) - (Prims.of_int (31))) - (Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.InteractiveHelpers.ExploreTerm.fst" - (Prims.of_int (379)) - (Prims.of_int (50)) - (Prims.of_int (379)) - (Prims.of_int (75))) - (Prims.mk_range "prims.fst" - (Prims.of_int (606)) - (Prims.of_int (19)) - (Prims.of_int (606)) - (Prims.of_int (31))) - (Obj.magic - (typ_or_comp_to_string tyc)) + (fun uu___ -> + match uu___ with + | FStar_InteractiveHelpers_Base.MetaAnalysis msg + -> + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.InteractiveHelpers.ExploreTerm.fst" + (Prims.of_int (379)) + (Prims.of_int (15)) + (Prims.of_int (379)) + (Prims.of_int (90))) + (Prims.mk_range + "FStar.InteractiveHelpers.ExploreTerm.fst" + (Prims.of_int (379)) + (Prims.of_int (9)) + (Prims.of_int (379)) + (Prims.of_int (90))) + (Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.InteractiveHelpers.ExploreTerm.fst" + (Prims.of_int (379)) + (Prims.of_int (50)) + (Prims.of_int (379)) + (Prims.of_int (89))) + (Prims.mk_range "prims.fst" + (Prims.of_int (606)) + (Prims.of_int (19)) + (Prims.of_int (606)) + (Prims.of_int (31))) + (Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.InteractiveHelpers.ExploreTerm.fst" + (Prims.of_int (379)) + (Prims.of_int (50)) + (Prims.of_int (379)) + (Prims.of_int (75))) + (Prims.mk_range + "prims.fst" + (Prims.of_int (606)) + (Prims.of_int (19)) + (Prims.of_int (606)) + (Prims.of_int (31))) + (Obj.magic + (typ_or_comp_to_string + tyc)) + (fun uu___1 -> + FStar_Tactics_Effect.lift_div_tac + (fun uu___2 -> + Prims.strcat + uu___1 + (Prims.strcat + ":\n" msg))))) + (fun uu___1 -> + FStar_Tactics_Effect.lift_div_tac + (fun uu___2 -> + Prims.strcat + "flush_typ_or_comp failed on: " + uu___1)))) + (fun uu___1 -> (fun uu___1 -> - FStar_Tactics_Effect.lift_div_tac - (fun uu___2 -> - Prims.strcat uu___1 - (Prims.strcat ":\n" msg))))) - (fun uu___1 -> - FStar_Tactics_Effect.lift_div_tac - (fun uu___2 -> - Prims.strcat - "flush_typ_or_comp failed on: " - uu___1)))) - (fun uu___1 -> - (fun uu___1 -> - Obj.magic - (FStar_InteractiveHelpers_Base.mfail - uu___1)) uu___1) - | err -> FStar_Tactics_Effect.raise err))) uu___) + Obj.magic + (FStar_InteractiveHelpers_Base.mfail + uu___1)) uu___1))) + | err -> + Obj.magic + (Obj.repr (FStar_Tactics_Effect.raise err))) + uu___))) uu___) let (safe_arg_typ_or_comp : Prims.bool -> FStar_Reflection_Types.env -> diff --git a/ocaml/fstar-lib/generated/FStar_InteractiveHelpers_PostProcess.ml b/ocaml/fstar-lib/generated/FStar_InteractiveHelpers_PostProcess.ml index a3e884e6b41..a2430d7f851 100644 --- a/ocaml/fstar-lib/generated/FStar_InteractiveHelpers_PostProcess.ml +++ b/ocaml/fstar-lib/generated/FStar_InteractiveHelpers_PostProcess.ml @@ -2146,22 +2146,29 @@ let (pp_analyze_effectful_term : (fun uu___2 -> Obj.magic (end_proof ())) uu___2))) uu___2)) (fun uu___1 -> - match uu___1 with - | FStar_InteractiveHelpers_Base.MetaAnalysis msg -> - FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.InteractiveHelpers.PostProcess.fst" - (Prims.of_int (256)) (Prims.of_int (29)) - (Prims.of_int (256)) (Prims.of_int (49))) - (Prims.mk_range - "FStar.InteractiveHelpers.PostProcess.fst" - (Prims.of_int (256)) (Prims.of_int (51)) - (Prims.of_int (256)) (Prims.of_int (63))) - (Obj.magic - (FStar_InteractiveHelpers_Output.printout_failure msg)) - (fun uu___2 -> - (fun uu___2 -> Obj.magic (end_proof ())) uu___2) - | err -> FStar_Tactics_Effect.raise err) + (fun uu___1 -> + match uu___1 with + | FStar_InteractiveHelpers_Base.MetaAnalysis msg -> + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.InteractiveHelpers.PostProcess.fst" + (Prims.of_int (256)) (Prims.of_int (29)) + (Prims.of_int (256)) (Prims.of_int (49))) + (Prims.mk_range + "FStar.InteractiveHelpers.PostProcess.fst" + (Prims.of_int (256)) (Prims.of_int (51)) + (Prims.of_int (256)) (Prims.of_int (63))) + (Obj.magic + (FStar_InteractiveHelpers_Output.printout_failure + msg)) + (fun uu___2 -> + (fun uu___2 -> Obj.magic (end_proof ())) + uu___2))) + | err -> + Obj.magic (Obj.repr (FStar_Tactics_Effect.raise err))) + uu___1) let _ = FStar_Tactics_Native.register_tactic "FStar.InteractiveHelpers.PostProcess.pp_analyze_effectful_term" @@ -2605,20 +2612,27 @@ let (pp_split_assert_conjs : (fun uu___2 -> Obj.magic (end_proof ())) uu___2))) uu___2)) (fun uu___1 -> - match uu___1 with - | FStar_InteractiveHelpers_Base.MetaAnalysis msg -> - FStar_Tactics_Effect.tac_bind - (Prims.mk_range "FStar.InteractiveHelpers.PostProcess.fst" - (Prims.of_int (356)) (Prims.of_int (29)) - (Prims.of_int (356)) (Prims.of_int (49))) - (Prims.mk_range "FStar.InteractiveHelpers.PostProcess.fst" - (Prims.of_int (356)) (Prims.of_int (51)) - (Prims.of_int (356)) (Prims.of_int (63))) - (Obj.magic - (FStar_InteractiveHelpers_Output.printout_failure msg)) - (fun uu___2 -> - (fun uu___2 -> Obj.magic (end_proof ())) uu___2) - | err -> FStar_Tactics_Effect.raise err) + (fun uu___1 -> + match uu___1 with + | FStar_InteractiveHelpers_Base.MetaAnalysis msg -> + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.InteractiveHelpers.PostProcess.fst" + (Prims.of_int (356)) (Prims.of_int (29)) + (Prims.of_int (356)) (Prims.of_int (49))) + (Prims.mk_range + "FStar.InteractiveHelpers.PostProcess.fst" + (Prims.of_int (356)) (Prims.of_int (51)) + (Prims.of_int (356)) (Prims.of_int (63))) + (Obj.magic + (FStar_InteractiveHelpers_Output.printout_failure + msg)) + (fun uu___2 -> + (fun uu___2 -> Obj.magic (end_proof ())) uu___2))) + | err -> Obj.magic (Obj.repr (FStar_Tactics_Effect.raise err))) + uu___1) let _ = FStar_Tactics_Native.register_tactic "FStar.InteractiveHelpers.PostProcess.pp_split_assert_conjs" @@ -6913,20 +6927,27 @@ let (pp_unfold_in_assert_or_assume : (fun uu___2 -> Obj.magic (end_proof ())) uu___2))) uu___2)) (fun uu___1 -> - match uu___1 with - | FStar_InteractiveHelpers_Base.MetaAnalysis msg -> - FStar_Tactics_Effect.tac_bind - (Prims.mk_range "FStar.InteractiveHelpers.PostProcess.fst" - (Prims.of_int (748)) (Prims.of_int (29)) - (Prims.of_int (748)) (Prims.of_int (49))) - (Prims.mk_range "FStar.InteractiveHelpers.PostProcess.fst" - (Prims.of_int (748)) (Prims.of_int (51)) - (Prims.of_int (748)) (Prims.of_int (63))) - (Obj.magic - (FStar_InteractiveHelpers_Output.printout_failure msg)) - (fun uu___2 -> - (fun uu___2 -> Obj.magic (end_proof ())) uu___2) - | err -> FStar_Tactics_Effect.raise err) + (fun uu___1 -> + match uu___1 with + | FStar_InteractiveHelpers_Base.MetaAnalysis msg -> + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.InteractiveHelpers.PostProcess.fst" + (Prims.of_int (748)) (Prims.of_int (29)) + (Prims.of_int (748)) (Prims.of_int (49))) + (Prims.mk_range + "FStar.InteractiveHelpers.PostProcess.fst" + (Prims.of_int (748)) (Prims.of_int (51)) + (Prims.of_int (748)) (Prims.of_int (63))) + (Obj.magic + (FStar_InteractiveHelpers_Output.printout_failure + msg)) + (fun uu___2 -> + (fun uu___2 -> Obj.magic (end_proof ())) uu___2))) + | err -> Obj.magic (Obj.repr (FStar_Tactics_Effect.raise err))) + uu___1) let _ = FStar_Tactics_Native.register_tactic "FStar.InteractiveHelpers.PostProcess.pp_unfold_in_assert_or_assume" diff --git a/ocaml/fstar-lib/generated/FStar_Tactics_BV.ml b/ocaml/fstar-lib/generated/FStar_Tactics_BV.ml index cc7174ea77a..496adf26e40 100644 --- a/ocaml/fstar-lib/generated/FStar_Tactics_BV.ml +++ b/ocaml/fstar-lib/generated/FStar_Tactics_BV.ml @@ -892,11 +892,9 @@ let (arith_to_bv_tac : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = "arith_to_bv_tac: unexpected: " uu___4)))) (fun uu___4 -> - (fun uu___4 -> - Obj.magic - (FStar_Tactics_Derived.fail - uu___4)) uu___4))) - uu___3))) uu___3))) uu___2)) + FStar_Tactics_Derived.fail + uu___4))) uu___3))) + uu___3))) uu___2)) let (bv_tac : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = fun uu___ -> FStar_Tactics_Derived.focus diff --git a/ocaml/fstar-lib/generated/FStar_Tactics_Canon.ml b/ocaml/fstar-lib/generated/FStar_Tactics_Canon.ml index 93e57a67c65..ef14886be48 100644 --- a/ocaml/fstar-lib/generated/FStar_Tactics_Canon.ml +++ b/ocaml/fstar-lib/generated/FStar_Tactics_Canon.ml @@ -1028,10 +1028,8 @@ let (canon_point_entry : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) "impossible: " uu___4)))) (fun uu___4 -> - (fun uu___4 -> - Obj.magic - (FStar_Tactics_Derived.fail - uu___4)) uu___4))) - uu___2))) uu___2))) uu___1) + FStar_Tactics_Derived.fail + uu___4))) uu___2))) uu___2))) + uu___1) let (canon : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = fun uu___ -> FStar_Tactics_Derived.pointwise canon_point_entry \ No newline at end of file diff --git a/ocaml/fstar-lib/generated/FStar_Tactics_CanonCommMonoid.ml b/ocaml/fstar-lib/generated/FStar_Tactics_CanonCommMonoid.ml index d71fd4339c0..a6482cacc8d 100644 --- a/ocaml/fstar-lib/generated/FStar_Tactics_CanonCommMonoid.ml +++ b/ocaml/fstar-lib/generated/FStar_Tactics_CanonCommMonoid.ml @@ -1669,28 +1669,30 @@ let canon_monoid_aux : t), t1, t2) -> Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.CanonCommMonoid.fst" - (Prims.of_int (340)) - (Prims.of_int (9)) - (Prims.of_int (340)) - (Prims.of_int (25))) - (Prims.mk_range - "FStar.Tactics.CanonCommMonoid.fst" - (Prims.of_int (340)) - (Prims.of_int (6)) - (Prims.of_int (411)) - (Prims.of_int (69))) - (Obj.magic - (FStar_Tactics_Builtins.term_eq_old - t ta)) - (fun uu___2 -> + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.CanonCommMonoid.fst" + (Prims.of_int (340)) + (Prims.of_int (9)) + (Prims.of_int (340)) + (Prims.of_int (25))) + (Prims.mk_range + "FStar.Tactics.CanonCommMonoid.fst" + (Prims.of_int (340)) + (Prims.of_int (6)) + (Prims.of_int (411)) + (Prims.of_int (69))) + (Obj.magic + (FStar_Tactics_Builtins.term_eq_old + t ta)) (fun uu___2 -> - if uu___2 - then - Obj.magic - (FStar_Tactics_Effect.tac_bind + (fun uu___2 -> + if uu___2 + then + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.CanonCommMonoid.fst" (Prims.of_int (341)) @@ -1731,6 +1733,7 @@ let canon_monoid_aux : (r1::r2::[], vm) -> Obj.magic + (Obj.repr (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.CanonCommMonoid.fst" @@ -1997,22 +2000,25 @@ let canon_monoid_aux : uu___4))) uu___4))) uu___4))) - uu___4)) + uu___4))) | uu___4 -> Obj.magic + (Obj.repr (FStar_Tactics_Derived.fail - "Unexpected")) - uu___3)) - else - Obj.magic - (FStar_Tactics_Derived.fail - "Goal should be an equality at the right monoid type")) - uu___2)) + "Unexpected"))) + uu___3))) + else + Obj.magic + (Obj.repr + (FStar_Tactics_Derived.fail + "Goal should be an equality at the right monoid type"))) + uu___2))) | uu___2 -> Obj.magic - (FStar_Tactics_Derived.fail - "Goal should be an equality")) + (Obj.repr + (FStar_Tactics_Derived.fail + "Goal should be an equality"))) uu___1))) uu___) let canon_monoid_with : 'b . diff --git a/ocaml/fstar-lib/generated/FStar_Tactics_CanonCommMonoidSimple.ml b/ocaml/fstar-lib/generated/FStar_Tactics_CanonCommMonoidSimple.ml index d4730d05f43..16012c10b64 100644 --- a/ocaml/fstar-lib/generated/FStar_Tactics_CanonCommMonoidSimple.ml +++ b/ocaml/fstar-lib/generated/FStar_Tactics_CanonCommMonoidSimple.ml @@ -624,89 +624,97 @@ let canon_monoid : (FStar_Pervasives_Native.Some t), t1, t2) -> Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.CanonCommMonoidSimple.fst" - (Prims.of_int (255)) (Prims.of_int (9)) - (Prims.of_int (255)) (Prims.of_int (28))) - (Prims.mk_range - "FStar.Tactics.CanonCommMonoidSimple.fst" - (Prims.of_int (255)) (Prims.of_int (6)) - (Prims.of_int (273)) (Prims.of_int (69))) - (Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.CanonCommMonoidSimple.fst" - (Prims.of_int (255)) - (Prims.of_int (19)) - (Prims.of_int (255)) - (Prims.of_int (28))) - (Prims.mk_range - "FStar.Tactics.CanonCommMonoidSimple.fst" - (Prims.of_int (255)) - (Prims.of_int (9)) - (Prims.of_int (255)) - (Prims.of_int (28))) - (FStar_Tactics_Effect.lift_div_tac - (fun uu___2 -> + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.CanonCommMonoidSimple.fst" + (Prims.of_int (255)) (Prims.of_int (9)) + (Prims.of_int (255)) + (Prims.of_int (28))) + (Prims.mk_range + "FStar.Tactics.CanonCommMonoidSimple.fst" + (Prims.of_int (255)) (Prims.of_int (6)) + (Prims.of_int (273)) + (Prims.of_int (69))) + (Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.CanonCommMonoidSimple.fst" + (Prims.of_int (255)) + (Prims.of_int (19)) + (Prims.of_int (255)) + (Prims.of_int (28))) + (Prims.mk_range + "FStar.Tactics.CanonCommMonoidSimple.fst" + (Prims.of_int (255)) + (Prims.of_int (9)) + (Prims.of_int (255)) + (Prims.of_int (28))) + (FStar_Tactics_Effect.lift_div_tac (fun uu___2 -> - Obj.magic - (failwith - "Cannot evaluate open quotation at runtime")) - uu___2)) - (fun uu___2 -> + (fun uu___2 -> + Obj.magic + (failwith + "Cannot evaluate open quotation at runtime")) + uu___2)) (fun uu___2 -> - Obj.magic (term_eq t uu___2)) - uu___2))) - (fun uu___2 -> + (fun uu___2 -> + Obj.magic (term_eq t uu___2)) + uu___2))) (fun uu___2 -> - if uu___2 - then - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.CanonCommMonoidSimple.fst" - (Prims.of_int (256)) - (Prims.of_int (27)) - (Prims.of_int (256)) - (Prims.of_int (67))) - (Prims.mk_range - "FStar.Tactics.CanonCommMonoidSimple.fst" - (Prims.of_int (256)) - (Prims.of_int (8)) - (Prims.of_int (271)) - (Prims.of_int (22))) - (Obj.magic - (reification m [] - (const - (FStar_Algebra_CommMonoid.__proj__CM__item__unit - m)) t1)) - (fun uu___3 -> - (fun uu___3 -> - match uu___3 with - | (r1, ts, am) -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.CanonCommMonoidSimple.fst" - (Prims.of_int (257)) - (Prims.of_int (26)) - (Prims.of_int (257)) - (Prims.of_int (48))) - (Prims.mk_range - "FStar.Tactics.CanonCommMonoidSimple.fst" - (Prims.of_int (257)) - (Prims.of_int (8)) - (Prims.of_int (271)) - (Prims.of_int (22))) - (Obj.magic - (reification m - ts am t2)) - (fun uu___4 -> - (fun uu___4 -> - match uu___4 - with - | (r2, + (fun uu___2 -> + if uu___2 + then + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.CanonCommMonoidSimple.fst" + (Prims.of_int (256)) + (Prims.of_int (27)) + (Prims.of_int (256)) + (Prims.of_int (67))) + (Prims.mk_range + "FStar.Tactics.CanonCommMonoidSimple.fst" + (Prims.of_int (256)) + (Prims.of_int (8)) + (Prims.of_int (271)) + (Prims.of_int (22))) + (Obj.magic + (reification m [] + (const + (FStar_Algebra_CommMonoid.__proj__CM__item__unit + m)) t1)) + (fun uu___3 -> + (fun uu___3 -> + match uu___3 with + | (r1, ts, am) -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.CanonCommMonoidSimple.fst" + (Prims.of_int (257)) + (Prims.of_int (26)) + (Prims.of_int (257)) + (Prims.of_int (48))) + (Prims.mk_range + "FStar.Tactics.CanonCommMonoidSimple.fst" + (Prims.of_int (257)) + (Prims.of_int (8)) + (Prims.of_int (271)) + (Prims.of_int (22))) + (Obj.magic + (reification + m ts am + t2)) + (fun uu___4 + -> + (fun + uu___4 -> + match uu___4 + with + | + (r2, uu___5, am1) -> Obj.magic @@ -900,15 +908,17 @@ let canon_monoid : uu___8))) uu___7))) uu___6))) - uu___4))) - uu___3)) - else - Obj.magic - (FStar_Tactics_Derived.fail - "Goal should be an equality at the right monoid type")) - uu___2)) + uu___4))) + uu___3))) + else + Obj.magic + (Obj.repr + (FStar_Tactics_Derived.fail + "Goal should be an equality at the right monoid type"))) + uu___2))) | uu___2 -> Obj.magic - (FStar_Tactics_Derived.fail - "Goal should be an equality")) uu___1))) + (Obj.repr + (FStar_Tactics_Derived.fail + "Goal should be an equality"))) uu___1))) uu___) \ No newline at end of file diff --git a/ocaml/fstar-lib/generated/FStar_Tactics_CanonCommMonoidSimple_Equiv.ml b/ocaml/fstar-lib/generated/FStar_Tactics_CanonCommMonoidSimple_Equiv.ml index 6bd65b0452a..5d9bec7196b 100644 --- a/ocaml/fstar-lib/generated/FStar_Tactics_CanonCommMonoidSimple_Equiv.ml +++ b/ocaml/fstar-lib/generated/FStar_Tactics_CanonCommMonoidSimple_Equiv.ml @@ -1075,67 +1075,74 @@ let (canon_monoid : (match rel_xy with | (rel_xy1, uu___2)::[] -> Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.CanonCommMonoidSimple.Equiv.fst" - (Prims.of_int (381)) - (Prims.of_int (21)) - (Prims.of_int (381)) - (Prims.of_int (43))) - (Prims.mk_range - "FStar.Tactics.CanonCommMonoidSimple.Equiv.fst" - (Prims.of_int (380)) - (Prims.of_int (21)) - (Prims.of_int (391)) - (Prims.of_int (6))) - (FStar_Tactics_Effect.lift_div_tac + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.CanonCommMonoidSimple.Equiv.fst" + (Prims.of_int (381)) + (Prims.of_int (21)) + (Prims.of_int (381)) + (Prims.of_int (43))) + (Prims.mk_range + "FStar.Tactics.CanonCommMonoidSimple.Equiv.fst" + (Prims.of_int (380)) + (Prims.of_int (21)) + (Prims.of_int (391)) + (Prims.of_int (6))) + (FStar_Tactics_Effect.lift_div_tac + (fun uu___3 -> + FStar_Reflection_Derived_Lemmas.collect_app_ref + rel_xy1)) (fun uu___3 -> - FStar_Reflection_Derived_Lemmas.collect_app_ref - rel_xy1)) - (fun uu___3 -> - (fun uu___3 -> - match uu___3 with - | (rel, xy) -> - if - (FStar_List_Tot_Base.length - xy) - >= - (Prims.of_int (2)) - then - (match - ((FStar_List_Tot_Base.index + (fun uu___3 -> + match uu___3 with + | (rel, xy) -> + if + (FStar_List_Tot_Base.length + xy) + >= + (Prims.of_int (2)) + then + Obj.magic + (Obj.repr + (match + ((FStar_List_Tot_Base.index xy ((FStar_List_Tot_Base.length xy) - (Prims.of_int (2)))), - (FStar_List_Tot_Base.index + (FStar_List_Tot_Base.index xy ((FStar_List_Tot_Base.length xy) - Prims.int_one))) - with - | ((lhs, - FStar_Reflection_Data.Q_Explicit), - (rhs, - FStar_Reflection_Data.Q_Explicit)) - -> - Obj.magic + with + | + ((lhs, + FStar_Reflection_Data.Q_Explicit), + (rhs, + FStar_Reflection_Data.Q_Explicit)) + -> + Obj.repr (canon_lhs_rhs eq m lhs rhs) - | uu___4 -> - Obj.magic + | + uu___4 -> + Obj.repr + (FStar_Tactics_Derived.fail + "Goal should have been an application of a binary relation to 2 explicit arguments"))) + else + Obj.magic + (Obj.repr (FStar_Tactics_Derived.fail - "Goal should have been an application of a binary relation to 2 explicit arguments")) - else - Obj.magic - (FStar_Tactics_Derived.fail - "Goal should have been an application of a binary relation to n implicit and 2 explicit arguments")) - uu___3)) + "Goal should have been an application of a binary relation to n implicit and 2 explicit arguments"))) + uu___3))) | uu___2 -> Obj.magic - (FStar_Tactics_Derived.fail - "Goal should be squash applied to a binary relation"))) + (Obj.repr + (FStar_Tactics_Derived.fail + "Goal should be squash applied to a binary relation")))) uu___1))) uu___1))) uu___) let _ = FStar_Tactics_Native.register_tactic diff --git a/ocaml/fstar-lib/generated/FStar_Tactics_CanonCommSemiring.ml b/ocaml/fstar-lib/generated/FStar_Tactics_CanonCommSemiring.ml index a7843079ede..547fce6978e 100644 --- a/ocaml/fstar-lib/generated/FStar_Tactics_CanonCommSemiring.ml +++ b/ocaml/fstar-lib/generated/FStar_Tactics_CanonCommSemiring.ml @@ -2136,29 +2136,31 @@ let canon_semiring_aux : t), t1, t2) -> Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.CanonCommSemiring.fst" - (Prims.of_int (1634)) - (Prims.of_int (9)) - (Prims.of_int (1634)) - (Prims.of_int (21))) - (Prims.mk_range - "FStar.Tactics.CanonCommSemiring.fst" - (Prims.of_int (1634)) - (Prims.of_int (6)) - (Prims.of_int (1671)) - (Prims.of_int (73))) - (Obj.magic - (term_eq t - ta)) - (fun uu___3 - -> + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.CanonCommSemiring.fst" + (Prims.of_int (1634)) + (Prims.of_int (9)) + (Prims.of_int (1634)) + (Prims.of_int (21))) + (Prims.mk_range + "FStar.Tactics.CanonCommSemiring.fst" + (Prims.of_int (1634)) + (Prims.of_int (6)) + (Prims.of_int (1671)) + (Prims.of_int (73))) + (Obj.magic + (term_eq + t ta)) (fun + uu___3 -> + (fun uu___3 -> if uu___3 then Obj.magic + (Obj.repr (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.CanonCommSemiring.fst" @@ -2191,6 +2193,7 @@ let canon_semiring_aux : (e1::e2::[], vm) -> Obj.magic + (Obj.repr (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.CanonCommSemiring.fst" @@ -2429,22 +2432,25 @@ let canon_semiring_aux : uu___5))) uu___5))) uu___5))) - uu___5)) + uu___5))) | uu___5 -> Obj.magic + (Obj.repr (FStar_Tactics_Derived.fail - "Unexpected")) - uu___4)) + "Unexpected"))) + uu___4))) else Obj.magic + (Obj.repr (FStar_Tactics_Derived.fail - "Found equality, but terms do not have the expected type")) - uu___3)) + "Found equality, but terms do not have the expected type"))) + uu___3))) | uu___3 -> Obj.magic - (FStar_Tactics_Derived.fail - "Goal should be an equality")) + (Obj.repr + (FStar_Tactics_Derived.fail + "Goal should be an equality"))) uu___2))) uu___2))) uu___1)) let canon_semiring : 'a . 'a cr -> (unit, unit) FStar_Tactics_Effect.tac_repr diff --git a/ocaml/fstar-lib/generated/FStar_Tactics_CanonMonoid.ml b/ocaml/fstar-lib/generated/FStar_Tactics_CanonMonoid.ml index 648100aa857..be918e0117b 100644 --- a/ocaml/fstar-lib/generated/FStar_Tactics_CanonMonoid.ml +++ b/ocaml/fstar-lib/generated/FStar_Tactics_CanonMonoid.ml @@ -448,85 +448,89 @@ let canon_monoid : me2) -> Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.CanonMonoid.fst" - (Prims.of_int (110)) - (Prims.of_int (9)) - (Prims.of_int (110)) - (Prims.of_int (32))) - (Prims.mk_range - "FStar.Tactics.CanonMonoid.fst" - (Prims.of_int (110)) - (Prims.of_int (6)) - (Prims.of_int (119)) - (Prims.of_int (69))) - (Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.CanonMonoid.fst" - (Prims.of_int (110)) - (Prims.of_int (23)) - (Prims.of_int (110)) - (Prims.of_int (32))) - (Prims.mk_range - "FStar.Tactics.CanonMonoid.fst" - (Prims.of_int (110)) - (Prims.of_int (9)) - (Prims.of_int (110)) - (Prims.of_int (32))) - (FStar_Tactics_Effect.lift_div_tac + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.CanonMonoid.fst" + (Prims.of_int (110)) + (Prims.of_int (9)) + (Prims.of_int (110)) + (Prims.of_int (32))) + (Prims.mk_range + "FStar.Tactics.CanonMonoid.fst" + (Prims.of_int (110)) + (Prims.of_int (6)) + (Prims.of_int (119)) + (Prims.of_int (69))) + (Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.CanonMonoid.fst" + (Prims.of_int (110)) + (Prims.of_int (23)) + (Prims.of_int (110)) + (Prims.of_int (32))) + (Prims.mk_range + "FStar.Tactics.CanonMonoid.fst" + (Prims.of_int (110)) + (Prims.of_int (9)) + (Prims.of_int (110)) + (Prims.of_int (32))) + (FStar_Tactics_Effect.lift_div_tac + (fun uu___2 -> + (fun uu___2 -> + Obj.magic + (failwith + "Cannot evaluate open quotation at runtime")) + uu___2)) (fun uu___2 -> (fun uu___2 -> Obj.magic - (failwith - "Cannot evaluate open quotation at runtime")) - uu___2)) - (fun uu___2 -> - (fun uu___2 -> - Obj.magic - (FStar_Tactics_Builtins.term_eq_old - t uu___2)) uu___2))) - (fun uu___2 -> + (FStar_Tactics_Builtins.term_eq_old + t uu___2)) + uu___2))) (fun uu___2 -> - if uu___2 - then - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.CanonMonoid.fst" - (Prims.of_int (111)) - (Prims.of_int (17)) - (Prims.of_int (111)) - (Prims.of_int (34))) - (Prims.mk_range - "FStar.Tactics.CanonMonoid.fst" - (Prims.of_int (112)) - (Prims.of_int (8)) - (Prims.of_int (118)) - (Prims.of_int (51))) - (Obj.magic - (reification m me1)) - (fun uu___3 -> - (fun r1 -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range + (fun uu___2 -> + if uu___2 + then + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.CanonMonoid.fst" + (Prims.of_int (111)) + (Prims.of_int (17)) + (Prims.of_int (111)) + (Prims.of_int (34))) + (Prims.mk_range + "FStar.Tactics.CanonMonoid.fst" + (Prims.of_int (112)) + (Prims.of_int (8)) + (Prims.of_int (118)) + (Prims.of_int (51))) + (Obj.magic + (reification + m me1)) + (fun uu___3 -> + (fun r1 -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "FStar.Tactics.CanonMonoid.fst" (Prims.of_int (112)) (Prims.of_int (17)) (Prims.of_int (112)) (Prims.of_int (34))) - (Prims.mk_range + (Prims.mk_range "FStar.Tactics.CanonMonoid.fst" (Prims.of_int (113)) (Prims.of_int (8)) (Prims.of_int (118)) (Prims.of_int (51))) - (Obj.magic + (Obj.magic (reification m me2)) - (fun + (fun uu___3 -> (fun r2 -> @@ -617,14 +621,16 @@ let canon_monoid : uu___4))) uu___3))) uu___3))) - uu___3)) - else - Obj.magic - (FStar_Tactics_Derived.fail - "Goal should be an equality at the right monoid type")) - uu___2)) + uu___3))) + else + Obj.magic + (Obj.repr + (FStar_Tactics_Derived.fail + "Goal should be an equality at the right monoid type"))) + uu___2))) | uu___2 -> Obj.magic - (FStar_Tactics_Derived.fail - "Goal should be an equality")) + (Obj.repr + (FStar_Tactics_Derived.fail + "Goal should be an equality"))) uu___1))) uu___1))) uu___) \ No newline at end of file diff --git a/ocaml/fstar-lib/generated/FStar_Tactics_Derived.ml b/ocaml/fstar-lib/generated/FStar_Tactics_Derived.ml index 4c8008fdce0..c8d79f468ed 100644 --- a/ocaml/fstar-lib/generated/FStar_Tactics_Derived.ml +++ b/ocaml/fstar-lib/generated/FStar_Tactics_Derived.ml @@ -36,7 +36,7 @@ let (goals : (Prims.of_int (42)) (Prims.of_int (40)) (Prims.of_int (50))) (Prims.mk_range "FStar.Tactics.Derived.fst" (Prims.of_int (40)) (Prims.of_int (33)) (Prims.of_int (40)) (Prims.of_int (50))) - (Obj.magic (FStar_Tactics_Effect.get ())) + (FStar_Tactics_Effect.get ()) (fun uu___1 -> FStar_Tactics_Effect.lift_div_tac (fun uu___2 -> FStar_Tactics_Types.goals_of uu___1)) @@ -50,11 +50,11 @@ let (smt_goals : (Prims.of_int (50)) (Prims.of_int (41)) (Prims.of_int (58))) (Prims.mk_range "FStar.Tactics.Derived.fst" (Prims.of_int (41)) (Prims.of_int (37)) (Prims.of_int (41)) (Prims.of_int (58))) - (Obj.magic (FStar_Tactics_Effect.get ())) + (FStar_Tactics_Effect.get ()) (fun uu___1 -> FStar_Tactics_Effect.lift_div_tac (fun uu___2 -> FStar_Tactics_Types.smt_goals_of uu___1)) -let fail : 'a . Prims.string -> ('a, unit) FStar_Tactics_Effect.tac_repr = +let fail : 'a . Prims.string -> ('a, Obj.t) FStar_Tactics_Effect.tac_repr = fun m -> FStar_Tactics_Effect.raise (FStar_Tactics_Common.TacticFailure m) let fail_silently : 'a . Prims.string -> ('a, unit) FStar_Tactics_Effect.tac_repr = @@ -66,10 +66,7 @@ let fail_silently : (Prims.of_int (4)) (Prims.of_int (50)) (Prims.of_int (30))) (Obj.magic (FStar_Tactics_Builtins.set_urgency Prims.int_zero)) (fun uu___ -> - (fun uu___ -> - Obj.magic - (FStar_Tactics_Effect.raise - (FStar_Tactics_Common.TacticFailure m))) uu___) + FStar_Tactics_Effect.raise (FStar_Tactics_Common.TacticFailure m)) let (_cur_goal : unit -> (FStar_Tactics_Types.goal, unit) FStar_Tactics_Effect.tac_repr) = fun uu___ -> @@ -80,14 +77,9 @@ let (_cur_goal : (Prims.of_int (4)) (Prims.of_int (56)) (Prims.of_int (15))) (Obj.magic (goals ())) (fun uu___1 -> - (fun uu___1 -> - match uu___1 with - | [] -> Obj.magic (Obj.repr (fail "no more goals")) - | g::uu___2 -> - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac (fun uu___3 -> g)))) - uu___1) + match uu___1 with + | [] -> fail "no more goals" + | g::uu___2 -> FStar_Tactics_Effect.lift_div_tac (fun uu___3 -> g)) let (cur_env : unit -> (FStar_Reflection_Types.env, unit) FStar_Tactics_Effect.tac_repr) = fun uu___ -> @@ -139,7 +131,7 @@ let (cur_goal_safe : (Prims.of_int (18)) (Prims.of_int (72)) (Prims.of_int (26))) (Prims.mk_range "FStar.Tactics.Derived.fst" (Prims.of_int (72)) (Prims.of_int (9)) (Prims.of_int (72)) (Prims.of_int (26))) - (Obj.magic (FStar_Tactics_Effect.get ())) + (FStar_Tactics_Effect.get ()) (fun uu___1 -> FStar_Tactics_Effect.lift_div_tac (fun uu___2 -> FStar_Tactics_Types.goals_of uu___1)))) @@ -275,14 +267,16 @@ let (trivial : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = match uu___2 with | FStar_Reflection_Formula.True_ -> Obj.magic - (exact - (FStar_Reflection_Builtins.pack_ln - (FStar_Reflection_Data.Tv_Const - FStar_Reflection_Data.C_Unit))) + (Obj.repr + (exact + (FStar_Reflection_Builtins.pack_ln + (FStar_Reflection_Data.Tv_Const + FStar_Reflection_Data.C_Unit)))) | uu___3 -> Obj.magic - (FStar_Tactics_Effect.raise - Goal_not_trivial)) uu___2))) + (Obj.repr + (FStar_Tactics_Effect.raise + Goal_not_trivial))) uu___2))) uu___2))) uu___1) let (dismiss : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = fun uu___ -> @@ -295,8 +289,9 @@ let (dismiss : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = (fun uu___1 -> (fun uu___1 -> match uu___1 with - | [] -> Obj.magic (fail "dismiss: no more goals") - | uu___2::gs -> Obj.magic (FStar_Tactics_Builtins.set_goals gs)) + | [] -> Obj.magic (Obj.repr (fail "dismiss: no more goals")) + | uu___2::gs -> + Obj.magic (Obj.repr (FStar_Tactics_Builtins.set_goals gs))) uu___1) let (flip : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = fun uu___ -> @@ -320,13 +315,17 @@ let (flip : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = (fun uu___1 -> (fun uu___1 -> match uu___1 with - | [] -> Obj.magic (fail "flip: less than two goals") + | [] -> + Obj.magic + (Obj.repr (fail "flip: less than two goals")) | uu___2::[] -> - Obj.magic (fail "flip: less than two goals") + Obj.magic + (Obj.repr (fail "flip: less than two goals")) | g1::g2::gs1 -> Obj.magic - (FStar_Tactics_Builtins.set_goals (g2 :: g1 :: - gs1))) uu___1))) uu___1) + (Obj.repr + (FStar_Tactics_Builtins.set_goals (g2 :: g1 + :: gs1)))) uu___1))) uu___1) let (qed : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = fun uu___ -> FStar_Tactics_Effect.tac_bind @@ -336,13 +335,9 @@ let (qed : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = (Prims.of_int (4)) (Prims.of_int (130)) (Prims.of_int (32))) (Obj.magic (goals ())) (fun uu___1 -> - (fun uu___1 -> - match uu___1 with - | [] -> - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac (fun uu___2 -> ()))) - | uu___2 -> Obj.magic (Obj.repr (fail "qed: not done!"))) uu___1) + match uu___1 with + | [] -> FStar_Tactics_Effect.lift_div_tac (fun uu___2 -> ()) + | uu___2 -> fail "qed: not done!") let (debug : Prims.string -> (unit, unit) FStar_Tactics_Effect.tac_repr) = fun m -> FStar_Tactics_Effect.tac_bind @@ -391,22 +386,24 @@ let (smt : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = (fun uu___1 -> (fun uu___1 -> match uu___1 with - | ([], uu___2) -> Obj.magic (fail "smt: no active goals") + | ([], uu___2) -> + Obj.magic (Obj.repr (fail "smt: no active goals")) | (g::gs, gs') -> Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "FStar.Tactics.Derived.fst" - (Prims.of_int (147)) (Prims.of_int (8)) - (Prims.of_int (147)) (Prims.of_int (20))) - (Prims.mk_range "FStar.Tactics.Derived.fst" - (Prims.of_int (148)) (Prims.of_int (8)) - (Prims.of_int (148)) (Prims.of_int (32))) - (Obj.magic (FStar_Tactics_Builtins.set_goals gs)) - (fun uu___2 -> + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "FStar.Tactics.Derived.fst" + (Prims.of_int (147)) (Prims.of_int (8)) + (Prims.of_int (147)) (Prims.of_int (20))) + (Prims.mk_range "FStar.Tactics.Derived.fst" + (Prims.of_int (148)) (Prims.of_int (8)) + (Prims.of_int (148)) (Prims.of_int (32))) + (Obj.magic (FStar_Tactics_Builtins.set_goals gs)) (fun uu___2 -> - Obj.magic - (FStar_Tactics_Builtins.set_smt_goals (g :: gs'))) - uu___2))) uu___1) + (fun uu___2 -> + Obj.magic + (FStar_Tactics_Builtins.set_smt_goals (g :: + gs'))) uu___2)))) uu___1) let (idtac : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = fun uu___ -> (fun uu___ -> @@ -425,9 +422,10 @@ let (later : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = match uu___1 with | g::gs -> Obj.magic - (FStar_Tactics_Builtins.set_goals - (FStar_List_Tot_Base.op_At gs [g])) - | uu___2 -> Obj.magic (fail "later: no goals")) uu___1) + (Obj.repr + (FStar_Tactics_Builtins.set_goals + (FStar_List_Tot_Base.op_At gs [g]))) + | uu___2 -> Obj.magic (Obj.repr (fail "later: no goals"))) uu___1) let (apply : FStar_Reflection_Types.term -> (unit, unit) FStar_Tactics_Effect.tac_repr) = fun t -> FStar_Tactics_Builtins.t_apply true false false t @@ -536,29 +534,21 @@ let (topdown_rewrite : (Prims.of_int (254)) (Prims.of_int (10))) (match i with | uu___2 when uu___2 = Prims.int_zero -> - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___3 -> - FStar_Tactics_Types.Continue))) + FStar_Tactics_Effect.lift_div_tac + (fun uu___3 -> + FStar_Tactics_Types.Continue) | uu___2 when uu___2 = Prims.int_one -> - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___3 -> - FStar_Tactics_Types.Skip))) + FStar_Tactics_Effect.lift_div_tac + (fun uu___3 -> + FStar_Tactics_Types.Skip) | uu___2 when uu___2 = (Prims.of_int (2)) -> - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___3 -> - FStar_Tactics_Types.Abort))) + FStar_Tactics_Effect.lift_div_tac + (fun uu___3 -> + FStar_Tactics_Types.Abort) | uu___2 -> - Obj.magic - (Obj.repr - (fail - "topdown_rewrite: bad value from ctrl"))) + fail + "topdown_rewrite: bad value from ctrl") (fun f -> FStar_Tactics_Effect.lift_div_tac (fun uu___2 -> (b, f))))) uu___1))) @@ -681,11 +671,8 @@ let divide : (Prims.mk_range "FStar.Tactics.Derived.fst" (Prims.of_int (289)) (Prims.of_int (4)) (Prims.of_int (301)) (Prims.of_int (10))) (if n < Prims.int_zero - then Obj.magic (Obj.repr (fail "divide: negative n")) - else - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac (fun uu___1 -> ())))) + then fail "divide: negative n" + else FStar_Tactics_Effect.lift_div_tac (fun uu___1 -> ())) (fun uu___ -> (fun uu___ -> Obj.magic @@ -1106,84 +1093,91 @@ let focus : (fun uu___ -> (fun uu___ -> match uu___ with - | [] -> Obj.magic (fail "focus: no goals") + | [] -> Obj.magic (Obj.repr (fail "focus: no goals")) | g::gs -> Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "FStar.Tactics.Derived.fst" - (Prims.of_int (314)) (Prims.of_int (18)) - (Prims.of_int (314)) (Prims.of_int (30))) - (Prims.mk_range "FStar.Tactics.Derived.fst" - (Prims.of_int (315)) (Prims.of_int (8)) - (Prims.of_int (318)) (Prims.of_int (9))) - (Obj.magic (smt_goals ())) - (fun uu___1 -> - (fun sgs -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "FStar.Tactics.Derived.fst" - (Prims.of_int (315)) (Prims.of_int (8)) - (Prims.of_int (315)) (Prims.of_int (21))) - (Prims.mk_range "FStar.Tactics.Derived.fst" - (Prims.of_int (315)) (Prims.of_int (23)) - (Prims.of_int (318)) (Prims.of_int (9))) - (Obj.magic - (FStar_Tactics_Builtins.set_goals [g])) - (fun uu___1 -> + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "FStar.Tactics.Derived.fst" + (Prims.of_int (314)) (Prims.of_int (18)) + (Prims.of_int (314)) (Prims.of_int (30))) + (Prims.mk_range "FStar.Tactics.Derived.fst" + (Prims.of_int (315)) (Prims.of_int (8)) + (Prims.of_int (318)) (Prims.of_int (9))) + (Obj.magic (smt_goals ())) + (fun uu___1 -> + (fun sgs -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.Derived.fst" + (Prims.of_int (315)) (Prims.of_int (8)) + (Prims.of_int (315)) + (Prims.of_int (21))) + (Prims.mk_range + "FStar.Tactics.Derived.fst" + (Prims.of_int (315)) + (Prims.of_int (23)) + (Prims.of_int (318)) (Prims.of_int (9))) + (Obj.magic + (FStar_Tactics_Builtins.set_goals [g])) (fun uu___1 -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.Derived.fst" - (Prims.of_int (315)) - (Prims.of_int (23)) - (Prims.of_int (315)) - (Prims.of_int (39))) - (Prims.mk_range - "FStar.Tactics.Derived.fst" - (Prims.of_int (316)) - (Prims.of_int (8)) - (Prims.of_int (318)) - (Prims.of_int (9))) - (Obj.magic - (FStar_Tactics_Builtins.set_smt_goals - [])) - (fun uu___2 -> + (fun uu___1 -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.Derived.fst" + (Prims.of_int (315)) + (Prims.of_int (23)) + (Prims.of_int (315)) + (Prims.of_int (39))) + (Prims.mk_range + "FStar.Tactics.Derived.fst" + (Prims.of_int (316)) + (Prims.of_int (8)) + (Prims.of_int (318)) + (Prims.of_int (9))) + (Obj.magic + (FStar_Tactics_Builtins.set_smt_goals + [])) (fun uu___2 -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.Derived.fst" - (Prims.of_int (316)) - (Prims.of_int (16)) - (Prims.of_int (316)) - (Prims.of_int (20))) - (Prims.mk_range - "FStar.Tactics.Derived.fst" - (Prims.of_int (317)) - (Prims.of_int (8)) - (Prims.of_int (318)) - (Prims.of_int (9))) - (Obj.magic (t ())) - (fun uu___3 -> - (fun x -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range + (fun uu___2 -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.Derived.fst" + (Prims.of_int (316)) + (Prims.of_int (16)) + (Prims.of_int (316)) + (Prims.of_int (20))) + (Prims.mk_range + "FStar.Tactics.Derived.fst" + (Prims.of_int (317)) + (Prims.of_int (8)) + (Prims.of_int (318)) + (Prims.of_int (9))) + (Obj.magic (t ())) + (fun uu___3 -> + (fun x -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + ( + Prims.mk_range "FStar.Tactics.Derived.fst" (Prims.of_int (317)) (Prims.of_int (8)) (Prims.of_int (317)) (Prims.of_int (33))) - (Prims.mk_range + ( + Prims.mk_range "FStar.Tactics.Derived.fst" (Prims.of_int (317)) (Prims.of_int (35)) (Prims.of_int (318)) (Prims.of_int (9))) - (Obj.magic ( - FStar_Tactics_Effect.tac_bind + Obj.magic + (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.Derived.fst" (Prims.of_int (317)) @@ -1227,8 +1221,9 @@ let focus : (FStar_Tactics_Builtins.set_goals uu___3)) uu___3))) - (fun uu___3 - -> + ( + fun + uu___3 -> (fun uu___3 -> Obj.magic @@ -1299,8 +1294,9 @@ let focus : uu___5 -> x)))) uu___3))) - uu___3))) uu___2))) - uu___1))) uu___1))) uu___) + uu___3))) + uu___2))) uu___1))) uu___1)))) + uu___) let (dump1 : Prims.string -> (unit, unit) FStar_Tactics_Effect.tac_repr) = fun m -> focus (fun uu___ -> FStar_Tactics_Builtins.dump m) let rec mapAll : @@ -1786,10 +1782,8 @@ let (guard : Prims.bool -> (unit, unit) FStar_Tactics_Effect.tac_repr) = fun uu___ -> (fun b -> if Prims.op_Negation b - then Obj.magic (Obj.repr (fail "guard failed")) - else - Obj.magic - (Obj.repr (FStar_Tactics_Effect.lift_div_tac (fun uu___1 -> ())))) + then Obj.magic (fail "guard failed") + else Obj.magic (FStar_Tactics_Effect.lift_div_tac (fun uu___1 -> ()))) uu___ let try_with : 'a . @@ -1863,7 +1857,8 @@ let first : = fun ts -> FStar_List_Tot_Base.fold_right op_Less_Bar_Greater ts - (fun uu___ -> fail "no tactics to try") () + (fun uu___ -> (fun uu___ -> Obj.magic (fail "no tactics to try")) uu___) + () let rec repeat : 'a . (unit -> ('a, unit) FStar_Tactics_Effect.tac_repr) -> @@ -2141,8 +2136,9 @@ let (skip_guard : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = (Obj.magic (is_guard ())) (fun uu___1 -> (fun uu___1 -> - if uu___1 then Obj.magic (smt ()) else Obj.magic (fail "")) - uu___1) + if uu___1 + then Obj.magic (Obj.repr (smt ())) + else Obj.magic (Obj.repr (fail ""))) uu___1) let (guards_to_smt : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = fun uu___ -> FStar_Tactics_Effect.tac_bind @@ -2387,50 +2383,57 @@ let rec (__assumption_aux : FStar_Reflection_Types.binders -> (unit, unit) FStar_Tactics_Effect.tac_repr) = - fun bs -> - match bs with - | [] -> fail "no assumption matches goal" - | b::bs1 -> - FStar_Tactics_Effect.tac_bind - (Prims.mk_range "FStar.Tactics.Derived.fst" (Prims.of_int (539)) - (Prims.of_int (16)) (Prims.of_int (539)) (Prims.of_int (32))) - (Prims.mk_range "FStar.Tactics.Derived.fst" (Prims.of_int (540)) - (Prims.of_int (8)) (Prims.of_int (543)) (Prims.of_int (27))) - (Obj.magic (binder_to_term b)) - (fun uu___ -> - (fun t -> - Obj.magic - (try_with (fun uu___ -> match () with | () -> exact t) - (fun uu___ -> - try_with - (fun uu___1 -> - match () with - | () -> - FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.Derived.fst" - (Prims.of_int (541)) - (Prims.of_int (13)) - (Prims.of_int (541)) - (Prims.of_int (48))) - (Prims.mk_range - "FStar.Tactics.Derived.fst" - (Prims.of_int (542)) - (Prims.of_int (13)) - (Prims.of_int (542)) - (Prims.of_int (20))) - (Obj.magic - (apply - (FStar_Reflection_Builtins.pack_ln - (FStar_Reflection_Data.Tv_FVar - (FStar_Reflection_Builtins.pack_fv - ["FStar"; - "Squash"; - "return_squash"]))))) - (fun uu___2 -> - (fun uu___2 -> Obj.magic (exact t)) - uu___2)) - (fun uu___1 -> __assumption_aux bs1)))) uu___) + fun uu___ -> + (fun bs -> + match bs with + | [] -> Obj.magic (Obj.repr (fail "no assumption matches goal")) + | b::bs1 -> + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "FStar.Tactics.Derived.fst" + (Prims.of_int (539)) (Prims.of_int (16)) + (Prims.of_int (539)) (Prims.of_int (32))) + (Prims.mk_range "FStar.Tactics.Derived.fst" + (Prims.of_int (540)) (Prims.of_int (8)) + (Prims.of_int (543)) (Prims.of_int (27))) + (Obj.magic (binder_to_term b)) + (fun uu___ -> + (fun t -> + Obj.magic + (try_with + (fun uu___ -> match () with | () -> exact t) + (fun uu___ -> + try_with + (fun uu___1 -> + match () with + | () -> + FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.Derived.fst" + (Prims.of_int (541)) + (Prims.of_int (13)) + (Prims.of_int (541)) + (Prims.of_int (48))) + (Prims.mk_range + "FStar.Tactics.Derived.fst" + (Prims.of_int (542)) + (Prims.of_int (13)) + (Prims.of_int (542)) + (Prims.of_int (20))) + (Obj.magic + (apply + (FStar_Reflection_Builtins.pack_ln + (FStar_Reflection_Data.Tv_FVar + (FStar_Reflection_Builtins.pack_fv + ["FStar"; + "Squash"; + "return_squash"]))))) + (fun uu___2 -> + (fun uu___2 -> + Obj.magic (exact t)) uu___2)) + (fun uu___1 -> __assumption_aux bs1)))) + uu___)))) uu___ let (assumption : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = fun uu___ -> FStar_Tactics_Effect.tac_bind @@ -2522,7 +2525,8 @@ let (rewrite' : (fun uu___2 -> Obj.magic (FStar_Tactics_Builtins.rewrite b)) uu___2))) uu___1))) - (fun uu___ -> fail "rewrite' failed") () + (fun uu___ -> (fun uu___ -> Obj.magic (fail "rewrite' failed")) uu___) + () let rec (try_rewrite_equality : FStar_Reflection_Types.term -> FStar_Reflection_Types.binders -> @@ -2638,23 +2642,26 @@ let (unfold_def : match uu___ with | FStar_Reflection_Data.Tv_FVar fv -> Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "FStar.Tactics.Derived.fst" - (Prims.of_int (601)) (Prims.of_int (16)) - (Prims.of_int (601)) (Prims.of_int (42))) - (Prims.mk_range "FStar.Tactics.Derived.fst" - (Prims.of_int (602)) (Prims.of_int (8)) - (Prims.of_int (602)) (Prims.of_int (30))) - (FStar_Tactics_Effect.lift_div_tac + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "FStar.Tactics.Derived.fst" + (Prims.of_int (601)) (Prims.of_int (16)) + (Prims.of_int (601)) (Prims.of_int (42))) + (Prims.mk_range "FStar.Tactics.Derived.fst" + (Prims.of_int (602)) (Prims.of_int (8)) + (Prims.of_int (602)) (Prims.of_int (30))) + (FStar_Tactics_Effect.lift_div_tac + (fun uu___1 -> + FStar_Reflection_Builtins.implode_qn + (FStar_Reflection_Builtins.inspect_fv fv))) (fun uu___1 -> - FStar_Reflection_Builtins.implode_qn - (FStar_Reflection_Builtins.inspect_fv fv))) - (fun uu___1 -> - (fun n -> - Obj.magic - (FStar_Tactics_Builtins.norm - [FStar_Pervasives.delta_fully [n]])) uu___1)) - | uu___1 -> Obj.magic (fail "unfold_def: term is not a fv")) + (fun n -> + Obj.magic + (FStar_Tactics_Builtins.norm + [FStar_Pervasives.delta_fully [n]])) + uu___1))) + | uu___1 -> + Obj.magic (Obj.repr (fail "unfold_def: term is not a fv"))) uu___) let (l_to_r : FStar_Reflection_Types.term Prims.list -> @@ -2885,79 +2892,83 @@ let (grewrite_eq : | FStar_Reflection_Formula.Comp (FStar_Reflection_Formula.Eq uu___3, l, r) -> Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.Derived.fst" - (Prims.of_int (655)) - (Prims.of_int (6)) - (Prims.of_int (655)) - (Prims.of_int (18))) - (Prims.mk_range - "FStar.Tactics.Derived.fst" - (Prims.of_int (656)) - (Prims.of_int (6)) - (Prims.of_int (657)) - (Prims.of_int (56))) - (Obj.magic (grewrite l r)) - (fun uu___4 -> + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.Derived.fst" + (Prims.of_int (655)) + (Prims.of_int (6)) + (Prims.of_int (655)) + (Prims.of_int (18))) + (Prims.mk_range + "FStar.Tactics.Derived.fst" + (Prims.of_int (656)) + (Prims.of_int (6)) + (Prims.of_int (657)) + (Prims.of_int (56))) + (Obj.magic (grewrite l r)) (fun uu___4 -> - Obj.magic - (iseq - [idtac; - (fun uu___5 -> - FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.Derived.fst" - (Prims.of_int (656)) - (Prims.of_int (30)) - (Prims.of_int (656)) - (Prims.of_int (55))) - (Prims.mk_range - "FStar.Tactics.Derived.fst" - (Prims.of_int (657)) - (Prims.of_int (30)) - (Prims.of_int (657)) - (Prims.of_int (54))) - (Obj.magic - (apply_lemma - (FStar_Reflection_Builtins.pack_ln - (FStar_Reflection_Data.Tv_FVar - (FStar_Reflection_Builtins.pack_fv - ["FStar"; - "Tactics"; - "Derived"; - "__un_sq_eq"]))))) - (fun uu___6 -> + (fun uu___4 -> + Obj.magic + (iseq + [idtac; + (fun uu___5 -> + FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.Derived.fst" + (Prims.of_int (656)) + (Prims.of_int (30)) + (Prims.of_int (656)) + (Prims.of_int (55))) + (Prims.mk_range + "FStar.Tactics.Derived.fst" + (Prims.of_int (657)) + (Prims.of_int (30)) + (Prims.of_int (657)) + (Prims.of_int (54))) + (Obj.magic + (apply_lemma + (FStar_Reflection_Builtins.pack_ln + (FStar_Reflection_Data.Tv_FVar + (FStar_Reflection_Builtins.pack_fv + ["FStar"; + "Tactics"; + "Derived"; + "__un_sq_eq"]))))) (fun uu___6 -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.Derived.fst" - (Prims.of_int (657)) - (Prims.of_int (36)) - (Prims.of_int (657)) - (Prims.of_int (54))) - (Prims.mk_range - "FStar.Tactics.Derived.fst" - (Prims.of_int (657)) - (Prims.of_int (30)) - (Prims.of_int (657)) - (Prims.of_int (54))) - (Obj.magic - (binder_to_term + (fun uu___6 -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.Derived.fst" + (Prims.of_int (657)) + (Prims.of_int (36)) + (Prims.of_int (657)) + (Prims.of_int (54))) + (Prims.mk_range + "FStar.Tactics.Derived.fst" + (Prims.of_int (657)) + (Prims.of_int (30)) + (Prims.of_int (657)) + (Prims.of_int (54))) + (Obj.magic + (binder_to_term b)) - (fun uu___7 -> (fun uu___7 -> + (fun + uu___7 -> Obj.magic (exact uu___7)) uu___7))) - uu___6))])) uu___4)) + uu___6))])) + uu___4))) | uu___3 -> Obj.magic - (fail - "grewrite_eq: binder type is not an equality")) + (Obj.repr + (fail + "grewrite_eq: binder type is not an equality"))) uu___2))) uu___) let rec (apply_squash_or_lem : Prims.nat -> @@ -2989,101 +3000,121 @@ let rec (apply_squash_or_lem : (fun uu___1 -> try_with (fun uu___2 -> match () with | () -> apply_lemma t) (fun uu___2 -> - if d <= Prims.int_zero - then fail "mapply: out of fuel" - else - FStar_Tactics_Effect.tac_bind - (Prims.mk_range "FStar.Tactics.Derived.fst" - (Prims.of_int (688)) (Prims.of_int (13)) - (Prims.of_int (688)) (Prims.of_int (30))) - (Prims.mk_range "FStar.Tactics.Derived.fst" - (Prims.of_int (689)) (Prims.of_int (4)) - (Prims.of_int (737)) (Prims.of_int (41))) - (Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "FStar.Tactics.Derived.fst" - (Prims.of_int (688)) (Prims.of_int (16)) - (Prims.of_int (688)) (Prims.of_int (28))) - (Prims.mk_range "FStar.Tactics.Derived.fst" - (Prims.of_int (688)) (Prims.of_int (13)) - (Prims.of_int (688)) (Prims.of_int (30))) - (Obj.magic (cur_env ())) - (fun uu___4 -> + (fun uu___2 -> + if d <= Prims.int_zero + then + Obj.magic (Obj.repr (fail "mapply: out of fuel")) + else + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "FStar.Tactics.Derived.fst" + (Prims.of_int (688)) (Prims.of_int (13)) + (Prims.of_int (688)) (Prims.of_int (30))) + (Prims.mk_range "FStar.Tactics.Derived.fst" + (Prims.of_int (689)) (Prims.of_int (4)) + (Prims.of_int (737)) (Prims.of_int (41))) + (Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.Derived.fst" + (Prims.of_int (688)) + (Prims.of_int (16)) + (Prims.of_int (688)) + (Prims.of_int (28))) + (Prims.mk_range + "FStar.Tactics.Derived.fst" + (Prims.of_int (688)) + (Prims.of_int (13)) + (Prims.of_int (688)) + (Prims.of_int (30))) + (Obj.magic (cur_env ())) + (fun uu___4 -> + (fun uu___4 -> + Obj.magic + (FStar_Tactics_Builtins.tc + uu___4 t)) uu___4))) (fun uu___4 -> - Obj.magic - (FStar_Tactics_Builtins.tc uu___4 t)) - uu___4))) - (fun uu___4 -> - (fun ty -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.Derived.fst" - (Prims.of_int (689)) - (Prims.of_int (17)) - (Prims.of_int (689)) - (Prims.of_int (31))) - (Prims.mk_range - "FStar.Tactics.Derived.fst" - (Prims.of_int (689)) - (Prims.of_int (4)) - (Prims.of_int (737)) - (Prims.of_int (41))) - (Obj.magic - (FStar_Tactics_SyntaxHelpers.collect_arr - ty)) - (fun uu___4 -> - (fun uu___4 -> - match uu___4 with - | (tys, c) -> - (match FStar_Reflection_Builtins.inspect_comp - c - with - | FStar_Reflection_Data.C_Lemma - (pre, post, uu___5) -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.Derived.fst" - (Prims.of_int (693)) - (Prims.of_int (18)) - (Prims.of_int (693)) - (Prims.of_int (32))) - (Prims.mk_range - "FStar.Tactics.Derived.fst" - (Prims.of_int (694)) - (Prims.of_int (7)) - (Prims.of_int (702)) - (Prims.of_int (41))) - (FStar_Tactics_Effect.lift_div_tac - (fun uu___6 -> - FStar_Reflection_Builtins.pack_ln - (FStar_Reflection_Data.Tv_App - (post, + (fun ty -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.Derived.fst" + (Prims.of_int (689)) + (Prims.of_int (17)) + (Prims.of_int (689)) + (Prims.of_int (31))) + (Prims.mk_range + "FStar.Tactics.Derived.fst" + (Prims.of_int (689)) + (Prims.of_int (4)) + (Prims.of_int (737)) + (Prims.of_int (41))) + (Obj.magic + (FStar_Tactics_SyntaxHelpers.collect_arr + ty)) + (fun uu___4 -> + (fun uu___4 -> + match uu___4 with + | (tys, c) -> + (match FStar_Reflection_Builtins.inspect_comp + c + with + | FStar_Reflection_Data.C_Lemma + (pre, post, + uu___5) + -> + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.tac_bind + ( + Prims.mk_range + "FStar.Tactics.Derived.fst" + (Prims.of_int (693)) + (Prims.of_int (18)) + (Prims.of_int (693)) + (Prims.of_int (32))) + ( + Prims.mk_range + "FStar.Tactics.Derived.fst" + (Prims.of_int (694)) + (Prims.of_int (7)) + (Prims.of_int (702)) + (Prims.of_int (41))) + ( + FStar_Tactics_Effect.lift_div_tac + (fun + uu___6 -> + FStar_Reflection_Builtins.pack_ln + (FStar_Reflection_Data.Tv_App + (post, ((FStar_Reflection_Builtins.pack_ln (FStar_Reflection_Data.Tv_Const FStar_Reflection_Data.C_Unit)), FStar_Reflection_Data.Q_Explicit))))) - (fun uu___6 -> - (fun post1 -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range + ( + fun + uu___6 -> + (fun + post1 -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "FStar.Tactics.Derived.fst" (Prims.of_int (694)) (Prims.of_int (18)) (Prims.of_int (694)) (Prims.of_int (35))) - (Prims.mk_range + (Prims.mk_range "FStar.Tactics.Derived.fst" (Prims.of_int (696)) (Prims.of_int (7)) (Prims.of_int (702)) (Prims.of_int (41))) - (Obj.magic + (Obj.magic (norm_term [] post1)) - (fun + (fun uu___6 -> (fun post2 -> @@ -3114,6 +3145,7 @@ let rec (apply_squash_or_lem : FStar_Reflection_Formula.Implies (p, q) -> Obj.magic + (Obj.repr (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.Derived.fst" @@ -3145,42 +3177,47 @@ let rec (apply_squash_or_lem : (d - Prims.int_one) t)) - uu___7)) + uu___7))) | uu___7 -> Obj.magic + (Obj.repr (fail - "mapply: can't apply (1)")) + "mapply: can't apply (1)"))) + uu___6))) uu___6))) uu___6))) - uu___6)) - | FStar_Reflection_Data.C_Total - rt -> - (match FStar_Reflection_Derived.unsquash_term - rt - with - | FStar_Pervasives_Native.Some - rt1 -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.Derived.fst" - (Prims.of_int (710)) - (Prims.of_int (18)) - (Prims.of_int (710)) - (Prims.of_int (33))) - (Prims.mk_range - "FStar.Tactics.Derived.fst" - (Prims.of_int (712)) - (Prims.of_int (9)) - (Prims.of_int (718)) - (Prims.of_int (43))) - (Obj.magic - (norm_term [] - rt1)) - (fun uu___5 -> - (fun rt2 -> - Obj.magic + | FStar_Reflection_Data.C_Total + rt -> + Obj.magic + (Obj.repr + (match + FStar_Reflection_Derived.unsquash_term + rt + with + | FStar_Pervasives_Native.Some + rt1 -> + FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.Derived.fst" + (Prims.of_int (710)) + (Prims.of_int (18)) + (Prims.of_int (710)) + (Prims.of_int (33))) + (Prims.mk_range + "FStar.Tactics.Derived.fst" + (Prims.of_int (712)) + (Prims.of_int (9)) + (Prims.of_int (718)) + (Prims.of_int (43))) + (Obj.magic + (norm_term + [] rt1)) + (fun + uu___5 -> + (fun rt2 + -> + Obj.magic (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.Derived.fst" @@ -3207,6 +3244,7 @@ let rec (apply_squash_or_lem : FStar_Reflection_Formula.Implies (p, q) -> Obj.magic + (Obj.repr (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.Derived.fst" @@ -3238,36 +3276,38 @@ let rec (apply_squash_or_lem : (d - Prims.int_one) t)) - uu___6)) + uu___6))) | uu___6 -> Obj.magic + (Obj.repr (fail - "mapply: can't apply (1)")) + "mapply: can't apply (1)"))) uu___5))) - uu___5)) - | FStar_Pervasives_Native.None - -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.Derived.fst" - (Prims.of_int (725)) - (Prims.of_int (18)) - (Prims.of_int (725)) - (Prims.of_int (33))) - (Prims.mk_range - "FStar.Tactics.Derived.fst" - (Prims.of_int (727)) - (Prims.of_int (9)) - (Prims.of_int (734)) - (Prims.of_int (20))) - (Obj.magic - (norm_term [] - rt)) - (fun uu___5 -> - (fun rt1 -> - Obj.magic + uu___5) + | FStar_Pervasives_Native.None + -> + FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.Derived.fst" + (Prims.of_int (725)) + (Prims.of_int (18)) + (Prims.of_int (725)) + (Prims.of_int (33))) + (Prims.mk_range + "FStar.Tactics.Derived.fst" + (Prims.of_int (727)) + (Prims.of_int (9)) + (Prims.of_int (734)) + (Prims.of_int (20))) + (Obj.magic + (norm_term + [] rt)) + (fun + uu___5 -> + (fun rt1 + -> + Obj.magic (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.Derived.fst" @@ -3358,12 +3398,13 @@ let rec (apply_squash_or_lem : (apply t)) uu___7))) uu___5))) - uu___5))) - | uu___5 -> - Obj.magic - (fail - "mapply: can't apply (2)"))) - uu___4))) uu___4)))) + uu___5))) + | uu___5 -> + Obj.magic + (Obj.repr + (fail + "mapply: can't apply (2)")))) + uu___4))) uu___4)))) uu___2))) let (mapply : FStar_Reflection_Types.term -> (unit, unit) FStar_Tactics_Effect.tac_repr) = fun t -> apply_squash_or_lem (Prims.of_int (10)) t @@ -3477,7 +3518,10 @@ let finish_by : (Prims.of_int (773)) (Prims.of_int (9))) (Obj.magic (or_else qed - (fun uu___ -> fail "finish_by: not finished"))) + (fun uu___ -> + (fun uu___ -> + Obj.magic (fail "finish_by: not finished")) + uu___))) (fun uu___ -> FStar_Tactics_Effect.lift_div_tac (fun uu___1 -> x)))) uu___) @@ -3632,11 +3676,12 @@ let (tlabel : Prims.string -> (unit, unit) FStar_Tactics_Effect.tac_repr) = (fun uu___ -> (fun uu___ -> match uu___ with - | [] -> Obj.magic (fail "tlabel: no goals") + | [] -> Obj.magic (Obj.repr (fail "tlabel: no goals")) | h::t -> Obj.magic - (FStar_Tactics_Builtins.set_goals - ((FStar_Tactics_Types.set_label l h) :: t))) uu___) + (Obj.repr + (FStar_Tactics_Builtins.set_goals + ((FStar_Tactics_Types.set_label l h) :: t)))) uu___) let (tlabel' : Prims.string -> (unit, unit) FStar_Tactics_Effect.tac_repr) = fun l -> FStar_Tactics_Effect.tac_bind @@ -3648,26 +3693,27 @@ let (tlabel' : Prims.string -> (unit, unit) FStar_Tactics_Effect.tac_repr) = (fun uu___ -> (fun uu___ -> match uu___ with - | [] -> Obj.magic (fail "tlabel': no goals") + | [] -> Obj.magic (Obj.repr (fail "tlabel': no goals")) | h::t -> Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "FStar.Tactics.Derived.fst" - (Prims.of_int (821)) (Prims.of_int (16)) - (Prims.of_int (821)) (Prims.of_int (45))) - (Prims.mk_range "FStar.Tactics.Derived.fst" - (Prims.of_int (822)) (Prims.of_int (8)) - (Prims.of_int (822)) (Prims.of_int (26))) - (FStar_Tactics_Effect.lift_div_tac + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "FStar.Tactics.Derived.fst" + (Prims.of_int (821)) (Prims.of_int (16)) + (Prims.of_int (821)) (Prims.of_int (45))) + (Prims.mk_range "FStar.Tactics.Derived.fst" + (Prims.of_int (822)) (Prims.of_int (8)) + (Prims.of_int (822)) (Prims.of_int (26))) + (FStar_Tactics_Effect.lift_div_tac + (fun uu___1 -> + FStar_Tactics_Types.set_label + (Prims.strcat l + (FStar_Tactics_Types.get_label h)) h)) (fun uu___1 -> - FStar_Tactics_Types.set_label - (Prims.strcat l - (FStar_Tactics_Types.get_label h)) h)) - (fun uu___1 -> - (fun h1 -> - Obj.magic - (FStar_Tactics_Builtins.set_goals (h1 :: t))) - uu___1))) uu___) + (fun h1 -> + Obj.magic + (FStar_Tactics_Builtins.set_goals (h1 :: t))) + uu___1)))) uu___) let (focus_all : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = fun uu___ -> FStar_Tactics_Effect.tac_bind @@ -3749,9 +3795,11 @@ let (bump_nth : Prims.pos -> (unit, unit) FStar_Tactics_Effect.tac_repr) = (fun uu___ -> match uu___ with | FStar_Pervasives_Native.None -> - Obj.magic (fail "bump_nth: not that many goals") + Obj.magic (Obj.repr (fail "bump_nth: not that many goals")) | FStar_Pervasives_Native.Some (h, t) -> - Obj.magic (FStar_Tactics_Builtins.set_goals (h :: t))) uu___) + Obj.magic + (Obj.repr (FStar_Tactics_Builtins.set_goals (h :: t)))) + uu___) let rec (destruct_list : FStar_Reflection_Types.term -> (FStar_Reflection_Types.term Prims.list, unit) @@ -3781,20 +3829,22 @@ let rec (destruct_list : (FStar_Reflection_Builtins.inspect_fv fv) = FStar_Reflection_Const.cons_qn then - FStar_Tactics_Effect.tac_bind - (Prims.mk_range "FStar.Tactics.Derived.fst" - (Prims.of_int (851)) (Prims.of_int (17)) - (Prims.of_int (851)) (Prims.of_int (33))) - (Prims.mk_range "FStar.Tactics.Derived.fst" - (Prims.of_int (851)) (Prims.of_int (14)) - (Prims.of_int (851)) (Prims.of_int (16))) - (Obj.magic (destruct_list a2)) - (fun uu___1 -> - FStar_Tactics_Effect.lift_div_tac - (fun uu___2 -> a1 :: uu___1)) + Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "FStar.Tactics.Derived.fst" + (Prims.of_int (851)) (Prims.of_int (17)) + (Prims.of_int (851)) (Prims.of_int (33))) + (Prims.mk_range "FStar.Tactics.Derived.fst" + (Prims.of_int (851)) (Prims.of_int (14)) + (Prims.of_int (851)) (Prims.of_int (16))) + (Obj.magic (destruct_list a2)) + (fun uu___1 -> + FStar_Tactics_Effect.lift_div_tac + (fun uu___2 -> a1 :: uu___1))) else - FStar_Tactics_Effect.raise - FStar_Tactics_Common.NotAListLiteral)) + Obj.repr + (FStar_Tactics_Effect.raise + FStar_Tactics_Common.NotAListLiteral))) | (FStar_Reflection_Data.Tv_FVar fv, (uu___1, FStar_Reflection_Data.Q_Implicit)::(a1, FStar_Reflection_Data.Q_Explicit):: @@ -3805,20 +3855,22 @@ let rec (destruct_list : (FStar_Reflection_Builtins.inspect_fv fv) = FStar_Reflection_Const.cons_qn then - FStar_Tactics_Effect.tac_bind - (Prims.mk_range "FStar.Tactics.Derived.fst" - (Prims.of_int (851)) (Prims.of_int (17)) - (Prims.of_int (851)) (Prims.of_int (33))) - (Prims.mk_range "FStar.Tactics.Derived.fst" - (Prims.of_int (851)) (Prims.of_int (14)) - (Prims.of_int (851)) (Prims.of_int (16))) - (Obj.magic (destruct_list a2)) - (fun uu___2 -> - FStar_Tactics_Effect.lift_div_tac - (fun uu___3 -> a1 :: uu___2)) + Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "FStar.Tactics.Derived.fst" + (Prims.of_int (851)) (Prims.of_int (17)) + (Prims.of_int (851)) (Prims.of_int (33))) + (Prims.mk_range "FStar.Tactics.Derived.fst" + (Prims.of_int (851)) (Prims.of_int (14)) + (Prims.of_int (851)) (Prims.of_int (16))) + (Obj.magic (destruct_list a2)) + (fun uu___2 -> + FStar_Tactics_Effect.lift_div_tac + (fun uu___3 -> a1 :: uu___2))) else - FStar_Tactics_Effect.raise - FStar_Tactics_Common.NotAListLiteral)) + Obj.repr + (FStar_Tactics_Effect.raise + FStar_Tactics_Common.NotAListLiteral))) | (FStar_Reflection_Data.Tv_FVar fv, uu___1) -> Obj.magic (Obj.repr @@ -3826,13 +3878,11 @@ let rec (destruct_list : (FStar_Reflection_Builtins.inspect_fv fv) = FStar_Reflection_Const.nil_qn then - Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___2 -> [])) + FStar_Tactics_Effect.lift_div_tac + (fun uu___2 -> []) else - Obj.repr - (FStar_Tactics_Effect.raise - FStar_Tactics_Common.NotAListLiteral))) + FStar_Tactics_Effect.raise + FStar_Tactics_Common.NotAListLiteral)) | uu___1 -> Obj.magic (Obj.repr @@ -3860,30 +3910,25 @@ let (get_match_body : (fun uu___1 -> (fun uu___1 -> match uu___1 with - | FStar_Pervasives_Native.None -> Obj.magic (fail "") + | FStar_Pervasives_Native.None -> Obj.magic (Obj.repr (fail "")) | FStar_Pervasives_Native.Some t -> Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "FStar.Tactics.Derived.fst" - (Prims.of_int (863)) (Prims.of_int (20)) - (Prims.of_int (863)) (Prims.of_int (39))) - (Prims.mk_range "FStar.Tactics.Derived.fst" - (Prims.of_int (863)) (Prims.of_int (14)) - (Prims.of_int (865)) (Prims.of_int (46))) - (Obj.magic (inspect_unascribe t)) - (fun uu___2 -> + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "FStar.Tactics.Derived.fst" + (Prims.of_int (863)) (Prims.of_int (20)) + (Prims.of_int (863)) (Prims.of_int (39))) + (Prims.mk_range "FStar.Tactics.Derived.fst" + (Prims.of_int (863)) (Prims.of_int (14)) + (Prims.of_int (865)) (Prims.of_int (46))) + (Obj.magic (inspect_unascribe t)) (fun uu___2 -> match uu___2 with | FStar_Reflection_Data.Tv_Match (sc, uu___3, uu___4) -> - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___5 -> sc))) - | uu___3 -> - Obj.magic - (Obj.repr (fail "Goal is not a match"))) - uu___2))) uu___1) + FStar_Tactics_Effect.lift_div_tac + (fun uu___5 -> sc) + | uu___3 -> fail "Goal is not a match")))) uu___1) let rec last : 'a . 'a Prims.list -> ('a, unit) FStar_Tactics_Effect.tac_repr = fun uu___ -> @@ -4016,26 +4061,17 @@ let (nth_binder : (Prims.of_int (896)) (Prims.of_int (2)) (Prims.of_int (898)) (Prims.of_int (15))) (if k < Prims.int_zero - then - Obj.magic - (Obj.repr (fail "not enough binders")) + then fail "not enough binders" else - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___1 -> k)))) - (fun uu___ -> - (fun k1 -> - match FStar_List_Tot_Base.nth bs k1 with - | FStar_Pervasives_Native.None -> - Obj.magic - (Obj.repr (fail "not enough binders")) - | FStar_Pervasives_Native.Some b -> - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___ -> b)))) uu___))) - uu___))) uu___) + FStar_Tactics_Effect.lift_div_tac + (fun uu___1 -> k)) + (fun k1 -> + match FStar_List_Tot_Base.nth bs k1 with + | FStar_Pervasives_Native.None -> + fail "not enough binders" + | FStar_Pervasives_Native.Some b -> + FStar_Tactics_Effect.lift_div_tac + (fun uu___ -> b)))) uu___))) uu___) exception Appears let (uu___is_Appears : Prims.exn -> Prims.bool) = fun projectee -> match projectee with | Appears -> true | uu___ -> false @@ -4077,15 +4113,10 @@ let (name_appears_in : (if (FStar_Reflection_Builtins.inspect_fv fv) = nm - then - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.raise Appears)) + then FStar_Tactics_Effect.raise Appears else - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___3 -> ())))) + FStar_Tactics_Effect.lift_div_tac + (fun uu___3 -> ())) (fun uu___2 -> FStar_Tactics_Effect.lift_div_tac (fun uu___3 -> t1))) @@ -4126,12 +4157,9 @@ let (name_appears_in : match uu___ with | Appears -> Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___1 -> true))) - | e -> - Obj.magic - (Obj.repr (FStar_Tactics_Effect.raise e))) + (FStar_Tactics_Effect.lift_div_tac + (fun uu___1 -> true)) + | e -> Obj.magic (FStar_Tactics_Effect.raise e)) uu___))) uu___) let rec (mk_abs : FStar_Reflection_Types.binder Prims.list -> diff --git a/ocaml/fstar-lib/generated/FStar_Tactics_Effect.ml b/ocaml/fstar-lib/generated/FStar_Tactics_Effect.ml index 226dda9374b..c20d9427b13 100644 --- a/ocaml/fstar-lib/generated/FStar_Tactics_Effect.ml +++ b/ocaml/fstar-lib/generated/FStar_Tactics_Effect.ml @@ -49,9 +49,9 @@ let __proj__TAC__item__bind = tac_bind type ('a, 'wp, 'uuuuu, 'uuuuu1) lift_div_tac_wp = 'wp let lift_div_tac : 'a 'wp . (unit -> 'a) -> ('a, 'wp) tac_repr = fun f -> fun ps -> FStar_Tactics_Result.Success ((f ()), ps) -let (get : unit -> (FStar_Tactics_Types.proofstate, unit) tac_repr) = +let (get : unit -> (FStar_Tactics_Types.proofstate, Obj.t) tac_repr) = fun uu___ -> fun ps -> FStar_Tactics_Result.Success (ps, ps) -let raise : 'a . Prims.exn -> ('a, unit) tac_repr = +let raise : 'a . Prims.exn -> ('a, Obj.t) tac_repr = fun e -> fun ps -> FStar_Tactics_Result.Failed (e, ps) type ('uuuuu, 'p) with_tactic = 'p let (rewrite_with_tactic : diff --git a/ocaml/fstar-lib/generated/FStar_Tactics_Logic.ml b/ocaml/fstar-lib/generated/FStar_Tactics_Logic.ml index 9848dd07536..453aaf00628 100644 --- a/ocaml/fstar-lib/generated/FStar_Tactics_Logic.ml +++ b/ocaml/fstar-lib/generated/FStar_Tactics_Logic.ml @@ -105,7 +105,10 @@ let (split : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = (FStar_Reflection_Data.Tv_FVar (FStar_Reflection_Builtins.pack_fv ["FStar"; "Tactics"; "Logic"; "split_lem"])))) - (fun uu___1 -> FStar_Tactics_Derived.fail "Could not split goal") + (fun uu___1 -> + (fun uu___1 -> + Obj.magic (FStar_Tactics_Derived.fail "Could not split goal")) + uu___1) let (implies_intro : unit -> (FStar_Reflection_Types.binder, unit) FStar_Tactics_Effect.tac_repr) = @@ -219,12 +222,9 @@ let (pose_lemma : (Prims.of_int (126)) (Prims.of_int (5))) (match FStar_Reflection_Builtins.inspect_comp c with | FStar_Reflection_Data.C_Lemma (pre, post, uu___) -> - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___1 -> (pre, post)))) - | uu___ -> - Obj.magic (Obj.repr (FStar_Tactics_Derived.fail ""))) + FStar_Tactics_Effect.lift_div_tac + (fun uu___1 -> (pre, post)) + | uu___ -> FStar_Tactics_Derived.fail "") (fun uu___ -> (fun uu___ -> match uu___ with @@ -729,74 +729,80 @@ let rec (simplify_eq_implication : match r with | FStar_Pervasives_Native.None -> Obj.magic - (FStar_Tactics_Derived.fail - "Not an equality implication") + (Obj.repr + (FStar_Tactics_Derived.fail + "Not an equality implication")) | FStar_Pervasives_Native.Some (uu___1, rhs) -> Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.Logic.fst" - (Prims.of_int (160)) - (Prims.of_int (19)) - (Prims.of_int (160)) - (Prims.of_int (35))) - (Prims.mk_range - "FStar.Tactics.Logic.fst" - (Prims.of_int (161)) - (Prims.of_int (8)) - (Prims.of_int (163)) - (Prims.of_int (37))) - (Obj.magic (implies_intro ())) - (fun uu___2 -> - (fun eq_h -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.Logic.fst" - (Prims.of_int (161)) - (Prims.of_int (8)) - (Prims.of_int (161)) - (Prims.of_int (20))) - (Prims.mk_range - "FStar.Tactics.Logic.fst" - (Prims.of_int (162)) - (Prims.of_int (8)) - (Prims.of_int (163)) - (Prims.of_int (37))) - (Obj.magic - (FStar_Tactics_Builtins.rewrite - eq_h)) - (fun uu___2 -> + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.Logic.fst" + (Prims.of_int (160)) + (Prims.of_int (19)) + (Prims.of_int (160)) + (Prims.of_int (35))) + (Prims.mk_range + "FStar.Tactics.Logic.fst" + (Prims.of_int (161)) + (Prims.of_int (8)) + (Prims.of_int (163)) + (Prims.of_int (37))) + (Obj.magic (implies_intro ())) + (fun uu___2 -> + (fun eq_h -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.Logic.fst" + (Prims.of_int (161)) + (Prims.of_int (8)) + (Prims.of_int (161)) + (Prims.of_int (20))) + (Prims.mk_range + "FStar.Tactics.Logic.fst" + (Prims.of_int (162)) + (Prims.of_int (8)) + (Prims.of_int (163)) + (Prims.of_int (37))) + (Obj.magic + (FStar_Tactics_Builtins.rewrite + eq_h)) (fun uu___2 -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range + (fun uu___2 -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + ( + Prims.mk_range "FStar.Tactics.Logic.fst" (Prims.of_int (162)) (Prims.of_int (8)) (Prims.of_int (162)) (Prims.of_int (20))) - (Prims.mk_range + ( + Prims.mk_range "FStar.Tactics.Logic.fst" (Prims.of_int (163)) (Prims.of_int (8)) (Prims.of_int (163)) (Prims.of_int (37))) - (Obj.magic ( - FStar_Tactics_Builtins.clear_top + Obj.magic + (FStar_Tactics_Builtins.clear_top ())) - (fun uu___3 - -> + ( + fun + uu___3 -> (fun uu___3 -> Obj.magic (visit simplify_eq_implication)) uu___3))) - uu___2))) uu___2))) - uu___1))) uu___1))) uu___1) + uu___2))) + uu___2)))) uu___1))) + uu___1))) uu___1) let (rewrite_all_equalities : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = fun uu___ -> visit simplify_eq_implication @@ -856,49 +862,52 @@ let rec (unfold_definition_and_simplify_eq : match r with | FStar_Pervasives_Native.None -> Obj.magic - (FStar_Tactics_Derived.fail - "Not an equality implication") + (Obj.repr + (FStar_Tactics_Derived.fail + "Not an equality implication")) | FStar_Pervasives_Native.Some (uu___2, rhs) -> Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.Logic.fst" - (Prims.of_int (180)) - (Prims.of_int (23)) - (Prims.of_int (180)) - (Prims.of_int (39))) - (Prims.mk_range - "FStar.Tactics.Logic.fst" - (Prims.of_int (181)) - (Prims.of_int (12)) - (Prims.of_int (183)) - (Prims.of_int (66))) - (Obj.magic - (implies_intro ())) - (fun uu___3 -> - (fun eq_h -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.Logic.fst" - (Prims.of_int (181)) - (Prims.of_int (12)) - (Prims.of_int (181)) - (Prims.of_int (24))) - (Prims.mk_range - "FStar.Tactics.Logic.fst" - (Prims.of_int (182)) - (Prims.of_int (12)) - (Prims.of_int (183)) - (Prims.of_int (66))) - (Obj.magic - (FStar_Tactics_Builtins.rewrite - eq_h)) - (fun uu___3 -> + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.Logic.fst" + (Prims.of_int (180)) + (Prims.of_int (23)) + (Prims.of_int (180)) + (Prims.of_int (39))) + (Prims.mk_range + "FStar.Tactics.Logic.fst" + (Prims.of_int (181)) + (Prims.of_int (12)) + (Prims.of_int (183)) + (Prims.of_int (66))) + (Obj.magic + (implies_intro ())) + (fun uu___3 -> + (fun eq_h -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.Logic.fst" + (Prims.of_int (181)) + (Prims.of_int (12)) + (Prims.of_int (181)) + (Prims.of_int (24))) + (Prims.mk_range + "FStar.Tactics.Logic.fst" + (Prims.of_int (182)) + (Prims.of_int (12)) + (Prims.of_int (183)) + (Prims.of_int (66))) + (Obj.magic + (FStar_Tactics_Builtins.rewrite + eq_h)) (fun uu___3 -> - Obj.magic + (fun + uu___3 -> + Obj.magic (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.Logic.fst" @@ -926,8 +935,8 @@ let rec (unfold_definition_and_simplify_eq : unfold_definition_and_simplify_eq tm))) uu___4))) - uu___3))) - uu___3))) uu___2)))) + uu___3))) + uu___3)))) uu___2)))) uu___))) uu___) let (unsquash : FStar_Reflection_Types.term -> @@ -1238,7 +1247,10 @@ let (instantiate : (fa, FStar_Reflection_Data.Q_Explicit)))), (x, FStar_Reflection_Data.Q_Explicit))))) (fun uu___1 -> - FStar_Tactics_Derived.fail "could not instantiate")) + (fun uu___1 -> + Obj.magic + (FStar_Tactics_Derived.fail "could not instantiate")) + uu___1)) let (instantiate_as : FStar_Reflection_Types.term -> FStar_Reflection_Types.term -> @@ -1361,19 +1373,12 @@ let rec (sk_binder' : (fun uu___4 -> uu___3 <> Prims.int_one)))) (fun uu___3 -> - (fun uu___3 -> - if uu___3 - then - Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "no")) - else - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___5 -> ())))) - uu___3))) + if uu___3 + then + FStar_Tactics_Derived.fail "no" + else + FStar_Tactics_Effect.lift_div_tac + (fun uu___5 -> ())))) (fun uu___3 -> (fun uu___3 -> Obj.magic @@ -1568,5 +1573,7 @@ let (using_lemma : "lem3_fa"]))), (t, FStar_Reflection_Data.Q_Explicit))))) (fun uu___2 -> - FStar_Tactics_Derived.fail - "using_lemma: failed to instantiate"))) \ No newline at end of file + (fun uu___2 -> + Obj.magic + (FStar_Tactics_Derived.fail + "using_lemma: failed to instantiate")) uu___2))) \ No newline at end of file diff --git a/ocaml/fstar-lib/generated/FStar_Tactics_PatternMatching.ml b/ocaml/fstar-lib/generated/FStar_Tactics_PatternMatching.ml index de9f29a87f5..976c932a28e 100644 --- a/ocaml/fstar-lib/generated/FStar_Tactics_PatternMatching.ml +++ b/ocaml/fstar-lib/generated/FStar_Tactics_PatternMatching.ml @@ -28,73 +28,79 @@ let (fetch_eq_side : | FStar_Reflection_Data.Tv_App (squash, (g1, uu___2)) -> Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.PatternMatching.fst" - (Prims.of_int (66)) (Prims.of_int (11)) - (Prims.of_int (66)) (Prims.of_int (25))) - (Prims.mk_range - "FStar.Tactics.PatternMatching.fst" - (Prims.of_int (66)) (Prims.of_int (4)) - (Prims.of_int (87)) (Prims.of_int (51))) - (Obj.magic - (FStar_Tactics_Builtins.inspect squash)) - (fun uu___3 -> + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.PatternMatching.fst" + (Prims.of_int (66)) (Prims.of_int (11)) + (Prims.of_int (66)) (Prims.of_int (25))) + (Prims.mk_range + "FStar.Tactics.PatternMatching.fst" + (Prims.of_int (66)) (Prims.of_int (4)) + (Prims.of_int (87)) (Prims.of_int (51))) + (Obj.magic + (FStar_Tactics_Builtins.inspect squash)) (fun uu___3 -> - match uu___3 with - | FStar_Reflection_Data.Tv_UInst - (squash1, uu___4) -> - if - (FStar_Reflection_Derived.fv_to_string - squash1) - = - (FStar_Reflection_Derived.flatten_name - FStar_Reflection_Const.squash_qn) - then - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.PatternMatching.fst" - (Prims.of_int (70)) - (Prims.of_int (16)) - (Prims.of_int (70)) - (Prims.of_int (25))) - (Prims.mk_range - "FStar.Tactics.PatternMatching.fst" - (Prims.of_int (70)) - (Prims.of_int (9)) - (Prims.of_int (85)) - (Prims.of_int (48))) - (Obj.magic - (FStar_Tactics_Builtins.inspect - g1)) - (fun uu___5 -> - (fun uu___5 -> - match uu___5 with - | FStar_Reflection_Data.Tv_App - (eq_type_x, - (y, uu___6)) - -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.PatternMatching.fst" - (Prims.of_int (72)) - (Prims.of_int (19)) - (Prims.of_int (72)) - (Prims.of_int (36))) - (Prims.mk_range - "FStar.Tactics.PatternMatching.fst" - (Prims.of_int (72)) - (Prims.of_int (12)) - (Prims.of_int (84)) - (Prims.of_int (39))) - (Obj.magic - (FStar_Tactics_Builtins.inspect + (fun uu___3 -> + match uu___3 with + | FStar_Reflection_Data.Tv_UInst + (squash1, uu___4) -> + Obj.magic + (Obj.repr + (if + (FStar_Reflection_Derived.fv_to_string + squash1) + = + (FStar_Reflection_Derived.flatten_name + FStar_Reflection_Const.squash_qn) + then + Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.PatternMatching.fst" + (Prims.of_int (70)) + (Prims.of_int (16)) + (Prims.of_int (70)) + (Prims.of_int (25))) + (Prims.mk_range + "FStar.Tactics.PatternMatching.fst" + (Prims.of_int (70)) + (Prims.of_int (9)) + (Prims.of_int (85)) + (Prims.of_int (48))) + (Obj.magic + (FStar_Tactics_Builtins.inspect + g1)) + (fun uu___5 -> + (fun uu___5 -> + match uu___5 + with + | FStar_Reflection_Data.Tv_App + (eq_type_x, + (y, + uu___6)) + -> + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.PatternMatching.fst" + (Prims.of_int (72)) + (Prims.of_int (19)) + (Prims.of_int (72)) + (Prims.of_int (36))) + (Prims.mk_range + "FStar.Tactics.PatternMatching.fst" + (Prims.of_int (72)) + (Prims.of_int (12)) + (Prims.of_int (84)) + (Prims.of_int (39))) + (Obj.magic + (FStar_Tactics_Builtins.inspect eq_type_x)) - (fun uu___7 - -> - (fun + (fun + uu___7 -> + (fun uu___7 -> match uu___7 with @@ -105,6 +111,7 @@ let (fetch_eq_side : uu___8)) -> Obj.magic + (Obj.repr (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.PatternMatching.fst" @@ -134,6 +141,7 @@ let (fetch_eq_side : uu___10)) -> Obj.magic + (Obj.repr (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.PatternMatching.fst" @@ -153,9 +161,6 @@ let (fetch_eq_side : (fun uu___11 -> - (fun - uu___11 - -> match uu___11 with | @@ -163,126 +168,123 @@ let (fetch_eq_side : (eq1, uu___12) -> - Obj.magic - (Obj.repr - (if + if (FStar_Reflection_Derived.fv_to_string eq1) = (FStar_Reflection_Derived.flatten_name FStar_Reflection_Const.eq2_qn) then - Obj.repr - (FStar_Tactics_Effect.lift_div_tac + FStar_Tactics_Effect.lift_div_tac (fun uu___13 - -> (x, y))) + -> (x, y)) else - Obj.repr - (FStar_Tactics_Derived.fail - "not an equality"))) + FStar_Tactics_Derived.fail + "not an equality" | FStar_Reflection_Data.Tv_FVar eq1 -> - Obj.magic - (Obj.repr - (if + if (FStar_Reflection_Derived.fv_to_string eq1) = (FStar_Reflection_Derived.flatten_name FStar_Reflection_Const.eq2_qn) then - Obj.repr - (FStar_Tactics_Effect.lift_div_tac + FStar_Tactics_Effect.lift_div_tac (fun uu___12 - -> (x, y))) + -> (x, y)) else - Obj.repr - (FStar_Tactics_Derived.fail - "not an equality"))) + FStar_Tactics_Derived.fail + "not an equality" | uu___12 -> - Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail + FStar_Tactics_Derived.fail "not an app2 of fvar: "))) - uu___11)) | uu___10 -> Obj.magic + (Obj.repr (FStar_Tactics_Derived.fail - "not an app3")) - uu___9)) + "not an app3"))) + uu___9))) | uu___8 -> Obj.magic + (Obj.repr (FStar_Tactics_Derived.fail - "not an app2")) - uu___7)) - | uu___6 -> - Obj.magic - (FStar_Tactics_Derived.fail - "not an app under squash")) - uu___5)) - else - Obj.magic - (FStar_Tactics_Derived.fail - "not a squash") - | FStar_Reflection_Data.Tv_FVar squash1 - -> - if - (FStar_Reflection_Derived.fv_to_string - squash1) - = - (FStar_Reflection_Derived.flatten_name - FStar_Reflection_Const.squash_qn) - then - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.PatternMatching.fst" - (Prims.of_int (70)) - (Prims.of_int (16)) - (Prims.of_int (70)) - (Prims.of_int (25))) - (Prims.mk_range - "FStar.Tactics.PatternMatching.fst" - (Prims.of_int (70)) - (Prims.of_int (9)) - (Prims.of_int (85)) - (Prims.of_int (48))) - (Obj.magic - (FStar_Tactics_Builtins.inspect - g1)) - (fun uu___4 -> - (fun uu___4 -> - match uu___4 with - | FStar_Reflection_Data.Tv_App - (eq_type_x, - (y, uu___5)) - -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.PatternMatching.fst" - (Prims.of_int (72)) - (Prims.of_int (19)) - (Prims.of_int (72)) - (Prims.of_int (36))) - (Prims.mk_range - "FStar.Tactics.PatternMatching.fst" - (Prims.of_int (72)) - (Prims.of_int (12)) - (Prims.of_int (84)) - (Prims.of_int (39))) - (Obj.magic - (FStar_Tactics_Builtins.inspect + "not an app2"))) + uu___7))) + | uu___6 -> + Obj.magic + (Obj.repr + (FStar_Tactics_Derived.fail + "not an app under squash"))) + uu___5)) + else + Obj.repr + (FStar_Tactics_Derived.fail + "not a squash"))) + | FStar_Reflection_Data.Tv_FVar + squash1 -> + Obj.magic + (Obj.repr + (if + (FStar_Reflection_Derived.fv_to_string + squash1) + = + (FStar_Reflection_Derived.flatten_name + FStar_Reflection_Const.squash_qn) + then + Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.PatternMatching.fst" + (Prims.of_int (70)) + (Prims.of_int (16)) + (Prims.of_int (70)) + (Prims.of_int (25))) + (Prims.mk_range + "FStar.Tactics.PatternMatching.fst" + (Prims.of_int (70)) + (Prims.of_int (9)) + (Prims.of_int (85)) + (Prims.of_int (48))) + (Obj.magic + (FStar_Tactics_Builtins.inspect + g1)) + (fun uu___4 -> + (fun uu___4 -> + match uu___4 + with + | FStar_Reflection_Data.Tv_App + (eq_type_x, + (y, + uu___5)) + -> + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.PatternMatching.fst" + (Prims.of_int (72)) + (Prims.of_int (19)) + (Prims.of_int (72)) + (Prims.of_int (36))) + (Prims.mk_range + "FStar.Tactics.PatternMatching.fst" + (Prims.of_int (72)) + (Prims.of_int (12)) + (Prims.of_int (84)) + (Prims.of_int (39))) + (Obj.magic + (FStar_Tactics_Builtins.inspect eq_type_x)) - (fun uu___6 - -> - (fun + (fun + uu___6 -> + (fun uu___6 -> match uu___6 with @@ -293,6 +295,7 @@ let (fetch_eq_side : uu___7)) -> Obj.magic + (Obj.repr (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.PatternMatching.fst" @@ -322,6 +325,7 @@ let (fetch_eq_side : uu___9)) -> Obj.magic + (Obj.repr (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.PatternMatching.fst" @@ -341,9 +345,6 @@ let (fetch_eq_side : (fun uu___10 -> - (fun - uu___10 - -> match uu___10 with | @@ -351,81 +352,76 @@ let (fetch_eq_side : (eq1, uu___11) -> - Obj.magic - (Obj.repr - (if + if (FStar_Reflection_Derived.fv_to_string eq1) = (FStar_Reflection_Derived.flatten_name FStar_Reflection_Const.eq2_qn) then - Obj.repr - (FStar_Tactics_Effect.lift_div_tac + FStar_Tactics_Effect.lift_div_tac (fun uu___12 - -> (x, y))) + -> (x, y)) else - Obj.repr - (FStar_Tactics_Derived.fail - "not an equality"))) + FStar_Tactics_Derived.fail + "not an equality" | FStar_Reflection_Data.Tv_FVar eq1 -> - Obj.magic - (Obj.repr - (if + if (FStar_Reflection_Derived.fv_to_string eq1) = (FStar_Reflection_Derived.flatten_name FStar_Reflection_Const.eq2_qn) then - Obj.repr - (FStar_Tactics_Effect.lift_div_tac + FStar_Tactics_Effect.lift_div_tac (fun uu___11 - -> (x, y))) + -> (x, y)) else - Obj.repr - (FStar_Tactics_Derived.fail - "not an equality"))) + FStar_Tactics_Derived.fail + "not an equality" | uu___11 -> - Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail + FStar_Tactics_Derived.fail "not an app2 of fvar: "))) - uu___10)) | uu___9 -> Obj.magic + (Obj.repr (FStar_Tactics_Derived.fail - "not an app3")) - uu___8)) + "not an app3"))) + uu___8))) | uu___7 -> Obj.magic + (Obj.repr (FStar_Tactics_Derived.fail - "not an app2")) - uu___6)) - | uu___5 -> - Obj.magic - (FStar_Tactics_Derived.fail - "not an app under squash")) - uu___4)) - else - Obj.magic - (FStar_Tactics_Derived.fail - "not a squash") - | uu___4 -> - Obj.magic - (FStar_Tactics_Derived.fail - "not an app of fvar at top level")) - uu___3)) + "not an app2"))) + uu___6))) + | uu___5 -> + Obj.magic + (Obj.repr + (FStar_Tactics_Derived.fail + "not an app under squash"))) + uu___4)) + else + Obj.repr + (FStar_Tactics_Derived.fail + "not a squash"))) + | uu___4 -> + Obj.magic + (Obj.repr + (FStar_Tactics_Derived.fail + "not an app of fvar at top level"))) + uu___3))) | uu___2 -> Obj.magic - (FStar_Tactics_Derived.fail - "not an app at top level")) uu___1))) uu___1) + (Obj.repr + (FStar_Tactics_Derived.fail + "not an app at top level"))) uu___1))) + uu___1) let mustfail : 'a . (unit -> ('a, unit) FStar_Tactics_Effect.tac_repr) -> @@ -441,15 +437,11 @@ let mustfail : (Prims.of_int (130)) (Prims.of_int (4)) (Prims.of_int (132)) (Prims.of_int (16))) (Obj.magic (FStar_Tactics_Derived.trytac t)) (fun uu___ -> - (fun uu___ -> - match uu___ with - | FStar_Pervasives_Native.Some uu___1 -> - Obj.magic (Obj.repr (FStar_Tactics_Derived.fail message)) - | FStar_Pervasives_Native.None -> - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac (fun uu___1 -> ())))) - uu___) + match uu___ with + | FStar_Pervasives_Native.Some uu___1 -> + FStar_Tactics_Derived.fail message + | FStar_Pervasives_Native.None -> + FStar_Tactics_Effect.lift_div_tac (fun uu___1 -> ())) let (implies_intro' : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = fun uu___ -> FStar_Tactics_Effect.tac_bind @@ -1045,10 +1037,8 @@ let lift_exn_tac : (Prims.of_int (268)) (Prims.of_int (18)) (Prims.of_int (268)) (Prims.of_int (61))) (Obj.magic (string_of_match_exception ex)) - (fun uu___ -> - (fun uu___ -> - Obj.magic (FStar_Tactics_Derived.fail uu___)) - uu___)))) uu___1 uu___ + (fun uu___ -> FStar_Tactics_Derived.fail uu___)))) + uu___1 uu___ let lift_exn_tactic : 'a 'b . ('a -> 'b match_res) -> 'a -> ('b, unit) FStar_Tactics_Effect.tac_repr @@ -1073,10 +1063,8 @@ let lift_exn_tactic : (Prims.of_int (273)) (Prims.of_int (18)) (Prims.of_int (273)) (Prims.of_int (61))) (Obj.magic (string_of_match_exception ex)) - (fun uu___ -> - (fun uu___ -> - Obj.magic (FStar_Tactics_Derived.fail uu___)) - uu___)))) uu___1 uu___ + (fun uu___ -> FStar_Tactics_Derived.fail uu___)))) + uu___1 uu___ type bindings = (varname * FStar_Reflection_Types.term) Prims.list let (string_of_bindings : bindings -> (Prims.string, unit) FStar_Tactics_Effect.tac_repr) = @@ -1496,10 +1484,8 @@ let (match_term : (Prims.of_int (338)) (Prims.of_int (20)) (Prims.of_int (338)) (Prims.of_int (63))) (Obj.magic (string_of_match_exception ex)) - (fun uu___1 -> - (fun uu___1 -> - Obj.magic (FStar_Tactics_Derived.fail uu___1)) - uu___1)))) uu___) + (fun uu___1 -> FStar_Tactics_Derived.fail uu___1)))) + uu___) let debug : 'uuuuu . 'uuuuu -> (unit, unit) FStar_Tactics_Effect.tac_repr = fun uu___ -> (fun msg -> @@ -1663,13 +1649,9 @@ let assoc_varname_fail : match FStar_List_Tot_Base.assoc key ls with | FStar_Pervasives_Native.None -> Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - (Prims.strcat "Not found: " key))) + (FStar_Tactics_Derived.fail (Prims.strcat "Not found: " key)) | FStar_Pervasives_Native.Some x -> - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac (fun uu___ -> x)))) + Obj.magic (FStar_Tactics_Effect.lift_div_tac (fun uu___ -> x))) uu___1 uu___ let ms_locate_hyp : 'a . @@ -1712,98 +1694,123 @@ let rec solve_mp_for_single_hyp : (matching_solution -> ('a, unit) FStar_Tactics_Effect.tac_repr) -> matching_solution -> ('a, unit) FStar_Tactics_Effect.tac_repr = - fun name -> - fun pat -> - fun hypotheses -> - fun body -> - fun part_sol -> - match hypotheses with - | [] -> FStar_Tactics_Derived.fail "No matching hypothesis" - | h::hs -> - FStar_Tactics_Derived.or_else - (fun uu___ -> - FStar_Tactics_Effect.tac_bind - (Prims.mk_range "FStar.Tactics.PatternMatching.fst" - (Prims.of_int (448)) (Prims.of_int (15)) - (Prims.of_int (448)) (Prims.of_int (73))) - (Prims.mk_range "FStar.Tactics.PatternMatching.fst" - (Prims.of_int (448)) (Prims.of_int (9)) - (Prims.of_int (453)) (Prims.of_int (73))) - (Obj.magic - (interp_pattern_aux pat part_sol.ms_vars - (FStar_Reflection_Derived.type_of_binder h))) - (fun uu___1 -> - (fun uu___1 -> - match uu___1 with - | Failure ex -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.PatternMatching.fst" - (Prims.of_int (450)) - (Prims.of_int (16)) - (Prims.of_int (450)) - (Prims.of_int (74))) - (Prims.mk_range - "FStar.Tactics.PatternMatching.fst" - (Prims.of_int (450)) - (Prims.of_int (11)) - (Prims.of_int (450)) - (Prims.of_int (74))) - (Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.PatternMatching.fst" - (Prims.of_int (450)) - (Prims.of_int (43)) - (Prims.of_int (450)) - (Prims.of_int (73))) - (Prims.mk_range "prims.fst" - (Prims.of_int (606)) - (Prims.of_int (19)) - (Prims.of_int (606)) - (Prims.of_int (31))) - (Obj.magic - (string_of_match_exception ex)) - (fun uu___2 -> - FStar_Tactics_Effect.lift_div_tac - (fun uu___3 -> - Prims.strcat - "Failed to match hyp: " - uu___2)))) - (fun uu___2 -> - (fun uu___2 -> - Obj.magic - (FStar_Tactics_Derived.fail - uu___2)) uu___2)) - | Success bindings1 -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.PatternMatching.fst" - (Prims.of_int (452)) - (Prims.of_int (35)) - (Prims.of_int (452)) - (Prims.of_int (37))) - (Prims.mk_range - "FStar.Tactics.PatternMatching.fst" - (Prims.of_int (453)) - (Prims.of_int (11)) - (Prims.of_int (453)) - (Prims.of_int (73))) - (FStar_Tactics_Effect.lift_div_tac - (fun uu___2 -> (name, h) :: - (part_sol.ms_hyps))) - (fun uu___2 -> - (fun ms_hyps -> - Obj.magic - (body - { - ms_vars = bindings1; - ms_hyps - })) uu___2))) uu___1)) - (fun uu___ -> - solve_mp_for_single_hyp name pat hs body part_sol) + fun uu___4 -> + fun uu___3 -> + fun uu___2 -> + fun uu___1 -> + fun uu___ -> + (fun name -> + fun pat -> + fun hypotheses -> + fun body -> + fun part_sol -> + match hypotheses with + | [] -> + Obj.magic + (Obj.repr + (FStar_Tactics_Derived.fail + "No matching hypothesis")) + | h::hs -> + Obj.magic + (Obj.repr + (FStar_Tactics_Derived.or_else + (fun uu___ -> + FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.PatternMatching.fst" + (Prims.of_int (448)) + (Prims.of_int (15)) + (Prims.of_int (448)) + (Prims.of_int (73))) + (Prims.mk_range + "FStar.Tactics.PatternMatching.fst" + (Prims.of_int (448)) + (Prims.of_int (9)) + (Prims.of_int (453)) + (Prims.of_int (73))) + (Obj.magic + (interp_pattern_aux pat + part_sol.ms_vars + (FStar_Reflection_Derived.type_of_binder + h))) + (fun uu___1 -> + (fun uu___1 -> + match uu___1 with + | Failure ex -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.PatternMatching.fst" + (Prims.of_int (450)) + (Prims.of_int (16)) + (Prims.of_int (450)) + (Prims.of_int (74))) + (Prims.mk_range + "FStar.Tactics.PatternMatching.fst" + (Prims.of_int (450)) + (Prims.of_int (11)) + (Prims.of_int (450)) + (Prims.of_int (74))) + (Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.PatternMatching.fst" + (Prims.of_int (450)) + (Prims.of_int (43)) + (Prims.of_int (450)) + (Prims.of_int (73))) + (Prims.mk_range + "prims.fst" + (Prims.of_int (606)) + (Prims.of_int (19)) + (Prims.of_int (606)) + (Prims.of_int (31))) + (Obj.magic + (string_of_match_exception + ex)) + (fun uu___2 -> + FStar_Tactics_Effect.lift_div_tac + (fun uu___3 + -> + Prims.strcat + "Failed to match hyp: " + uu___2)))) + (fun uu___2 -> + FStar_Tactics_Derived.fail + uu___2)) + | Success bindings1 -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.PatternMatching.fst" + (Prims.of_int (452)) + (Prims.of_int (35)) + (Prims.of_int (452)) + (Prims.of_int (37))) + (Prims.mk_range + "FStar.Tactics.PatternMatching.fst" + (Prims.of_int (453)) + (Prims.of_int (11)) + (Prims.of_int (453)) + (Prims.of_int (73))) + (FStar_Tactics_Effect.lift_div_tac + (fun uu___2 -> + (name, h) :: + (part_sol.ms_hyps))) + (fun uu___2 -> + (fun ms_hyps -> + Obj.magic + (body + { + ms_vars = + bindings1; + ms_hyps + })) uu___2))) + uu___1)) + (fun uu___ -> + solve_mp_for_single_hyp name pat hs + body part_sol)))) uu___4 uu___3 + uu___2 uu___1 uu___ let rec solve_mp_for_hyps : 'a . (varname * pattern) Prims.list -> @@ -1898,10 +1905,8 @@ let solve_mp : "Failed to match goal: " uu___1)))) (fun uu___1 -> - (fun uu___1 -> - Obj.magic - (FStar_Tactics_Derived.fail - uu___1)) uu___1))) + FStar_Tactics_Derived.fail + uu___1))) | Success bindings1 -> Obj.magic (Obj.repr @@ -2042,10 +2047,8 @@ let (pattern_of_term : (Prims.of_int (532)) (Prims.of_int (20)) (Prims.of_int (532)) (Prims.of_int (63))) (Obj.magic (string_of_match_exception ex)) - (fun uu___1 -> - (fun uu___1 -> - Obj.magic (FStar_Tactics_Derived.fail uu___1)) - uu___1)))) uu___) + (fun uu___1 -> FStar_Tactics_Derived.fail uu___1)))) + uu___) type 'a hyp = FStar_Reflection_Types.binder type 'a pm_goal = unit let (hyp_qn : Prims.string) = "FStar.Tactics.PatternMatching.hyp" @@ -2205,8 +2208,6 @@ let (classify_abspat_binder : ( fun uu___2 -> - (fun - uu___2 -> match uu___2 with | @@ -2214,32 +2215,25 @@ let (classify_abspat_binder : ((uu___3, goal_typ)::[]) -> - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac + FStar_Tactics_Effect.lift_div_tac (fun uu___4 -> (ABKGoal, - goal_typ)))) + goal_typ)) | Success uu___3 -> - Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "classifiy_abspat_binder: impossible (2)")) + FStar_Tactics_Derived.fail + "classifiy_abspat_binder: impossible (2)" | Failure uu___3 -> - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac + FStar_Tactics_Effect.lift_div_tac (fun uu___4 -> ((ABKVar typ), - typ))))) - uu___2)))) + typ)))))) uu___))) uu___))) uu___))) uu___))) uu___) let rec (binders_and_body_of_abs : diff --git a/ocaml/fstar-lib/generated/FStar_Tactics_Simplifier.ml b/ocaml/fstar-lib/generated/FStar_Tactics_Simplifier.ml index d3708421bb0..e6d04f4ca30 100644 --- a/ocaml/fstar-lib/generated/FStar_Tactics_Simplifier.ml +++ b/ocaml/fstar-lib/generated/FStar_Tactics_Simplifier.ml @@ -144,45 +144,55 @@ let (inhabit : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = match FStar_Reflection_Builtins.inspect_ln t with | FStar_Reflection_Data.Tv_FVar fv -> Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "FStar.Tactics.Simplifier.fst" - (Prims.of_int (200)) (Prims.of_int (17)) - (Prims.of_int (200)) (Prims.of_int (30))) - (Prims.mk_range "FStar.Tactics.Simplifier.fst" - (Prims.of_int (201)) (Prims.of_int (13)) - (Prims.of_int (204)) (Prims.of_int (20))) - (FStar_Tactics_Effect.lift_div_tac + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "FStar.Tactics.Simplifier.fst" + (Prims.of_int (200)) (Prims.of_int (17)) + (Prims.of_int (200)) (Prims.of_int (30))) + (Prims.mk_range "FStar.Tactics.Simplifier.fst" + (Prims.of_int (201)) (Prims.of_int (13)) + (Prims.of_int (204)) (Prims.of_int (20))) + (FStar_Tactics_Effect.lift_div_tac + (fun uu___1 -> + FStar_Reflection_Builtins.inspect_fv fv)) (fun uu___1 -> - FStar_Reflection_Builtins.inspect_fv fv)) - (fun uu___1 -> - (fun qn -> - if qn = FStar_Reflection_Const.int_lid - then - Obj.magic - (FStar_Tactics_Derived.exact - (FStar_Reflection_Builtins.pack_ln - (FStar_Reflection_Data.Tv_Const - (FStar_Reflection_Data.C_Int - (Prims.of_int (42)))))) - else - if qn = FStar_Reflection_Const.bool_lid - then - Obj.magic - (FStar_Tactics_Derived.exact - (FStar_Reflection_Builtins.pack_ln - (FStar_Reflection_Data.Tv_Const - FStar_Reflection_Data.C_True))) - else - if qn = FStar_Reflection_Const.unit_lid - then - Obj.magic - (FStar_Tactics_Derived.exact - (FStar_Reflection_Builtins.pack_ln - (FStar_Reflection_Data.Tv_Const - FStar_Reflection_Data.C_Unit))) - else Obj.magic (FStar_Tactics_Derived.fail "")) - uu___1)) - | uu___1 -> Obj.magic (FStar_Tactics_Derived.fail "")) uu___1) + (fun qn -> + if qn = FStar_Reflection_Const.int_lid + then + Obj.magic + (Obj.repr + (FStar_Tactics_Derived.exact + (FStar_Reflection_Builtins.pack_ln + (FStar_Reflection_Data.Tv_Const + (FStar_Reflection_Data.C_Int + (Prims.of_int (42))))))) + else + Obj.magic + (Obj.repr + (if qn = FStar_Reflection_Const.bool_lid + then + Obj.repr + (FStar_Tactics_Derived.exact + (FStar_Reflection_Builtins.pack_ln + (FStar_Reflection_Data.Tv_Const + FStar_Reflection_Data.C_True))) + else + Obj.repr + (if + qn = + FStar_Reflection_Const.unit_lid + then + Obj.repr + (FStar_Tactics_Derived.exact + (FStar_Reflection_Builtins.pack_ln + (FStar_Reflection_Data.Tv_Const + FStar_Reflection_Data.C_Unit))) + else + Obj.repr + (FStar_Tactics_Derived.fail ""))))) + uu___1))) + | uu___1 -> Obj.magic (Obj.repr (FStar_Tactics_Derived.fail ""))) + uu___1) let rec (simplify_point : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = fun uu___ -> @@ -239,29 +249,31 @@ let rec (simplify_point : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) | FStar_Reflection_Formula.Iff (l, r) -> Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.Simplifier.fst" - (Prims.of_int (217)) - (Prims.of_int (20)) - (Prims.of_int (217)) - (Prims.of_int (38))) - (Prims.mk_range - "FStar.Tactics.Simplifier.fst" - (Prims.of_int (217)) - (Prims.of_int (14)) - (Prims.of_int (264)) - (Prims.of_int (22))) - (Obj.magic - (FStar_Reflection_Formula.term_as_formula' - l)) - (fun uu___3 -> + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.Simplifier.fst" + (Prims.of_int (217)) + (Prims.of_int (20)) + (Prims.of_int (217)) + (Prims.of_int (38))) + (Prims.mk_range + "FStar.Tactics.Simplifier.fst" + (Prims.of_int (217)) + (Prims.of_int (14)) + (Prims.of_int (264)) + (Prims.of_int (22))) + (Obj.magic + (FStar_Reflection_Formula.term_as_formula' + l)) (fun uu___3 -> - match uu___3 with - | FStar_Reflection_Formula.And - (p, q) -> - Obj.magic - (FStar_Tactics_Effect.tac_bind + (fun uu___3 -> + match uu___3 + with + | FStar_Reflection_Formula.And + (p, q) -> + Obj.magic + (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.Simplifier.fst" (Prims.of_int (219)) @@ -401,10 +413,10 @@ let rec (simplify_point : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) uu___8))) uu___6))) uu___4)) - | FStar_Reflection_Formula.Or - (p, q) -> - Obj.magic - (FStar_Tactics_Effect.tac_bind + | FStar_Reflection_Formula.Or + (p, q) -> + Obj.magic + (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.Simplifier.fst" (Prims.of_int (226)) @@ -544,10 +556,10 @@ let rec (simplify_point : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) uu___8))) uu___6))) uu___4)) - | FStar_Reflection_Formula.Implies - (p, q) -> - Obj.magic - (FStar_Tactics_Effect.tac_bind + | FStar_Reflection_Formula.Implies + (p, q) -> + Obj.magic + (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.Simplifier.fst" (Prims.of_int (233)) @@ -650,10 +662,10 @@ let rec (simplify_point : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) uu___8))) uu___6))) uu___4)) - | FStar_Reflection_Formula.Forall - (b, p) -> - Obj.magic - (FStar_Tactics_Effect.tac_bind + | FStar_Reflection_Formula.Forall + (b, p) -> + Obj.magic + (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.Simplifier.fst" (Prims.of_int (239)) @@ -748,10 +760,10 @@ let rec (simplify_point : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) (tiff ())) uu___6))) uu___4)) - | FStar_Reflection_Formula.Exists - (b, p) -> - Obj.magic - (FStar_Tactics_Effect.tac_bind + | FStar_Reflection_Formula.Exists + (b, p) -> + Obj.magic + (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.Simplifier.fst" (Prims.of_int (244)) @@ -846,10 +858,10 @@ let rec (simplify_point : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) (tiff ())) uu___6))) uu___4)) - | FStar_Reflection_Formula.Not - p -> - Obj.magic - (FStar_Tactics_Effect.tac_bind + | FStar_Reflection_Formula.Not + p -> + Obj.magic + (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.Simplifier.fst" (Prims.of_int (249)) @@ -918,10 +930,10 @@ let rec (simplify_point : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) (tiff ())) uu___6))) uu___4)) - | FStar_Reflection_Formula.Iff - (p, q) -> - Obj.magic - (FStar_Tactics_Effect.tac_bind + | FStar_Reflection_Formula.Iff + (p, q) -> + Obj.magic + (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.Simplifier.fst" (Prims.of_int (256)) @@ -1104,14 +1116,15 @@ let rec (simplify_point : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) ())) uu___5))) uu___4)) - | uu___4 -> - Obj.magic - (tiff ())) - uu___3)) + | uu___4 -> + Obj.magic + (tiff ())) + uu___3))) | uu___3 -> Obj.magic - (FStar_Tactics_Derived.fail - "simplify_point: failed precondition: goal should be `g <==> ?u`")) + (Obj.repr + (FStar_Tactics_Derived.fail + "simplify_point: failed precondition: goal should be `g <==> ?u`"))) uu___3))) uu___3))) uu___2))) uu___1) and (recurse : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = @@ -1169,31 +1182,33 @@ and (recurse : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = | FStar_Reflection_Formula.Iff (l, r) -> Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.Simplifier.fst" - (Prims.of_int (275)) - (Prims.of_int (20)) - (Prims.of_int (275)) - (Prims.of_int (38))) - (Prims.mk_range - "FStar.Tactics.Simplifier.fst" - (Prims.of_int (275)) - (Prims.of_int (14)) - (Prims.of_int (302)) - (Prims.of_int (22))) - (Obj.magic - (FStar_Reflection_Formula.term_as_formula' - l)) - (fun uu___3 -> + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.Simplifier.fst" + (Prims.of_int (275)) + (Prims.of_int (20)) + (Prims.of_int (275)) + (Prims.of_int (38))) + (Prims.mk_range + "FStar.Tactics.Simplifier.fst" + (Prims.of_int (275)) + (Prims.of_int (14)) + (Prims.of_int (302)) + (Prims.of_int (22))) + (Obj.magic + (FStar_Reflection_Formula.term_as_formula' + l)) (fun uu___3 -> - match uu___3 with - | FStar_Reflection_Formula.And - (uu___4, - uu___5) - -> - Obj.magic - (FStar_Tactics_Derived.seq + (fun uu___3 -> + match uu___3 + with + | FStar_Reflection_Formula.And + (uu___4, + uu___5) + -> + Obj.magic + (FStar_Tactics_Derived.seq (fun uu___6 -> FStar_Tactics_Derived.apply_lemma @@ -1205,12 +1220,12 @@ and (recurse : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = "Simplifier"; "and_cong"])))) simplify_point) - | FStar_Reflection_Formula.Or - (uu___4, - uu___5) - -> - Obj.magic - (FStar_Tactics_Derived.seq + | FStar_Reflection_Formula.Or + (uu___4, + uu___5) + -> + Obj.magic + (FStar_Tactics_Derived.seq (fun uu___6 -> FStar_Tactics_Derived.apply_lemma @@ -1222,12 +1237,12 @@ and (recurse : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = "Simplifier"; "or_cong"])))) simplify_point) - | FStar_Reflection_Formula.Implies - (uu___4, - uu___5) - -> - Obj.magic - (FStar_Tactics_Derived.seq + | FStar_Reflection_Formula.Implies + (uu___4, + uu___5) + -> + Obj.magic + (FStar_Tactics_Derived.seq (fun uu___6 -> FStar_Tactics_Derived.apply_lemma @@ -1239,12 +1254,12 @@ and (recurse : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = "Simplifier"; "imp_cong"])))) simplify_point) - | FStar_Reflection_Formula.Forall - (uu___4, - uu___5) - -> - Obj.magic - (FStar_Tactics_Effect.tac_bind + | FStar_Reflection_Formula.Forall + (uu___4, + uu___5) + -> + Obj.magic + (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.Simplifier.fst" (Prims.of_int (286)) @@ -1296,12 +1311,12 @@ and (recurse : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = ())) uu___7))) uu___6)) - | FStar_Reflection_Formula.Exists - (uu___4, - uu___5) - -> - Obj.magic - (FStar_Tactics_Effect.tac_bind + | FStar_Reflection_Formula.Exists + (uu___4, + uu___5) + -> + Obj.magic + (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.Simplifier.fst" (Prims.of_int (291)) @@ -1353,10 +1368,10 @@ and (recurse : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = ())) uu___7))) uu___6)) - | FStar_Reflection_Formula.Not - uu___4 -> - Obj.magic - (FStar_Tactics_Effect.tac_bind + | FStar_Reflection_Formula.Not + uu___4 -> + Obj.magic + (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.Simplifier.fst" (Prims.of_int (296)) @@ -1386,12 +1401,12 @@ and (recurse : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = (simplify_point ())) uu___5)) - | FStar_Reflection_Formula.Iff - (uu___4, - uu___5) - -> - Obj.magic - (FStar_Tactics_Derived.seq + | FStar_Reflection_Formula.Iff + (uu___4, + uu___5) + -> + Obj.magic + (FStar_Tactics_Derived.seq (fun uu___6 -> FStar_Tactics_Derived.apply_lemma @@ -1403,14 +1418,15 @@ and (recurse : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = "Simplifier"; "iff_cong"])))) simplify_point) - | uu___4 -> - Obj.magic - (tiff ())) - uu___3)) + | uu___4 -> + Obj.magic + (tiff ())) + uu___3))) | uu___3 -> Obj.magic - (FStar_Tactics_Derived.fail - "recurse: failed precondition: goal should be `g <==> ?u`")) + (Obj.repr + (FStar_Tactics_Derived.fail + "recurse: failed precondition: goal should be `g <==> ?u`"))) uu___3))) uu___3))) uu___2))) uu___1) let (simplify : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = diff --git a/ocaml/fstar-lib/generated/FStar_Tactics_SyntaxHelpers.ml b/ocaml/fstar-lib/generated/FStar_Tactics_SyntaxHelpers.ml index 6f6a6fa5759..805016faf18 100644 --- a/ocaml/fstar-lib/generated/FStar_Tactics_SyntaxHelpers.ml +++ b/ocaml/fstar-lib/generated/FStar_Tactics_SyntaxHelpers.ml @@ -123,7 +123,11 @@ let (collect_abs : match uu___ with | (bs, t') -> ((FStar_List_Tot_Base.rev bs), t'))) let fail : 'a . Prims.string -> ('a, unit) FStar_Tactics_Effect.tac_repr = - fun m -> FStar_Tactics_Effect.raise (FStar_Tactics_Common.TacticFailure m) + fun uu___ -> + (fun m -> + Obj.magic + (FStar_Tactics_Effect.raise (FStar_Tactics_Common.TacticFailure m))) + uu___ let rec (mk_arr : FStar_Reflection_Types.binder Prims.list -> FStar_Reflection_Types.comp -> diff --git a/ocaml/fstar-lib/generated/FStar_Tactics_Typeclasses.ml b/ocaml/fstar-lib/generated/FStar_Tactics_Typeclasses.ml index df399efc6e2..b296a249006 100644 --- a/ocaml/fstar-lib/generated/FStar_Tactics_Typeclasses.ml +++ b/ocaml/fstar-lib/generated/FStar_Tactics_Typeclasses.ml @@ -4,13 +4,18 @@ let rec first : ('a -> ('b, unit) FStar_Tactics_Effect.tac_repr) -> 'a Prims.list -> ('b, unit) FStar_Tactics_Effect.tac_repr = - fun f -> - fun l -> - match l with - | [] -> FStar_Tactics_Derived.fail "no cands" - | x::xs -> - FStar_Tactics_Derived.or_else (fun uu___ -> f x) - (fun uu___ -> first f xs) + fun uu___1 -> + fun uu___ -> + (fun f -> + fun l -> + match l with + | [] -> + Obj.magic (Obj.repr (FStar_Tactics_Derived.fail "no cands")) + | x::xs -> + Obj.magic + (Obj.repr + (FStar_Tactics_Derived.or_else (fun uu___ -> f x) + (fun uu___ -> first f xs)))) uu___1 uu___ let rec (tcresolve' : FStar_Reflection_Types.term Prims.list -> Prims.int -> (unit, unit) FStar_Tactics_Effect.tac_repr) @@ -23,10 +28,8 @@ let rec (tcresolve' : (Prims.mk_range "FStar.Tactics.Typeclasses.fst" (Prims.of_int (45)) (Prims.of_int (4)) (Prims.of_int (50)) (Prims.of_int (137))) (if fuel <= Prims.int_zero - then Obj.magic (Obj.repr (FStar_Tactics_Derived.fail "out of fuel")) - else - Obj.magic - (Obj.repr (FStar_Tactics_Effect.lift_div_tac (fun uu___1 -> ())))) + then FStar_Tactics_Derived.fail "out of fuel" + else FStar_Tactics_Effect.lift_div_tac (fun uu___1 -> ())) (fun uu___ -> (fun uu___ -> Obj.magic @@ -72,15 +75,10 @@ let rec (tcresolve' : (FStar_Reflection_Builtins.term_eq g) seen then - Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "loop")) + FStar_Tactics_Derived.fail "loop" else - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___3 -> ())))) + FStar_Tactics_Effect.lift_div_tac + (fun uu___3 -> ())) (fun uu___2 -> (fun uu___2 -> Obj.magic @@ -155,11 +153,7 @@ let rec (tcresolve' : uu___5)))) (fun uu___5 -> - (fun - uu___5 -> - Obj.magic - (FStar_Tactics_Derived.fail - uu___5)) + FStar_Tactics_Derived.fail uu___5))))) uu___3))) uu___2))) uu___2))) uu___1))) uu___) @@ -360,11 +354,15 @@ let (tcresolve : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = (fun uu___2 -> match () with | () -> tcresolve' [] (Prims.of_int (16))) (fun uu___2 -> - match uu___2 with - | FStar_Tactics_Common.TacticFailure s -> - FStar_Tactics_Derived.fail - (Prims.strcat "Typeclass resolution failed: " s) - | e -> FStar_Tactics_Effect.raise e))) uu___1) + (fun uu___2 -> + match uu___2 with + | FStar_Tactics_Common.TacticFailure s -> + Obj.magic + (FStar_Tactics_Derived.fail + (Prims.strcat "Typeclass resolution failed: " + s)) + | e -> Obj.magic (FStar_Tactics_Effect.raise e)) + uu___2))) uu___1) let _ = FStar_Tactics_Native.register_tactic "FStar.Tactics.Typeclasses.tcresolve" (Prims.of_int (2)) @@ -1100,11 +1098,14 @@ let (mk_class : FStar_Pervasives_Native.None -> Obj.magic + (Obj.repr (FStar_Tactics_Derived.fail - "mk_class: proj not found?") + "mk_class: proj not found?")) | FStar_Pervasives_Native.Some se1 -> + Obj.magic + (Obj.repr (match FStar_Reflection_Builtins.inspect_sigelt se1 @@ -1113,7 +1114,7 @@ let (mk_class : FStar_Reflection_Data.Sg_Let (uu___10, lbs) -> - Obj.magic + Obj.repr (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.Typeclasses.fst" @@ -1154,9 +1155,9 @@ let (mk_class : | uu___10 -> - Obj.magic + Obj.repr (FStar_Tactics_Derived.fail - "mk_class: proj not Sg_Let?"))) + "mk_class: proj not Sg_Let?")))) uu___9))) (fun uu___9 -> @@ -1241,12 +1242,14 @@ let (mk_class : | [] -> Obj.magic + (Obj.repr (FStar_Tactics_Derived.fail - "mk_class: impossible, no binders") + "mk_class: impossible, no binders")) | b1::bs' -> Obj.magic + (Obj.repr (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.Typeclasses.fst" @@ -1311,7 +1314,7 @@ let (mk_class : :: bs')) cod2)) uu___12))) - uu___11)))) + uu___11))))) uu___10))) uu___9))) (fun diff --git a/ocaml/fstar-lib/generated/Steel_Effect_Common.ml b/ocaml/fstar-lib/generated/Steel_Effect_Common.ml index afb14f087b1..cdd2331bbcf 100644 --- a/ocaml/fstar-lib/generated/Steel_Effect_Common.ml +++ b/ocaml/fstar-lib/generated/Steel_Effect_Common.ml @@ -804,26 +804,15 @@ let (name_appears_in : (Prims.of_int (730)) (Prims.of_int (13))) (Obj.magic (FStar_Tactics_Builtins.inspect t1)) (fun uu___1 -> - (fun uu___1 -> - match uu___1 with - | FStar_Reflection_Data.Tv_FVar fv -> - Obj.magic - (Obj.repr - (if - (FStar_Reflection_Builtins.inspect_fv fv) - = nm - then - Obj.repr - (FStar_Tactics_Effect.raise Appears) - else - Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___3 -> ())))) - | t2 -> - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___2 -> ())))) uu___1))) + match uu___1 with + | FStar_Reflection_Data.Tv_FVar fv -> + if (FStar_Reflection_Builtins.inspect_fv fv) = nm + then FStar_Tactics_Effect.raise Appears + else + FStar_Tactics_Effect.lift_div_tac + (fun uu___3 -> ()) + | t2 -> + FStar_Tactics_Effect.lift_div_tac (fun uu___2 -> ())))) (fun uu___ -> (fun ff -> Obj.magic @@ -858,12 +847,9 @@ let (name_appears_in : match uu___ with | Appears -> Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___1 -> true))) - | e -> - Obj.magic - (Obj.repr (FStar_Tactics_Effect.raise e))) + (FStar_Tactics_Effect.lift_div_tac + (fun uu___1 -> true)) + | e -> Obj.magic (FStar_Tactics_Effect.raise e)) uu___))) uu___) let (term_appears_in : FStar_Reflection_Types.term -> @@ -1312,17 +1298,13 @@ let rec (try_candidates : (select t am) (select hd am))) (fun uu___3 -> - (fun uu___3 -> - if uu___3 - then - Obj.magic - (FStar_Tactics_Effect.raise - Success) - else - Obj.magic - (FStar_Tactics_Effect.raise - Failed)) - uu___3)) + if uu___3 + then + FStar_Tactics_Effect.raise + Success + else + FStar_Tactics_Effect.raise + Failed)) (fun uu___2 -> (fun uu___2 -> Obj.magic @@ -1384,9 +1366,8 @@ let rec (remove_from_list : (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (863)) (Prims.of_int (72)) (Prims.of_int (863)) (Prims.of_int (74))) - (Obj.magic - (FStar_Tactics_Derived.fail - "atom in remove_from_list not found: should not happen")) + (FStar_Tactics_Derived.fail + "atom in remove_from_list not found: should not happen") (fun uu___ -> FStar_Tactics_Effect.lift_div_tac (fun uu___1 -> [])))) @@ -1748,16 +1729,10 @@ let rec (try_unifying_remaining : (FStar_Tactics_Derived.unify u (select hd am))) (fun uu___1 -> - (fun uu___1 -> - if uu___1 - then - Obj.magic - (FStar_Tactics_Effect.raise - Success) - else - Obj.magic - (FStar_Tactics_Effect.raise - Failed)) uu___1)) + if uu___1 + then + FStar_Tactics_Effect.raise Success + else FStar_Tactics_Effect.raise Failed)) (fun uu___ -> match uu___ with | Success -> try_unifying_remaining tl u am @@ -1798,11 +1773,8 @@ let rec (try_unifying_remaining : "could not find candidate for scrutinee " uu___2)))) (fun uu___2 -> - (fun uu___2 -> - Obj.magic - (FStar_Tactics_Derived.fail - uu___2)) uu___2))))) uu___2 - uu___1 uu___ + FStar_Tactics_Derived.fail uu___2))))) + uu___2 uu___1 uu___ let (is_smt_binder : FStar_Reflection_Types.binder -> (Prims.bool, unit) FStar_Tactics_Effect.tac_repr) @@ -2085,13 +2057,11 @@ let rec (new_args_for_smt_attrs : with | FStar_Reflection_Data.C_Total ty2 -> - Obj.magic + FStar_Tactics_Effect.lift_div_tac ( - Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun + fun uu___1 -> - ty2))) + ty2) | FStar_Reflection_Data.C_Eff (uu___1, eff_name, @@ -2099,30 +2069,22 @@ let rec (new_args_for_smt_attrs : uu___2, uu___3) -> - Obj.magic - ( - Obj.repr - (if + if eff_name = ["Prims"; "Tot"] - then - Obj.repr - (FStar_Tactics_Effect.lift_div_tac + then + FStar_Tactics_Effect.lift_div_tac (fun uu___4 -> - ty2)) - else - Obj.repr - (FStar_Tactics_Derived.fail - "computation type not supported in definition of slprops"))) + ty2) + else + FStar_Tactics_Derived.fail + "computation type not supported in definition of slprops" | uu___1 -> - Obj.magic - ( - Obj.repr - (FStar_Tactics_Derived.fail - "computation type not supported in definition of slprops"))) + FStar_Tactics_Derived.fail + "computation type not supported in definition of slprops") (fun uu___1 -> (fun ty2 -> Obj.magic @@ -2360,9 +2322,7 @@ let fail_atoms : (fun uu___1 -> Prims.strcat "could not find a solution for unifying\n" uu___)))) - (fun uu___ -> - (fun uu___ -> Obj.magic (FStar_Tactics_Derived.fail uu___)) - uu___) + (fun uu___ -> FStar_Tactics_Derived.fail uu___) let rec (equivalent_lists_fallback : Prims.nat -> atom Prims.list -> @@ -2471,11 +2431,7 @@ let rec (equivalent_lists_fallback : "could not find candidates for " uu___2)))) (fun uu___2 -> - (fun uu___2 - -> - Obj.magic - (FStar_Tactics_Derived.fail - uu___2)) + FStar_Tactics_Derived.fail uu___2)))) uu___)) | uu___ -> @@ -2517,10 +2473,8 @@ let rec (equivalent_lists_fallback : "could not find candidates for " uu___1)))) (fun uu___1 -> - (fun uu___1 -> - Obj.magic - (FStar_Tactics_Derived.fail - uu___1)) uu___1)))) + FStar_Tactics_Derived.fail + uu___1)))) | uu___ -> Obj.magic (Obj.repr @@ -2767,11 +2721,7 @@ let rec (equivalent_lists' : uu___2)))) (fun uu___2 -> - (fun - uu___2 -> - Obj.magic - (FStar_Tactics_Derived.fail - uu___2)) + FStar_Tactics_Derived.fail uu___2)))) uu___)) | uu___ -> @@ -2813,10 +2763,8 @@ let rec (equivalent_lists' : "could not find candidates for " uu___1)))) (fun uu___1 -> - (fun uu___1 -> - Obj.magic - (FStar_Tactics_Derived.fail - uu___1)) uu___1)))) + FStar_Tactics_Derived.fail + uu___1)))) | uu___ -> Obj.magic (Obj.repr @@ -3171,18 +3119,13 @@ let rec (unifies_with_all_uvars : t hd_t)) (fun uu___3 -> - (fun - uu___3 -> if uu___3 then - Obj.magic - (FStar_Tactics_Effect.raise - Success) + FStar_Tactics_Effect.raise + Success else - Obj.magic - (FStar_Tactics_Effect.raise + FStar_Tactics_Effect.raise Failed)) - uu___3)) (fun uu___2 -> (fun @@ -4343,96 +4286,93 @@ let rec (unify_pr_with_true : | (hd, tl) -> if is_and hd then - (match tl with - | (pr_l, uu___1)::(pr_r, uu___2)::[] -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (1795)) (Prims.of_int (6)) - (Prims.of_int (1795)) (Prims.of_int (29))) - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (1796)) (Prims.of_int (6)) - (Prims.of_int (1796)) (Prims.of_int (29))) - (Obj.magic (unify_pr_with_true pr_l)) - (fun uu___3 -> - (fun uu___3 -> - Obj.magic (unify_pr_with_true pr_r)) uu___3)) - | uu___1 -> - Obj.magic - (FStar_Tactics_Derived.fail - "unify_pr_with_true: ill-formed /\\")) + Obj.magic + (Obj.repr + (match tl with + | (pr_l, uu___1)::(pr_r, uu___2)::[] -> + Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (1795)) (Prims.of_int (6)) + (Prims.of_int (1795)) (Prims.of_int (29))) + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (1796)) (Prims.of_int (6)) + (Prims.of_int (1796)) (Prims.of_int (29))) + (Obj.magic (unify_pr_with_true pr_l)) + (fun uu___3 -> + (fun uu___3 -> + Obj.magic (unify_pr_with_true pr_r)) + uu___3)) + | uu___1 -> + Obj.repr + (FStar_Tactics_Derived.fail + "unify_pr_with_true: ill-formed /\\"))) else Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (1799)) (Prims.of_int (10)) - (Prims.of_int (1799)) (Prims.of_int (30))) - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (1799)) (Prims.of_int (4)) - (Prims.of_int (1811)) (Prims.of_int (73))) - (Obj.magic - (FStar_Tactics_Derived.inspect_unascribe hd)) - (fun uu___2 -> + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (1799)) (Prims.of_int (10)) + (Prims.of_int (1799)) (Prims.of_int (30))) + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (1799)) (Prims.of_int (4)) + (Prims.of_int (1811)) (Prims.of_int (73))) + (Obj.magic + (FStar_Tactics_Derived.inspect_unascribe hd)) (fun uu___2 -> - match uu___2 with - | FStar_Reflection_Data.Tv_Uvar (uu___3, uu___4) - -> - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (1801)) - (Prims.of_int (9)) - (Prims.of_int (1801)) - (Prims.of_int (27))) - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (1801)) - (Prims.of_int (6)) - (Prims.of_int (1805)) - (Prims.of_int (9))) - (Obj.magic - (FStar_Tactics_Derived.unify pr - (FStar_Reflection_Builtins.pack_ln - (FStar_Reflection_Data.Tv_FVar - (FStar_Reflection_Builtins.pack_fv - ["Steel"; - "Effect"; - "Common"; - "true_p"]))))) - (fun uu___5 -> + (fun uu___2 -> + match uu___2 with + | FStar_Reflection_Data.Tv_Uvar + (uu___3, uu___4) -> + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (1801)) + (Prims.of_int (9)) + (Prims.of_int (1801)) + (Prims.of_int (27))) + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (1801)) + (Prims.of_int (6)) + (Prims.of_int (1805)) + (Prims.of_int (9))) + (Obj.magic + (FStar_Tactics_Derived.unify + pr + (FStar_Reflection_Builtins.pack_ln + (FStar_Reflection_Data.Tv_FVar + (FStar_Reflection_Builtins.pack_fv + ["Steel"; + "Effect"; + "Common"; + "true_p"]))))) (fun uu___5 -> if uu___5 then - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___6 -> ()))) + FStar_Tactics_Effect.lift_div_tac + (fun uu___6 -> ()) else - Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "unify_pr_with_true: could not unify SMT prop with True"))) - uu___5))) - | uu___3 -> - Obj.magic - (Obj.repr - (if - (FStar_List_Tot_Base.length - (FStar_Reflection_Builtins.free_uvars - pr)) - = Prims.int_zero - then - Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___4 -> ())) - else - Obj.repr - (FStar_Tactics_Effect.raise + FStar_Tactics_Derived.fail + "unify_pr_with_true: could not unify SMT prop with True"))) + | uu___3 -> + Obj.magic + (Obj.repr + (if + (FStar_List_Tot_Base.length + (FStar_Reflection_Builtins.free_uvars + pr)) + = Prims.int_zero + then + FStar_Tactics_Effect.lift_div_tac + (fun uu___4 -> ()) + else + FStar_Tactics_Effect.raise (Postpone - "unify_pr_with_true: some uvars are still there"))))) - uu___2))) uu___) + "unify_pr_with_true: some uvars are still there")))) + uu___2)))) uu___) let rec (set_abduction_variable_term : FStar_Reflection_Types.term -> (FStar_Reflection_Types.term, unit) FStar_Tactics_Effect.tac_repr) @@ -4451,27 +4391,112 @@ let rec (set_abduction_variable_term : | (hd, tl) -> if is_and hd then - (match tl with - | (pr_l, FStar_Reflection_Data.Q_Explicit)::(pr_r, - FStar_Reflection_Data.Q_Explicit)::[] - -> - if - (FStar_List_Tot_Base.length - (FStar_Reflection_Builtins.free_uvars pr_r)) - = Prims.int_zero - then - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (1838)) (Prims.of_int (18)) - (Prims.of_int (1838)) (Prims.of_int (50))) - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (1839)) (Prims.of_int (8)) - (Prims.of_int (1839)) (Prims.of_int (53))) - (Obj.magic (set_abduction_variable_term pr_l)) - (fun arg -> + Obj.magic + (Obj.repr + (match tl with + | (pr_l, FStar_Reflection_Data.Q_Explicit)::(pr_r, + FStar_Reflection_Data.Q_Explicit)::[] + -> + Obj.repr + (if + (FStar_List_Tot_Base.length + (FStar_Reflection_Builtins.free_uvars + pr_r)) + = Prims.int_zero + then + Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (1838)) + (Prims.of_int (18)) + (Prims.of_int (1838)) + (Prims.of_int (50))) + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (1839)) + (Prims.of_int (8)) + (Prims.of_int (1839)) + (Prims.of_int (53))) + (Obj.magic + (set_abduction_variable_term pr_l)) + (fun arg -> + FStar_Tactics_Effect.lift_div_tac + (fun uu___1 -> + FStar_Reflection_Derived.mk_app + (FStar_Reflection_Builtins.pack_ln + (FStar_Reflection_Data.Tv_FVar + (FStar_Reflection_Builtins.pack_fv + ["Steel"; + "Effect"; + "Common"; + "elim_and_l_squash"]))) + [(arg, + FStar_Reflection_Data.Q_Explicit)]))) + else + Obj.repr + (if + (FStar_List_Tot_Base.length + (FStar_Reflection_Builtins.free_uvars + pr_l)) + = Prims.int_zero + then + Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (1842)) + (Prims.of_int (18)) + (Prims.of_int (1842)) + (Prims.of_int (50))) + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (1843)) + (Prims.of_int (8)) + (Prims.of_int (1843)) + (Prims.of_int (53))) + (Obj.magic + (set_abduction_variable_term + pr_r)) + (fun arg -> + FStar_Tactics_Effect.lift_div_tac + (fun uu___2 -> + FStar_Reflection_Derived.mk_app + (FStar_Reflection_Builtins.pack_ln + (FStar_Reflection_Data.Tv_FVar + (FStar_Reflection_Builtins.pack_fv + ["Steel"; + "Effect"; + "Common"; + "elim_and_r_squash"]))) + [(arg, + FStar_Reflection_Data.Q_Explicit)]))) + else + Obj.repr + (FStar_Tactics_Effect.raise + (Postpone + "set_abduction_variable_term: there are still uvars on both sides of l_and")))) + | uu___1 -> + Obj.repr + (FStar_Tactics_Derived.fail + "set_abduction_variable: ill-formed /\\"))) + else + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (1831)) (Prims.of_int (6)) + (Prims.of_int (1831)) (Prims.of_int (8))) + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (1849)) (Prims.of_int (4)) + (Prims.of_int (1852)) (Prims.of_int (54))) + (Obj.magic (FStar_Tactics_Builtins.inspect hd)) + (fun uu___2 -> + match uu___2 with + | FStar_Reflection_Data.Tv_Uvar (uu___3, uu___4) + -> FStar_Tactics_Effect.lift_div_tac - (fun uu___1 -> + (fun uu___5 -> FStar_Reflection_Derived.mk_app (FStar_Reflection_Builtins.pack_ln (FStar_Reflection_Data.Tv_FVar @@ -4479,83 +4504,15 @@ let rec (set_abduction_variable_term : ["Steel"; "Effect"; "Common"; - "elim_and_l_squash"]))) - [(arg, - FStar_Reflection_Data.Q_Explicit)]))) - else - if - (FStar_List_Tot_Base.length - (FStar_Reflection_Builtins.free_uvars pr_l)) - = Prims.int_zero - then - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (1842)) (Prims.of_int (18)) - (Prims.of_int (1842)) (Prims.of_int (50))) - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (1843)) (Prims.of_int (8)) - (Prims.of_int (1843)) (Prims.of_int (53))) - (Obj.magic (set_abduction_variable_term pr_r)) - (fun arg -> - FStar_Tactics_Effect.lift_div_tac - (fun uu___2 -> - FStar_Reflection_Derived.mk_app - (FStar_Reflection_Builtins.pack_ln - (FStar_Reflection_Data.Tv_FVar - (FStar_Reflection_Builtins.pack_fv - ["Steel"; - "Effect"; - "Common"; - "elim_and_r_squash"]))) - [(arg, - FStar_Reflection_Data.Q_Explicit)]))) - else - Obj.magic - (FStar_Tactics_Effect.raise - (Postpone - "set_abduction_variable_term: there are still uvars on both sides of l_and")) - | uu___1 -> - Obj.magic - (FStar_Tactics_Derived.fail - "set_abduction_variable: ill-formed /\\")) - else - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (1831)) (Prims.of_int (6)) - (Prims.of_int (1831)) (Prims.of_int (8))) - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (1849)) (Prims.of_int (4)) - (Prims.of_int (1852)) (Prims.of_int (54))) - (Obj.magic (FStar_Tactics_Builtins.inspect hd)) - (fun uu___2 -> - (fun uu___2 -> - match uu___2 with - | FStar_Reflection_Data.Tv_Uvar (uu___3, uu___4) - -> - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___5 -> - FStar_Reflection_Derived.mk_app - (FStar_Reflection_Builtins.pack_ln - (FStar_Reflection_Data.Tv_FVar - (FStar_Reflection_Builtins.pack_fv - ["Steel"; - "Effect"; - "Common"; - "_return_squash"]))) - [((FStar_Reflection_Builtins.pack_ln - (FStar_Reflection_Data.Tv_Const - FStar_Reflection_Data.C_Unit)), - FStar_Reflection_Data.Q_Explicit)]))) + "_return_squash"]))) + [((FStar_Reflection_Builtins.pack_ln + (FStar_Reflection_Data.Tv_Const + FStar_Reflection_Data.C_Unit)), + FStar_Reflection_Data.Q_Explicit)]) | uu___3 -> - Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "set_abduction_variable: cannot unify"))) - uu___2))) uu___) + FStar_Tactics_Derived.fail + "set_abduction_variable: cannot unify")))) + uu___) let (set_abduction_variable : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = fun uu___ -> @@ -4581,62 +4538,68 @@ let (set_abduction_variable : match uu___1 with | FStar_Reflection_Data.Tv_Arrow (b, uu___2) -> Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (1858)) (Prims.of_int (18)) - (Prims.of_int (1858)) (Prims.of_int (34))) - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (1858)) (Prims.of_int (4)) - (Prims.of_int (1861)) (Prims.of_int (42))) - (FStar_Tactics_Effect.lift_div_tac - (fun uu___3 -> - FStar_Reflection_Builtins.inspect_binder - b)) - (fun uu___3 -> + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (1858)) + (Prims.of_int (18)) + (Prims.of_int (1858)) + (Prims.of_int (34))) + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (1858)) + (Prims.of_int (4)) + (Prims.of_int (1861)) + (Prims.of_int (42))) + (FStar_Tactics_Effect.lift_div_tac + (fun uu___3 -> + FStar_Reflection_Builtins.inspect_binder + b)) (fun uu___3 -> - match uu___3 with - | (bv, uu___4) -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (1859)) - (Prims.of_int (13)) - (Prims.of_int (1859)) - (Prims.of_int (26))) - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (1860)) - (Prims.of_int (4)) - (Prims.of_int (1861)) - (Prims.of_int (42))) - (FStar_Tactics_Effect.lift_div_tac + (fun uu___3 -> + match uu___3 with + | (bv, uu___4) -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (1859)) + (Prims.of_int (13)) + (Prims.of_int (1859)) + (Prims.of_int (26))) + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (1860)) + (Prims.of_int (4)) + (Prims.of_int (1861)) + (Prims.of_int (42))) + (FStar_Tactics_Effect.lift_div_tac + (fun uu___5 -> + FStar_Reflection_Builtins.inspect_bv + bv)) (fun uu___5 -> - FStar_Reflection_Builtins.inspect_bv - bv)) - (fun uu___5 -> - (fun bv1 -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (1860)) - (Prims.of_int (13)) - (Prims.of_int (1860)) - (Prims.of_int (23))) - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (1861)) - (Prims.of_int (4)) - (Prims.of_int (1861)) - (Prims.of_int (42))) - (FStar_Tactics_Effect.lift_div_tac + (fun bv1 -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (1860)) + (Prims.of_int (13)) + (Prims.of_int (1860)) + (Prims.of_int (23))) + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (1861)) + (Prims.of_int (4)) + (Prims.of_int (1861)) + (Prims.of_int (42))) + (FStar_Tactics_Effect.lift_div_tac + (fun uu___5 + -> + bv1.FStar_Reflection_Data.bv_sort)) (fun uu___5 -> - bv1.FStar_Reflection_Data.bv_sort)) - (fun uu___5 -> - (fun pr -> - Obj.magic - (FStar_Tactics_Effect.tac_bind + (fun pr -> + Obj.magic + (FStar_Tactics_Effect.tac_bind (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (1861)) @@ -4660,12 +4623,13 @@ let (set_abduction_variable : (FStar_Tactics_Derived.exact uu___5)) uu___5))) - uu___5))) - uu___5))) uu___3)) + uu___5))) + uu___5))) uu___3))) | uu___2 -> Obj.magic - (FStar_Tactics_Derived.fail "Not an arrow goal")) - uu___1))) uu___1) + (Obj.repr + (FStar_Tactics_Derived.fail + "Not an arrow goal"))) uu___1))) uu___1) let (canon_l_r : Prims.bool -> FStar_Reflection_Types.term -> @@ -4850,15 +4814,11 @@ let (canon_l_r : (flatten r2_raw) am2)) - (fun - uu___4 -> (fun res -> - Obj.magic - (FStar_Tactics_Effect.raise + FStar_Tactics_Effect.raise (Result res))) - uu___4)) (fun uu___3 -> (fun @@ -4869,24 +4829,21 @@ let (canon_l_r : FStar_Tactics_Common.TacticFailure m1 -> Obj.magic - (Obj.repr (FStar_Tactics_Derived.fail - m1)) + m1) | Result res -> Obj.magic - (Obj.repr (FStar_Tactics_Effect.lift_div_tac (fun uu___4 -> - res))) + res)) | uu___4 -> Obj.magic - (Obj.repr (FStar_Tactics_Derived.fail - "uncaught exception in equivalent_lists"))) + "uncaught exception in equivalent_lists")) uu___3))) (fun uu___3 -> @@ -6577,19 +6534,25 @@ let (canon_l_r : (fun uu___11 -> + (fun + uu___11 + -> match uu___11 with | FStar_Tactics_Common.TacticFailure msg -> - FStar_Tactics_Derived.fail + Obj.magic + (FStar_Tactics_Derived.fail (Prims.strcat "Cannot unify pr with true: " - msg) + msg)) | e -> - FStar_Tactics_Effect.raise - e))) + Obj.magic + (FStar_Tactics_Effect.raise + e)) + uu___11))) | l -> Obj.magic @@ -6832,34 +6795,40 @@ let (canon_monoid : (match rel_xy with | (rel_xy1, uu___2)::[] -> Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2090)) - (Prims.of_int (21)) - (Prims.of_int (2090)) - (Prims.of_int (43))) - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2090)) - (Prims.of_int (7)) - (Prims.of_int (2100)) - (Prims.of_int (8))) - (FStar_Tactics_Effect.lift_div_tac - (fun uu___3 -> - FStar_Reflection_Derived_Lemmas.collect_app_ref - rel_xy1)) - (fun uu___3 -> + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2090)) + (Prims.of_int (21)) + (Prims.of_int (2090)) + (Prims.of_int (43))) + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2090)) + (Prims.of_int (7)) + (Prims.of_int (2100)) + (Prims.of_int (8))) + (FStar_Tactics_Effect.lift_div_tac + (fun uu___3 + -> + FStar_Reflection_Derived_Lemmas.collect_app_ref + rel_xy1)) (fun uu___3 -> - match uu___3 - with - | (rel, xy) + (fun uu___3 + -> + match uu___3 + with + | + (rel, xy) -> if (FStar_List_Tot_Base.length xy) >= (Prims.of_int (2)) then + Obj.magic + (Obj.repr (match ((FStar_List_Tot_Base.index xy @@ -6878,7 +6847,7 @@ let (canon_monoid : (rhs, FStar_Reflection_Data.Q_Explicit)) -> - Obj.magic + Obj.repr (canon_l_r use_smt carrier_t @@ -6888,18 +6857,20 @@ let (canon_monoid : rhs) | uu___4 -> - Obj.magic + Obj.repr (FStar_Tactics_Derived.fail - "Goal should have been an application of a binary relation to 2 explicit arguments")) + "Goal should have been an application of a binary relation to 2 explicit arguments"))) else Obj.magic + (Obj.repr (FStar_Tactics_Derived.fail - "Goal should have been an application of a binary relation to n implicit and 2 explicit arguments")) - uu___3)) + "Goal should have been an application of a binary relation to n implicit and 2 explicit arguments"))) + uu___3))) | uu___2 -> Obj.magic - (FStar_Tactics_Derived.fail - "Goal should be squash applied to a binary relation"))) + (Obj.repr + (FStar_Tactics_Derived.fail + "Goal should be squash applied to a binary relation")))) uu___1))) uu___1))) uu___) let (canon' : Prims.bool -> @@ -7765,13 +7736,15 @@ let rec (extract_contexts : FStar_Pervasives_Native.None -> Obj.magic + (Obj.repr (FStar_Tactics_Derived.fail - "no context on the right either") + "no context on the right either")) | FStar_Pervasives_Native.Some g -> Obj.magic - (g ())) + (Obj.repr + (g ()))) uu___5))))))) uu___1))) uu___1)) | uu___1 -> @@ -8018,15 +7991,9 @@ let (open_existentials : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) "Common"; "solve_can_be_split_lookup"]))) e) then - Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "Tactic disabled: no available lemmas in context")) - else - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___2 -> ())))) + FStar_Tactics_Derived.fail + "Tactic disabled: no available lemmas in context" + else FStar_Tactics_Effect.lift_div_tac (fun uu___2 -> ())) (fun uu___1 -> (fun uu___1 -> Obj.magic @@ -8069,40 +8036,44 @@ let (open_existentials : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) FStar_Reflection_Data.Q_Explicit)::[]) -> Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2309)) - (Prims.of_int (22)) - (Prims.of_int (2309)) - (Prims.of_int (36))) - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2309)) - (Prims.of_int (9)) - (Prims.of_int (2325)) - (Prims.of_int (60))) - (FStar_Tactics_Effect.lift_div_tac - (fun uu___4 -> - FStar_Reflection_Derived.collect_app - t1)) - (fun uu___4 -> + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2309)) + (Prims.of_int (22)) + (Prims.of_int (2309)) + (Prims.of_int (36))) + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2309)) + (Prims.of_int (9)) + (Prims.of_int (2325)) + (Prims.of_int (60))) + (FStar_Tactics_Effect.lift_div_tac + (fun uu___4 -> + FStar_Reflection_Derived.collect_app + t1)) (fun uu___4 -> - match uu___4 with - | (hd, tl) -> - if - FStar_Reflection_Derived.is_fvar + (fun uu___4 -> + match uu___4 + with + | (hd, tl) -> + if + FStar_Reflection_Derived.is_fvar hd "Steel.Effect.Common.can_be_split" - then - (match tl - with - | - uu___5:: + then + Obj.magic + (Obj.repr + (match tl + with + | + uu___5:: (rhs, FStar_Reflection_Data.Q_Explicit)::[] -> - Obj.magic + Obj.repr (FStar_Tactics_Effect.tac_bind (Prims.mk_range "Steel.Effect.Common.fsti" @@ -8129,12 +8100,14 @@ let (open_existentials : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) FStar_Pervasives_Native.None -> Obj.magic + (Obj.repr (FStar_Tactics_Derived.fail - "open_existentials: no context found") + "open_existentials: no context found")) | FStar_Pervasives_Native.Some f -> Obj.magic + (Obj.repr (FStar_Tactics_Effect.tac_bind (Prims.mk_range "Steel.Effect.Common.fsti" @@ -8232,23 +8205,24 @@ let (open_existentials : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) uu___10))) uu___9))) uu___8))) - uu___7))) + uu___7)))) uu___6)) - | - uu___5 -> + | + uu___5 -> + Obj.repr + (FStar_Tactics_Derived.fail + "open_existentials: ill-formed can_be_split"))) + else Obj.magic + (Obj.repr (FStar_Tactics_Derived.fail - "open_existentials: ill-formed can_be_split")) - else - Obj.magic - ( - FStar_Tactics_Derived.fail - "open_existentials: not a can_be_split goal")) - uu___4)) + "open_existentials: not a can_be_split goal"))) + uu___4))) | uu___3 -> Obj.magic - (FStar_Tactics_Derived.fail - "open_existentials: not a squash goal")) + (Obj.repr + (FStar_Tactics_Derived.fail + "open_existentials: not a squash goal"))) uu___3))) uu___2))) uu___1))) uu___1) let (try_open_existentials : @@ -8567,85 +8541,104 @@ let (solve_can_be_split_dep : FStar_Reflection_Data.argv Prims.list -> (Prims.bool, unit) FStar_Tactics_Effect.tac_repr) = - fun args -> - match args with - | (p, uu___)::(t1, uu___1)::(t2, uu___2)::[] -> - FStar_Tactics_Effect.tac_bind - (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2388)) - (Prims.of_int (17)) (Prims.of_int (2388)) (Prims.of_int (36))) - (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2389)) - (Prims.of_int (6)) (Prims.of_int (2420)) (Prims.of_int (18))) - (Obj.magic (slterm_nbr_uvars t1)) - (fun uu___3 -> - (fun lnbr -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2389)) (Prims.of_int (17)) - (Prims.of_int (2389)) (Prims.of_int (36))) - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2390)) (Prims.of_int (6)) - (Prims.of_int (2420)) (Prims.of_int (18))) - (Obj.magic (slterm_nbr_uvars t2)) - (fun uu___3 -> - (fun rnbr -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2391)) (Prims.of_int (8)) - (Prims.of_int (2393)) (Prims.of_int (18))) - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2390)) (Prims.of_int (6)) - (Prims.of_int (2420)) (Prims.of_int (18))) - (if (lnbr + rnbr) <= Prims.int_one - then Obj.magic (Obj.repr (unfold_guard ())) - else - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___4 -> false)))) - (fun uu___3 -> - (fun uu___3 -> - if uu___3 - then - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2396)) - (Prims.of_int (8)) - (Prims.of_int (2416)) - (Prims.of_int (36))) - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2418)) - (Prims.of_int (8)) - (Prims.of_int (2418)) - (Prims.of_int (12))) - (Obj.magic - (FStar_Tactics_Derived.focus - (fun uu___4 -> - FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2397)) - (Prims.of_int (23)) - (Prims.of_int (2397)) - (Prims.of_int (39))) - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2398)) - (Prims.of_int (10)) - (Prims.of_int (2416)) - (Prims.of_int (35))) - (Obj.magic - (FStar_Tactics_Logic.implies_intro - ())) - (fun uu___5 -> - (fun p_bind -> - Obj.magic - (FStar_Tactics_Effect.tac_bind + fun uu___ -> + (fun args -> + match args with + | (p, uu___)::(t1, uu___1)::(t2, uu___2)::[] -> + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2388)) (Prims.of_int (17)) + (Prims.of_int (2388)) (Prims.of_int (36))) + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2389)) (Prims.of_int (6)) + (Prims.of_int (2420)) (Prims.of_int (18))) + (Obj.magic (slterm_nbr_uvars t1)) + (fun uu___3 -> + (fun lnbr -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2389)) (Prims.of_int (17)) + (Prims.of_int (2389)) (Prims.of_int (36))) + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2390)) (Prims.of_int (6)) + (Prims.of_int (2420)) (Prims.of_int (18))) + (Obj.magic (slterm_nbr_uvars t2)) + (fun uu___3 -> + (fun rnbr -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2391)) + (Prims.of_int (8)) + (Prims.of_int (2393)) + (Prims.of_int (18))) + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2390)) + (Prims.of_int (6)) + (Prims.of_int (2420)) + (Prims.of_int (18))) + (if (lnbr + rnbr) <= Prims.int_one + then + Obj.magic + (Obj.repr (unfold_guard ())) + else + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___4 -> false)))) + (fun uu___3 -> + (fun uu___3 -> + if uu___3 + then + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2396)) + (Prims.of_int (8)) + (Prims.of_int (2416)) + (Prims.of_int (36))) + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2418)) + (Prims.of_int (8)) + (Prims.of_int (2418)) + (Prims.of_int (12))) + (Obj.magic + (FStar_Tactics_Derived.focus + (fun uu___4 -> + FStar_Tactics_Effect.tac_bind + ( + Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2397)) + (Prims.of_int (23)) + (Prims.of_int (2397)) + (Prims.of_int (39))) + ( + Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2398)) + (Prims.of_int (10)) + (Prims.of_int (2416)) + (Prims.of_int (35))) + ( + Obj.magic + (FStar_Tactics_Logic.implies_intro + ())) + ( + fun + uu___5 -> + (fun + p_bind -> + Obj.magic + (FStar_Tactics_Effect.tac_bind (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2398)) @@ -8740,17 +8733,13 @@ let (solve_can_be_split_dep : Prims.op_Negation b then - Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "could not unify SMT prop with True")) + FStar_Tactics_Derived.fail + "could not unify SMT prop with True" else - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac + FStar_Tactics_Effect.lift_div_tac (fun uu___9 -> - ())))) + ())) (fun uu___8 -> (fun @@ -8912,121 +8901,144 @@ let (solve_can_be_split_dep : uu___8)))) uu___6))) uu___5))) - uu___5)))) - (fun uu___4 -> - FStar_Tactics_Effect.lift_div_tac - (fun uu___5 -> true)))) - else - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___5 -> false)))) - uu___3))) uu___3))) uu___3) - | uu___ -> FStar_Tactics_Derived.fail "ill-formed can_be_split_dep" + uu___5)))) + (fun uu___4 -> + FStar_Tactics_Effect.lift_div_tac + (fun uu___5 -> + true)))) + else + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___5 -> false)))) + uu___3))) uu___3))) uu___3))) + | uu___ -> + Obj.magic + (Obj.repr + (FStar_Tactics_Derived.fail "ill-formed can_be_split_dep"))) + uu___ let (solve_can_be_split_forall : FStar_Reflection_Data.argv Prims.list -> (Prims.bool, unit) FStar_Tactics_Effect.tac_repr) = - fun args -> - match args with - | uu___::(t1, uu___1)::(t2, uu___2)::[] -> - FStar_Tactics_Effect.tac_bind - (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2432)) - (Prims.of_int (17)) (Prims.of_int (2432)) (Prims.of_int (36))) - (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2433)) - (Prims.of_int (6)) (Prims.of_int (2458)) (Prims.of_int (18))) - (Obj.magic (slterm_nbr_uvars t1)) - (fun uu___3 -> - (fun lnbr -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2433)) (Prims.of_int (17)) - (Prims.of_int (2433)) (Prims.of_int (36))) - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2434)) (Prims.of_int (6)) - (Prims.of_int (2458)) (Prims.of_int (18))) - (Obj.magic (slterm_nbr_uvars t2)) - (fun uu___3 -> - (fun rnbr -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2435)) (Prims.of_int (8)) - (Prims.of_int (2437)) (Prims.of_int (18))) - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2434)) (Prims.of_int (6)) - (Prims.of_int (2458)) (Prims.of_int (18))) - (if (lnbr + rnbr) <= Prims.int_one - then Obj.magic (Obj.repr (unfold_guard ())) - else - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___4 -> false)))) - (fun uu___3 -> - (fun uu___3 -> - if uu___3 - then - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2440)) - (Prims.of_int (8)) - (Prims.of_int (2456)) - (Prims.of_int (46))) - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2457)) - (Prims.of_int (8)) - (Prims.of_int (2457)) - (Prims.of_int (12))) - (Obj.magic - (FStar_Tactics_Derived.focus - (fun uu___4 -> - FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2441)) - (Prims.of_int (10)) - (Prims.of_int (2441)) - (Prims.of_int (33))) - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2442)) - (Prims.of_int (10)) - (Prims.of_int (2456)) - (Prims.of_int (45))) - (Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range + fun uu___ -> + (fun args -> + match args with + | uu___::(t1, uu___1)::(t2, uu___2)::[] -> + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2432)) (Prims.of_int (17)) + (Prims.of_int (2432)) (Prims.of_int (36))) + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2433)) (Prims.of_int (6)) + (Prims.of_int (2458)) (Prims.of_int (18))) + (Obj.magic (slterm_nbr_uvars t1)) + (fun uu___3 -> + (fun lnbr -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2433)) (Prims.of_int (17)) + (Prims.of_int (2433)) (Prims.of_int (36))) + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2434)) (Prims.of_int (6)) + (Prims.of_int (2458)) (Prims.of_int (18))) + (Obj.magic (slterm_nbr_uvars t2)) + (fun uu___3 -> + (fun rnbr -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2435)) + (Prims.of_int (8)) + (Prims.of_int (2437)) + (Prims.of_int (18))) + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2434)) + (Prims.of_int (6)) + (Prims.of_int (2458)) + (Prims.of_int (18))) + (if (lnbr + rnbr) <= Prims.int_one + then + Obj.magic + (Obj.repr (unfold_guard ())) + else + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___4 -> false)))) + (fun uu___3 -> + (fun uu___3 -> + if uu___3 + then + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2440)) + (Prims.of_int (8)) + (Prims.of_int (2456)) + (Prims.of_int (46))) + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2457)) + (Prims.of_int (8)) + (Prims.of_int (2457)) + (Prims.of_int (12))) + (Obj.magic + (FStar_Tactics_Derived.focus + (fun uu___4 -> + FStar_Tactics_Effect.tac_bind + ( + Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2441)) + (Prims.of_int (10)) + (Prims.of_int (2441)) + (Prims.of_int (33))) + ( + Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2442)) + (Prims.of_int (10)) + (Prims.of_int (2456)) + (Prims.of_int (45))) + ( + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2441)) (Prims.of_int (17)) (Prims.of_int (2441)) (Prims.of_int (33))) - (Prims.mk_range + (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2441)) (Prims.of_int (10)) (Prims.of_int (2441)) (Prims.of_int (33))) - (Obj.magic - ( - FStar_Tactics_Logic.forall_intro + (Obj.magic + (FStar_Tactics_Logic.forall_intro ())) - (fun uu___5 - -> + (fun + uu___5 -> FStar_Tactics_Effect.lift_div_tac (fun uu___6 -> ())))) - (fun uu___5 -> - (fun uu___5 -> - Obj.magic - (FStar_Tactics_Effect.tac_bind + ( + fun + uu___5 -> + (fun + uu___5 -> + Obj.magic + (FStar_Tactics_Effect.tac_bind (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2442)) @@ -9225,19 +9237,23 @@ let (solve_can_be_split_forall : uu___9)))) uu___7))) uu___6))) - uu___5)))) - (fun uu___4 -> - FStar_Tactics_Effect.lift_div_tac - (fun uu___5 -> true)))) - else - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___5 -> false)))) - uu___3))) uu___3))) uu___3) - | uu___ -> - FStar_Tactics_Derived.fail - "Ill-formed can_be_split_forall, should not happen" + uu___5)))) + (fun uu___4 -> + FStar_Tactics_Effect.lift_div_tac + (fun uu___5 -> + true)))) + else + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___5 -> false)))) + uu___3))) uu___3))) uu___3))) + | uu___ -> + Obj.magic + (Obj.repr + (FStar_Tactics_Derived.fail + "Ill-formed can_be_split_forall, should not happen"))) + uu___ let (extract_cbs_forall_dep_contexts : FStar_Reflection_Types.term -> ((unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) @@ -9298,15 +9314,9 @@ let (open_existentials_forall_dep : "solve_can_be_split_forall_dep_lookup"]))) e) then - Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "Tactic disabled: no available lemmas in context")) - else - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___2 -> ())))) + FStar_Tactics_Derived.fail + "Tactic disabled: no available lemmas in context" + else FStar_Tactics_Effect.lift_div_tac (fun uu___2 -> ())) (fun uu___1 -> (fun uu___1 -> Obj.magic @@ -9354,40 +9364,44 @@ let (open_existentials_forall_dep : FStar_Reflection_Data.Q_Explicit)::[]) -> Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2491)) - (Prims.of_int (17)) - (Prims.of_int (2491)) - (Prims.of_int (31))) - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2491)) - (Prims.of_int (4)) - (Prims.of_int (2513)) - (Prims.of_int (78))) - (FStar_Tactics_Effect.lift_div_tac - (fun uu___4 -> - FStar_Reflection_Derived.collect_app - t1)) - (fun uu___4 -> + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2491)) + (Prims.of_int (17)) + (Prims.of_int (2491)) + (Prims.of_int (31))) + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2491)) + (Prims.of_int (4)) + (Prims.of_int (2513)) + (Prims.of_int (78))) + (FStar_Tactics_Effect.lift_div_tac + (fun uu___4 -> + FStar_Reflection_Derived.collect_app + t1)) (fun uu___4 -> - match uu___4 with - | (hd, tl) -> - if - FStar_Reflection_Derived.is_fvar + (fun uu___4 -> + match uu___4 + with + | (hd, tl) -> + if + FStar_Reflection_Derived.is_fvar hd "Steel.Effect.Common.can_be_split_forall_dep" - then - (match tl - with - | - uu___5::uu___6:: + then + Obj.magic + (Obj.repr + (match tl + with + | + uu___5::uu___6:: (rhs, FStar_Reflection_Data.Q_Explicit)::[] -> - Obj.magic + Obj.repr (FStar_Tactics_Effect.tac_bind (Prims.mk_range "Steel.Effect.Common.fsti" @@ -9415,6 +9429,7 @@ let (open_existentials_forall_dep : (uu___8, body) -> Obj.magic + (Obj.repr (FStar_Tactics_Effect.tac_bind (Prims.mk_range "Steel.Effect.Common.fsti" @@ -9441,12 +9456,14 @@ let (open_existentials_forall_dep : FStar_Pervasives_Native.None -> Obj.magic + (Obj.repr (FStar_Tactics_Derived.fail - "open_existentials_forall_dep: no candidate") + "open_existentials_forall_dep: no candidate")) | FStar_Pervasives_Native.Some f -> Obj.magic + (Obj.repr (FStar_Tactics_Effect.tac_bind (Prims.mk_range "Steel.Effect.Common.fsti" @@ -9632,21 +9649,22 @@ let (open_existentials_forall_dep : uu___13))) uu___12))) uu___11))) - uu___10))) - uu___9)) + uu___10)))) + uu___9))) | uu___8 -> Obj.magic + (Obj.repr (FStar_Tactics_Derived.fail - "open_existentials_forall_dep : not an abstraction")) + "open_existentials_forall_dep : not an abstraction"))) uu___7)) - | - (uu___5, + | + (uu___5, FStar_Reflection_Data.Q_Implicit)::uu___6::uu___7:: (rhs, FStar_Reflection_Data.Q_Explicit)::[] -> - Obj.magic + Obj.repr (FStar_Tactics_Effect.tac_bind (Prims.mk_range "Steel.Effect.Common.fsti" @@ -9674,6 +9692,7 @@ let (open_existentials_forall_dep : (uu___9, body) -> Obj.magic + (Obj.repr (FStar_Tactics_Effect.tac_bind (Prims.mk_range "Steel.Effect.Common.fsti" @@ -9702,12 +9721,14 @@ let (open_existentials_forall_dep : FStar_Pervasives_Native.None -> Obj.magic + (Obj.repr (FStar_Tactics_Derived.fail - "open_existentials_forall_dep: no candidate") + "open_existentials_forall_dep: no candidate")) | FStar_Pervasives_Native.Some f -> Obj.magic + (Obj.repr (FStar_Tactics_Effect.tac_bind (Prims.mk_range "Steel.Effect.Common.fsti" @@ -9893,29 +9914,31 @@ let (open_existentials_forall_dep : uu___14))) uu___13))) uu___12))) - uu___11))) - uu___10)) + uu___11)))) + uu___10))) | uu___9 -> Obj.magic + (Obj.repr (FStar_Tactics_Derived.fail - "open_existentials_forall_dep : not an abstraction")) + "open_existentials_forall_dep : not an abstraction"))) uu___8)) - | - uu___5 -> + | + uu___5 -> + Obj.repr + (FStar_Tactics_Derived.fail + "open_existentials_forall_dep : wrong number of arguments to can_be_split_forall_dep"))) + else Obj.magic + (Obj.repr (FStar_Tactics_Derived.fail - "open_existentials_forall_dep : wrong number of arguments to can_be_split_forall_dep")) - else - Obj.magic - ( - FStar_Tactics_Derived.fail - "open_existentials_forall_dep : not a can_be_split_forall_dep goal")) - uu___4)) + "open_existentials_forall_dep : not a can_be_split_forall_dep goal"))) + uu___4))) | uu___3 -> Obj.magic - (FStar_Tactics_Derived.fail - "open_existentials_forall_dep : not a squash/auto_squash goal")) + (Obj.repr + (FStar_Tactics_Derived.fail + "open_existentials_forall_dep : not a squash/auto_squash goal"))) uu___3))) uu___2))) uu___1))) uu___1) let (try_open_existentials_forall_dep : @@ -9946,90 +9969,100 @@ let rec (solve_can_be_split_forall_dep : FStar_Reflection_Data.argv Prims.list -> (Prims.bool, unit) FStar_Tactics_Effect.tac_repr) = - fun args -> - match args with - | uu___::(pr, uu___1)::(t1, uu___2)::(t2, uu___3)::[] -> - FStar_Tactics_Effect.tac_bind - (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2530)) - (Prims.of_int (17)) (Prims.of_int (2530)) (Prims.of_int (36))) - (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2531)) - (Prims.of_int (6)) (Prims.of_int (2575)) (Prims.of_int (18))) - (Obj.magic (slterm_nbr_uvars t1)) - (fun uu___4 -> - (fun lnbr -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2531)) (Prims.of_int (17)) - (Prims.of_int (2531)) (Prims.of_int (36))) - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2532)) (Prims.of_int (6)) - (Prims.of_int (2575)) (Prims.of_int (18))) - (Obj.magic (slterm_nbr_uvars t2)) - (fun uu___4 -> - (fun rnbr -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2533)) (Prims.of_int (8)) - (Prims.of_int (2535)) (Prims.of_int (18))) - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2532)) (Prims.of_int (6)) - (Prims.of_int (2575)) (Prims.of_int (18))) - (if (lnbr + rnbr) <= Prims.int_one - then Obj.magic (Obj.repr (unfold_guard ())) - else - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___5 -> false)))) - (fun uu___4 -> - (fun uu___4 -> - if uu___4 - then - Obj.magic - (Obj.repr - (FStar_Tactics_Derived.try_with - (fun uu___5 -> - match () with - | () -> - FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2539)) - (Prims.of_int (9)) - (Prims.of_int (2563)) - (Prims.of_int (37))) - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2565)) - (Prims.of_int (9)) - (Prims.of_int (2565)) - (Prims.of_int (13))) - (Obj.magic - (FStar_Tactics_Derived.focus - (fun uu___6 -> - FStar_Tactics_Effect.tac_bind - ( - Prims.mk_range + fun uu___ -> + (fun args -> + match args with + | uu___::(pr, uu___1)::(t1, uu___2)::(t2, uu___3)::[] -> + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2530)) (Prims.of_int (17)) + (Prims.of_int (2530)) (Prims.of_int (36))) + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2531)) (Prims.of_int (6)) + (Prims.of_int (2575)) (Prims.of_int (18))) + (Obj.magic (slterm_nbr_uvars t1)) + (fun uu___4 -> + (fun lnbr -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2531)) (Prims.of_int (17)) + (Prims.of_int (2531)) (Prims.of_int (36))) + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2532)) (Prims.of_int (6)) + (Prims.of_int (2575)) (Prims.of_int (18))) + (Obj.magic (slterm_nbr_uvars t2)) + (fun uu___4 -> + (fun rnbr -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2533)) + (Prims.of_int (8)) + (Prims.of_int (2535)) + (Prims.of_int (18))) + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2532)) + (Prims.of_int (6)) + (Prims.of_int (2575)) + (Prims.of_int (18))) + (if (lnbr + rnbr) <= Prims.int_one + then + Obj.magic + (Obj.repr (unfold_guard ())) + else + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___5 -> false)))) + (fun uu___4 -> + (fun uu___4 -> + if uu___4 + then + Obj.magic + (Obj.repr + (FStar_Tactics_Derived.try_with + (fun uu___5 -> + match () with + | () -> + FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2540)) - (Prims.of_int (10)) - (Prims.of_int (2540)) - (Prims.of_int (17))) - ( - Prims.mk_range + (Prims.of_int (2539)) + (Prims.of_int (9)) + (Prims.of_int (2563)) + (Prims.of_int (37))) + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2565)) + (Prims.of_int (9)) + (Prims.of_int (2565)) + (Prims.of_int (13))) + (Obj.magic + (FStar_Tactics_Derived.focus + (fun + uu___6 -> + FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2540)) + (Prims.of_int (10)) + (Prims.of_int (2540)) + (Prims.of_int (17))) + (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2541)) (Prims.of_int (10)) (Prims.of_int (2563)) (Prims.of_int (36))) - ( - Obj.magic + (Obj.magic (FStar_Tactics_Builtins.norm [])) - ( - fun + (fun uu___7 -> (fun uu___7 -> @@ -10290,17 +10323,13 @@ let rec (solve_can_be_split_forall_dep : Prims.op_Negation b then - Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "could not unify SMT prop with True")) + FStar_Tactics_Derived.fail + "could not unify SMT prop with True" else - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac + FStar_Tactics_Effect.lift_div_tac (fun uu___12 - -> ())))) + -> ())) (fun uu___11 -> @@ -10477,158 +10506,186 @@ let rec (solve_can_be_split_forall_dep : uu___8))) uu___8))) uu___7)))) - (fun uu___6 -> - FStar_Tactics_Effect.lift_div_tac - (fun uu___7 -> - true))) - (fun uu___5 -> - (fun uu___5 -> - match uu___5 with - | Postpone msg -> - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac (fun uu___6 - -> false))) - | FStar_Tactics_Common.TacticFailure - msg -> - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range + -> + FStar_Tactics_Effect.lift_div_tac + (fun + uu___7 -> + true))) + (fun uu___5 -> + (fun uu___5 -> + match uu___5 + with + | Postpone msg + -> + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun + uu___6 -> + false))) + | FStar_Tactics_Common.TacticFailure + msg -> + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2570)) (Prims.of_int (22)) (Prims.of_int (2570)) (Prims.of_int (57))) - (Prims.mk_range + (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2571)) (Prims.of_int (9)) (Prims.of_int (2573)) (Prims.of_int (22))) - (Obj.magic + (Obj.magic (try_open_existentials_forall_dep ())) - (fun uu___6 - -> + (fun + uu___6 -> (fun opened -> if opened then Obj.magic + (Obj.repr (solve_can_be_split_forall_dep - args) + args)) else Obj.magic + (Obj.repr (FStar_Tactics_Derived.fail - msg)) + msg))) uu___6))) - | uu___6 -> - Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "Unexpected exception in framing tactic"))) - uu___5))) - else - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___6 -> false)))) - uu___4))) uu___4))) uu___4) - | uu___ -> - FStar_Tactics_Derived.fail - "Ill-formed can_be_split_forall_dep, should not happen" + | uu___6 -> + Obj.magic + (Obj.repr + (FStar_Tactics_Derived.fail + "Unexpected exception in framing tactic"))) + uu___5))) + else + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___6 -> false)))) + uu___4))) uu___4))) uu___4))) + | uu___ -> + Obj.magic + (Obj.repr + (FStar_Tactics_Derived.fail + "Ill-formed can_be_split_forall_dep, should not happen"))) + uu___ let (solve_equiv_forall : FStar_Reflection_Data.argv Prims.list -> (Prims.bool, unit) FStar_Tactics_Effect.tac_repr) = - fun args -> - match args with - | uu___::(t1, uu___1)::(t2, uu___2)::[] -> - FStar_Tactics_Effect.tac_bind - (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2583)) - (Prims.of_int (17)) (Prims.of_int (2583)) (Prims.of_int (36))) - (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2584)) - (Prims.of_int (6)) (Prims.of_int (2612)) (Prims.of_int (18))) - (Obj.magic (slterm_nbr_uvars t1)) - (fun uu___3 -> - (fun lnbr -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2584)) (Prims.of_int (17)) - (Prims.of_int (2584)) (Prims.of_int (36))) - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2585)) (Prims.of_int (6)) - (Prims.of_int (2612)) (Prims.of_int (18))) - (Obj.magic (slterm_nbr_uvars t2)) - (fun uu___3 -> - (fun rnbr -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2586)) (Prims.of_int (8)) - (Prims.of_int (2588)) (Prims.of_int (18))) - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2585)) (Prims.of_int (6)) - (Prims.of_int (2612)) (Prims.of_int (18))) - (if (lnbr + rnbr) <= Prims.int_one - then Obj.magic (Obj.repr (unfold_guard ())) - else - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___4 -> false)))) - (fun uu___3 -> - (fun uu___3 -> - if uu___3 - then - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2591)) - (Prims.of_int (8)) - (Prims.of_int (2610)) - (Prims.of_int (62))) - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2611)) - (Prims.of_int (8)) - (Prims.of_int (2611)) - (Prims.of_int (12))) - (Obj.magic - (FStar_Tactics_Derived.focus - (fun uu___4 -> - FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2591)) - (Prims.of_int (24)) - (Prims.of_int (2591)) - (Prims.of_int (56))) - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2592)) - (Prims.of_int (22)) - (Prims.of_int (2610)) - (Prims.of_int (61))) - (Obj.magic - (FStar_Tactics_Derived.apply_lemma - (FStar_Reflection_Builtins.pack_ln + fun uu___ -> + (fun args -> + match args with + | uu___::(t1, uu___1)::(t2, uu___2)::[] -> + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2583)) (Prims.of_int (17)) + (Prims.of_int (2583)) (Prims.of_int (36))) + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2584)) (Prims.of_int (6)) + (Prims.of_int (2612)) (Prims.of_int (18))) + (Obj.magic (slterm_nbr_uvars t1)) + (fun uu___3 -> + (fun lnbr -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2584)) (Prims.of_int (17)) + (Prims.of_int (2584)) (Prims.of_int (36))) + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2585)) (Prims.of_int (6)) + (Prims.of_int (2612)) (Prims.of_int (18))) + (Obj.magic (slterm_nbr_uvars t2)) + (fun uu___3 -> + (fun rnbr -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2586)) + (Prims.of_int (8)) + (Prims.of_int (2588)) + (Prims.of_int (18))) + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2585)) + (Prims.of_int (6)) + (Prims.of_int (2612)) + (Prims.of_int (18))) + (if (lnbr + rnbr) <= Prims.int_one + then + Obj.magic + (Obj.repr (unfold_guard ())) + else + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___4 -> false)))) + (fun uu___3 -> + (fun uu___3 -> + if uu___3 + then + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2591)) + (Prims.of_int (8)) + (Prims.of_int (2610)) + (Prims.of_int (62))) + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2611)) + (Prims.of_int (8)) + (Prims.of_int (2611)) + (Prims.of_int (12))) + (Obj.magic + (FStar_Tactics_Derived.focus + (fun uu___4 -> + FStar_Tactics_Effect.tac_bind + ( + Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2591)) + (Prims.of_int (24)) + (Prims.of_int (2591)) + (Prims.of_int (56))) + ( + Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2592)) + (Prims.of_int (22)) + (Prims.of_int (2610)) + (Prims.of_int (61))) ( - FStar_Reflection_Data.Tv_FVar + Obj.magic + (FStar_Tactics_Derived.apply_lemma + (FStar_Reflection_Builtins.pack_ln + (FStar_Reflection_Data.Tv_FVar (FStar_Reflection_Builtins.pack_fv ["Steel"; "Effect"; "Common"; "equiv_forall_elim"]))))) - (fun uu___5 -> - (fun uu___5 -> - Obj.magic - (FStar_Tactics_Effect.tac_bind + ( + fun + uu___5 -> + (fun + uu___5 -> + Obj.magic + (FStar_Tactics_Effect.tac_bind (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2592)) @@ -10880,114 +10937,134 @@ let (solve_equiv_forall : uu___9))) uu___8)))) uu___6))) - uu___5)))) - (fun uu___4 -> - FStar_Tactics_Effect.lift_div_tac - (fun uu___5 -> true)))) - else - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___5 -> false)))) - uu___3))) uu___3))) uu___3) - | uu___ -> - FStar_Tactics_Derived.fail - "Ill-formed equiv_forall, should not happen" + uu___5)))) + (fun uu___4 -> + FStar_Tactics_Effect.lift_div_tac + (fun uu___5 -> + true)))) + else + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___5 -> false)))) + uu___3))) uu___3))) uu___3))) + | uu___ -> + Obj.magic + (Obj.repr + (FStar_Tactics_Derived.fail + "Ill-formed equiv_forall, should not happen"))) uu___ let (solve_equiv : FStar_Reflection_Data.argv Prims.list -> (Prims.bool, unit) FStar_Tactics_Effect.tac_repr) = - fun args -> - match args with - | (t1, uu___)::(t2, uu___1)::[] -> - FStar_Tactics_Effect.tac_bind - (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2620)) - (Prims.of_int (17)) (Prims.of_int (2620)) (Prims.of_int (36))) - (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2621)) - (Prims.of_int (6)) (Prims.of_int (2645)) (Prims.of_int (18))) - (Obj.magic (slterm_nbr_uvars t1)) - (fun uu___2 -> - (fun lnbr -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2621)) (Prims.of_int (17)) - (Prims.of_int (2621)) (Prims.of_int (36))) - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2622)) (Prims.of_int (6)) - (Prims.of_int (2645)) (Prims.of_int (18))) - (Obj.magic (slterm_nbr_uvars t2)) - (fun uu___2 -> - (fun rnbr -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2623)) (Prims.of_int (8)) - (Prims.of_int (2625)) (Prims.of_int (18))) - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2622)) (Prims.of_int (6)) - (Prims.of_int (2645)) (Prims.of_int (18))) - (if (lnbr + rnbr) <= Prims.int_one - then Obj.magic (Obj.repr (unfold_guard ())) - else - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___3 -> false)))) - (fun uu___2 -> - (fun uu___2 -> - if uu___2 - then - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2628)) - (Prims.of_int (8)) - (Prims.of_int (2642)) - (Prims.of_int (48))) - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2643)) - (Prims.of_int (8)) - (Prims.of_int (2643)) - (Prims.of_int (12))) - (Obj.magic - (FStar_Tactics_Derived.focus - (fun uu___3 -> - FStar_Tactics_Derived.or_else - (fun uu___4 -> - FStar_Tactics_Derived.apply_lemma - (FStar_Reflection_Builtins.pack_ln - (FStar_Reflection_Data.Tv_FVar + fun uu___ -> + (fun args -> + match args with + | (t1, uu___)::(t2, uu___1)::[] -> + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2620)) (Prims.of_int (17)) + (Prims.of_int (2620)) (Prims.of_int (36))) + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2621)) (Prims.of_int (6)) + (Prims.of_int (2645)) (Prims.of_int (18))) + (Obj.magic (slterm_nbr_uvars t1)) + (fun uu___2 -> + (fun lnbr -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2621)) (Prims.of_int (17)) + (Prims.of_int (2621)) (Prims.of_int (36))) + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2622)) (Prims.of_int (6)) + (Prims.of_int (2645)) (Prims.of_int (18))) + (Obj.magic (slterm_nbr_uvars t2)) + (fun uu___2 -> + (fun rnbr -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2623)) + (Prims.of_int (8)) + (Prims.of_int (2625)) + (Prims.of_int (18))) + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2622)) + (Prims.of_int (6)) + (Prims.of_int (2645)) + (Prims.of_int (18))) + (if (lnbr + rnbr) <= Prims.int_one + then + Obj.magic + (Obj.repr (unfold_guard ())) + else + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___3 -> false)))) + (fun uu___2 -> + (fun uu___2 -> + if uu___2 + then + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2628)) + (Prims.of_int (8)) + (Prims.of_int (2642)) + (Prims.of_int (48))) + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2643)) + (Prims.of_int (8)) + (Prims.of_int (2643)) + (Prims.of_int (12))) + (Obj.magic + (FStar_Tactics_Derived.focus + (fun uu___3 -> + FStar_Tactics_Derived.or_else + ( + fun + uu___4 -> + FStar_Tactics_Derived.apply_lemma + (FStar_Reflection_Builtins.pack_ln + (FStar_Reflection_Data.Tv_FVar (FStar_Reflection_Builtins.pack_fv ["Steel"; "Effect"; "Common"; "equiv_refl"])))) - (fun uu___4 -> - FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2632)) - (Prims.of_int (14)) - (Prims.of_int (2632)) - (Prims.of_int (68))) - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2633)) - (Prims.of_int (14)) - (Prims.of_int (2642)) - (Prims.of_int (46))) - (if - (lnbr <> + ( + fun + uu___4 -> + FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2632)) + (Prims.of_int (14)) + (Prims.of_int (2632)) + (Prims.of_int (68))) + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2633)) + (Prims.of_int (14)) + (Prims.of_int (2642)) + (Prims.of_int (46))) + (if + (lnbr <> Prims.int_zero) && (rnbr = Prims.int_zero) - then - Obj.magic + then + Obj.magic (Obj.repr (FStar_Tactics_Derived.apply_lemma (FStar_Reflection_Builtins.pack_ln @@ -10997,16 +11074,16 @@ let (solve_equiv : "Effect"; "Common"; "equiv_sym"]))))) - else - Obj.magic + else + Obj.magic (Obj.repr (FStar_Tactics_Effect.lift_div_tac (fun uu___6 -> ())))) - (fun uu___5 - -> - (fun + (fun + uu___5 -> + (fun uu___5 -> Obj.magic (FStar_Tactics_Effect.tac_bind @@ -11075,121 +11152,144 @@ let (solve_equiv : (fun uu___7 -> (fun - uu___7 -> + uu___7 -> + Obj.magic + (canon' + false + (FStar_Reflection_Builtins.pack_ln + (FStar_Reflection_Data.Tv_FVar + (FStar_Reflection_Builtins.pack_fv + ["Steel"; + "Effect"; + "Common"; + "true_p"]))) + (FStar_Reflection_Builtins.pack_ln + (FStar_Reflection_Data.Tv_FVar + (FStar_Reflection_Builtins.pack_fv + ["Steel"; + "Effect"; + "Common"; + "true_p"]))))) + uu___7))) + uu___6))) + uu___5))))) + (fun uu___3 -> + FStar_Tactics_Effect.lift_div_tac + (fun uu___4 -> + true)))) + else + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___4 -> false)))) + uu___2))) uu___2))) uu___2))) + | uu___ -> + Obj.magic + (Obj.repr + (FStar_Tactics_Derived.fail + "Ill-formed equiv, should not happen"))) uu___ +let (solve_can_be_split_post : + FStar_Reflection_Data.argv Prims.list -> + (Prims.bool, unit) FStar_Tactics_Effect.tac_repr) + = + fun uu___ -> + (fun args -> + match args with + | uu___::uu___1::(t1, uu___2)::(t2, uu___3)::[] -> + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2653)) (Prims.of_int (17)) + (Prims.of_int (2653)) (Prims.of_int (36))) + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2654)) (Prims.of_int (6)) + (Prims.of_int (2685)) (Prims.of_int (18))) + (Obj.magic (slterm_nbr_uvars t1)) + (fun uu___4 -> + (fun lnbr -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2654)) (Prims.of_int (17)) + (Prims.of_int (2654)) (Prims.of_int (36))) + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2655)) (Prims.of_int (6)) + (Prims.of_int (2685)) (Prims.of_int (18))) + (Obj.magic (slterm_nbr_uvars t2)) + (fun uu___4 -> + (fun rnbr -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2656)) + (Prims.of_int (8)) + (Prims.of_int (2658)) + (Prims.of_int (18))) + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2655)) + (Prims.of_int (6)) + (Prims.of_int (2685)) + (Prims.of_int (18))) + (if (lnbr + rnbr) <= Prims.int_one + then + Obj.magic + (Obj.repr (unfold_guard ())) + else + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___5 -> false)))) + (fun uu___4 -> + (fun uu___4 -> + if uu___4 + then + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2661)) + (Prims.of_int (8)) + (Prims.of_int (2683)) + (Prims.of_int (62))) + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2684)) + (Prims.of_int (8)) + (Prims.of_int (2684)) + (Prims.of_int (12))) + (Obj.magic + (FStar_Tactics_Derived.focus + (fun uu___5 -> + FStar_Tactics_Effect.tac_bind + ( + Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2661)) + (Prims.of_int (24)) + (Prims.of_int (2661)) + (Prims.of_int (30))) + ( + Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2662)) + (Prims.of_int (22)) + (Prims.of_int (2683)) + (Prims.of_int (61))) + ( + Obj.magic + (FStar_Tactics_Builtins.norm + [])) + ( + fun + uu___6 -> + (fun + uu___6 -> Obj.magic - (canon' - false - (FStar_Reflection_Builtins.pack_ln - (FStar_Reflection_Data.Tv_FVar - (FStar_Reflection_Builtins.pack_fv - ["Steel"; - "Effect"; - "Common"; - "true_p"]))) - (FStar_Reflection_Builtins.pack_ln - (FStar_Reflection_Data.Tv_FVar - (FStar_Reflection_Builtins.pack_fv - ["Steel"; - "Effect"; - "Common"; - "true_p"]))))) - uu___7))) - uu___6))) - uu___5))))) - (fun uu___3 -> - FStar_Tactics_Effect.lift_div_tac - (fun uu___4 -> true)))) - else - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___4 -> false)))) - uu___2))) uu___2))) uu___2) - | uu___ -> - FStar_Tactics_Derived.fail "Ill-formed equiv, should not happen" -let (solve_can_be_split_post : - FStar_Reflection_Data.argv Prims.list -> - (Prims.bool, unit) FStar_Tactics_Effect.tac_repr) - = - fun args -> - match args with - | uu___::uu___1::(t1, uu___2)::(t2, uu___3)::[] -> - FStar_Tactics_Effect.tac_bind - (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2653)) - (Prims.of_int (17)) (Prims.of_int (2653)) (Prims.of_int (36))) - (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2654)) - (Prims.of_int (6)) (Prims.of_int (2685)) (Prims.of_int (18))) - (Obj.magic (slterm_nbr_uvars t1)) - (fun uu___4 -> - (fun lnbr -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2654)) (Prims.of_int (17)) - (Prims.of_int (2654)) (Prims.of_int (36))) - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2655)) (Prims.of_int (6)) - (Prims.of_int (2685)) (Prims.of_int (18))) - (Obj.magic (slterm_nbr_uvars t2)) - (fun uu___4 -> - (fun rnbr -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2656)) (Prims.of_int (8)) - (Prims.of_int (2658)) (Prims.of_int (18))) - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2655)) (Prims.of_int (6)) - (Prims.of_int (2685)) (Prims.of_int (18))) - (if (lnbr + rnbr) <= Prims.int_one - then Obj.magic (Obj.repr (unfold_guard ())) - else - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___5 -> false)))) - (fun uu___4 -> - (fun uu___4 -> - if uu___4 - then - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2661)) - (Prims.of_int (8)) - (Prims.of_int (2683)) - (Prims.of_int (62))) - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2684)) - (Prims.of_int (8)) - (Prims.of_int (2684)) - (Prims.of_int (12))) - (Obj.magic - (FStar_Tactics_Derived.focus - (fun uu___5 -> - FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2661)) - (Prims.of_int (24)) - (Prims.of_int (2661)) - (Prims.of_int (30))) - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2662)) - (Prims.of_int (22)) - (Prims.of_int (2683)) - (Prims.of_int (61))) - (Obj.magic - (FStar_Tactics_Builtins.norm - [])) - (fun uu___6 -> - (fun uu___6 -> - Obj.magic - (FStar_Tactics_Effect.tac_bind + (FStar_Tactics_Effect.tac_bind (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2662)) @@ -11538,17 +11638,22 @@ let (solve_can_be_split_post : uu___8))) uu___7))) uu___7))) - uu___6)))) - (fun uu___5 -> - FStar_Tactics_Effect.lift_div_tac - (fun uu___6 -> true)))) - else - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___6 -> false)))) - uu___4))) uu___4))) uu___4) - | uu___ -> FStar_Tactics_Derived.fail "ill-formed can_be_split_post" + uu___6)))) + (fun uu___5 -> + FStar_Tactics_Effect.lift_div_tac + (fun uu___6 -> + true)))) + else + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___6 -> false)))) + uu___4))) uu___4))) uu___4))) + | uu___ -> + Obj.magic + (Obj.repr + (FStar_Tactics_Derived.fail "ill-formed can_be_split_post"))) + uu___ let (is_return_eq : FStar_Reflection_Types.term -> FStar_Reflection_Types.term -> @@ -11924,116 +12029,118 @@ let (goal_to_equiv : match f with | FStar_Reflection_Formula.App (hd0, t1) -> Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2746)) (Prims.of_int (6)) - (Prims.of_int (2747)) (Prims.of_int (70))) - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2748)) (Prims.of_int (6)) - (Prims.of_int (2768)) (Prims.of_int (51))) - (if - Prims.op_Negation - (FStar_Reflection_Derived.is_fvar hd0 - "Prims.squash") - then - Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - (Prims.strcat loc - " unexpected non-squash goal in goal_to_equiv"))) - else - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___1 -> ())))) - (fun uu___ -> + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2746)) + (Prims.of_int (6)) + (Prims.of_int (2747)) + (Prims.of_int (70))) + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2748)) + (Prims.of_int (6)) + (Prims.of_int (2768)) + (Prims.of_int (51))) + (if + Prims.op_Negation + (FStar_Reflection_Derived.is_fvar hd0 + "Prims.squash") + then + FStar_Tactics_Derived.fail + (Prims.strcat loc + " unexpected non-squash goal in goal_to_equiv") + else + FStar_Tactics_Effect.lift_div_tac + (fun uu___1 -> ())) (fun uu___ -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2748)) - (Prims.of_int (21)) - (Prims.of_int (2748)) - (Prims.of_int (34))) - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2748)) - (Prims.of_int (6)) - (Prims.of_int (2768)) - (Prims.of_int (51))) - (FStar_Tactics_Effect.lift_div_tac - (fun uu___1 -> - FStar_Reflection_Derived.collect_app - t1)) - (fun uu___1 -> + (fun uu___ -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2748)) + (Prims.of_int (21)) + (Prims.of_int (2748)) + (Prims.of_int (34))) + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2748)) + (Prims.of_int (6)) + (Prims.of_int (2768)) + (Prims.of_int (51))) + (FStar_Tactics_Effect.lift_div_tac + (fun uu___1 -> + FStar_Reflection_Derived.collect_app + t1)) (fun uu___1 -> - match uu___1 with - | (hd, args) -> - if - FStar_Reflection_Derived.is_fvar - hd - "Steel.Effect.Common.can_be_split" - then - Obj.magic - (FStar_Tactics_Derived.apply_lemma - (FStar_Reflection_Builtins.pack_ln - (FStar_Reflection_Data.Tv_FVar - (FStar_Reflection_Builtins.pack_fv + (fun uu___1 -> + match uu___1 with + | (hd, args) -> + if + FStar_Reflection_Derived.is_fvar + hd + "Steel.Effect.Common.can_be_split" + then + Obj.magic + (Obj.repr + (FStar_Tactics_Derived.apply_lemma + (FStar_Reflection_Builtins.pack_ln + (FStar_Reflection_Data.Tv_FVar + (FStar_Reflection_Builtins.pack_fv ["Steel"; "Effect"; "Common"; - "equiv_can_be_split"])))) - else - if - FStar_Reflection_Derived.is_fvar - hd - "Steel.Effect.Common.can_be_split_forall" - then - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2752)) - (Prims.of_int (8)) - (Prims.of_int (2752)) - (Prims.of_int (32))) - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2753)) - (Prims.of_int (8)) - (Prims.of_int (2753)) - (Prims.of_int (41))) - (Obj.magic - (FStar_Tactics_Effect.tac_bind + "equiv_can_be_split"]))))) + else + Obj.magic + (Obj.repr + (if + FStar_Reflection_Derived.is_fvar + hd + "Steel.Effect.Common.can_be_split_forall" + then + Obj.repr ( - Prims.mk_range + FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2752)) + (Prims.of_int (8)) + (Prims.of_int (2752)) + (Prims.of_int (32))) + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2753)) + (Prims.of_int (8)) + (Prims.of_int (2753)) + (Prims.of_int (41))) + (Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2752)) (Prims.of_int (15)) (Prims.of_int (2752)) (Prims.of_int (32))) - ( - Prims.mk_range + (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2752)) (Prims.of_int (8)) (Prims.of_int (2752)) (Prims.of_int (32))) - ( - Obj.magic + (Obj.magic (FStar_Tactics_Logic.forall_intro ())) - ( - fun + (fun uu___3 -> FStar_Tactics_Effect.lift_div_tac (fun uu___4 -> ())))) - (fun uu___3 -> - (fun uu___3 - -> + (fun + uu___3 -> + (fun + uu___3 -> Obj.magic (FStar_Tactics_Derived.apply_lemma (FStar_Reflection_Builtins.pack_ln @@ -12043,29 +12150,31 @@ let (goal_to_equiv : "Effect"; "Common"; "equiv_can_be_split"]))))) - uu___3)) - else - if - FStar_Reflection_Derived.is_fvar - hd - "Steel.Effect.Common.equiv_forall" - then - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2755)) - (Prims.of_int (8)) - (Prims.of_int (2755)) - (Prims.of_int (40))) - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2756)) - (Prims.of_int (8)) - (Prims.of_int (2756)) - (Prims.of_int (32))) - (Obj.magic - (FStar_Tactics_Derived.apply_lemma + uu___3)) + else + Obj.repr + ( + if + FStar_Reflection_Derived.is_fvar + hd + "Steel.Effect.Common.equiv_forall" + then + Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2755)) + (Prims.of_int (8)) + (Prims.of_int (2755)) + (Prims.of_int (40))) + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2756)) + (Prims.of_int (8)) + (Prims.of_int (2756)) + (Prims.of_int (32))) + (Obj.magic + (FStar_Tactics_Derived.apply_lemma (FStar_Reflection_Builtins.pack_ln (FStar_Reflection_Data.Tv_FVar (FStar_Reflection_Builtins.pack_fv @@ -12073,9 +12182,9 @@ let (goal_to_equiv : "Effect"; "Common"; "equiv_forall_elim"]))))) - (fun uu___4 - -> - (fun + (fun + uu___4 -> + (fun uu___4 -> Obj.magic (FStar_Tactics_Effect.tac_bind @@ -12101,27 +12210,28 @@ let (goal_to_equiv : uu___6 -> ())))) uu___4)) - else - if - FStar_Reflection_Derived.is_fvar - hd - "Steel.Effect.Common.can_be_split_post" - then - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range + else + Obj.repr + (if + FStar_Reflection_Derived.is_fvar + hd + "Steel.Effect.Common.can_be_split_post" + then + Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2758)) (Prims.of_int (8)) (Prims.of_int (2758)) (Prims.of_int (45))) - (Prims.mk_range + (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2759)) (Prims.of_int (8)) (Prims.of_int (2761)) (Prims.of_int (32))) - (Obj.magic + (Obj.magic (FStar_Tactics_Derived.apply_lemma (FStar_Reflection_Builtins.pack_ln (FStar_Reflection_Data.Tv_FVar @@ -12130,8 +12240,8 @@ let (goal_to_equiv : "Effect"; "Common"; "can_be_split_post_elim"]))))) - (fun uu___5 - -> + (fun + uu___5 -> (fun uu___5 -> Obj.magic @@ -12222,41 +12332,39 @@ let (goal_to_equiv : uu___7))) uu___6))) uu___5)) - else - if - FStar_Reflection_Derived.is_fvar - hd - "Steel.Effect.Common.can_be_split_dep" - then - Obj.magic - (FStar_Tactics_Derived.fail - ( - Prims.strcat + else + Obj.repr + (if + FStar_Reflection_Derived.is_fvar + hd + "Steel.Effect.Common.can_be_split_dep" + then + FStar_Tactics_Derived.fail + (Prims.strcat "can_be_split_dep not supported in " - loc)) - else - if - FStar_Reflection_Derived.is_fvar - hd - "Steel.Effect.Common.can_be_split_forall_dep" - then - Obj.magic - (FStar_Tactics_Derived.fail + loc) + else + if + FStar_Reflection_Derived.is_fvar + hd + "Steel.Effect.Common.can_be_split_forall_dep" + then + FStar_Tactics_Derived.fail (Prims.strcat "can_be_split_forall_dep not supported in " - loc)) - else - Obj.magic - (FStar_Tactics_Derived.fail + loc) + else + FStar_Tactics_Derived.fail (Prims.strcat loc - " goal in unexpected position"))) - uu___1))) uu___)) + " goal in unexpected position"))))))) + uu___1))) uu___))) | uu___ -> Obj.magic - (FStar_Tactics_Derived.fail - (Prims.strcat loc " unexpected goal"))) uu___))) - uu___) + (Obj.repr + (FStar_Tactics_Derived.fail + (Prims.strcat loc " unexpected goal")))) + uu___))) uu___) let rec term_dict_assoc : 'a . FStar_Reflection_Types.term -> @@ -13311,11 +13419,13 @@ let rec (resolve_tac : if uu___2 then Obj.magic - (resolve_tac dict) + (Obj.repr + (resolve_tac dict)) else Obj.magic - (FStar_Tactics_Derived.fail - "Could not make progress, no solvable goal found")) + (Obj.repr + (FStar_Tactics_Derived.fail + "Could not make progress, no solvable goal found"))) uu___2))) uu___2))) uu___1)))) uu___) let rec (pick_next_logical : @@ -13604,14 +13714,16 @@ let (is_true : match uu___1 with | FStar_Reflection_Formula.True_ -> Obj.magic - (FStar_Tactics_Derived.exact - (FStar_Reflection_Builtins.pack_ln - (FStar_Reflection_Data.Tv_Const - FStar_Reflection_Data.C_Unit))) + (Obj.repr + (FStar_Tactics_Derived.exact + (FStar_Reflection_Builtins.pack_ln + (FStar_Reflection_Data.Tv_Const + FStar_Reflection_Data.C_Unit)))) | uu___2 -> Obj.magic - (FStar_Tactics_Effect.raise - FStar_Tactics_Derived.Goal_not_trivial)) uu___1) + (Obj.repr + (FStar_Tactics_Effect.raise + FStar_Tactics_Derived.Goal_not_trivial))) uu___1) let rec (solve_maybe_emps : Prims.nat -> (unit, unit) FStar_Tactics_Effect.tac_repr) = fun uu___ -> @@ -14529,11 +14641,13 @@ let (ite_soundness_tac : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) match uu___4 with | [] -> Obj.magic - (FStar_Tactics_Derived.fail - "should not happen") + (Obj.repr + (FStar_Tactics_Derived.fail + "should not happen")) | uu___5::tl -> Obj.magic - (FStar_Tactics_Effect.tac_bind + (Obj.repr + (FStar_Tactics_Effect.tac_bind (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (3153)) @@ -14751,7 +14865,7 @@ let (ite_soundness_tac : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) uu___9))) uu___8))) uu___7))) - uu___6))) + uu___6)))) uu___4))) uu___4))) uu___3))) uu___2))) uu___1) let (vc_norm : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = diff --git a/ocaml/fstar-lib/generated/Steel_ST_GenElim_Base.ml b/ocaml/fstar-lib/generated/Steel_ST_GenElim_Base.ml index ad134920e48..2ca503ea1dc 100644 --- a/ocaml/fstar-lib/generated/Steel_ST_GenElim_Base.ml +++ b/ocaml/fstar-lib/generated/Steel_ST_GenElim_Base.ml @@ -240,24 +240,28 @@ let rec (solve_gen_unit_elim : (t2, FStar_Reflection_Data.Q_Explicit)::[] -> - FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.ST.GenElim.Base.fsti" - (Prims.of_int (335)) - (Prims.of_int (20)) - (Prims.of_int (335)) - (Prims.of_int (42))) - (Prims.mk_range - "Steel.ST.GenElim.Base.fsti" - (Prims.of_int (336)) - (Prims.of_int (10)) - (Prims.of_int (337)) - (Prims.of_int (68))) - (Obj.magic - (solve_gen_unit_elim + Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.ST.GenElim.Base.fsti" + (Prims.of_int (335)) + (Prims.of_int (20)) + (Prims.of_int (335)) + (Prims.of_int (42))) + (Prims.mk_range + "Steel.ST.GenElim.Base.fsti" + (Prims.of_int (336)) + (Prims.of_int (10)) + (Prims.of_int (337)) + (Prims.of_int (68))) + (Obj.magic + ( + solve_gen_unit_elim t1)) - (fun uu___5 -> - (fun t1' -> + (fun uu___5 + -> + (fun t1' + -> Obj.magic (FStar_Tactics_Effect.tac_bind (Prims.mk_range @@ -294,10 +298,11 @@ let rec (solve_gen_unit_elim : FStar_Reflection_Data.Q_Explicit); (t2', FStar_Reflection_Data.Q_Explicit)])))) - uu___5) + uu___5)) | uu___5 -> - FStar_Tactics_Derived.fail - "ill-formed star")) + Obj.repr + (FStar_Tactics_Derived.fail + "ill-formed star"))) else Obj.magic (Obj.repr @@ -429,25 +434,19 @@ let rec (solve_gen_elim : (body, FStar_Reflection_Data.Q_Explicit)::[] -> - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___3 -> - ([(ty, - FStar_Reflection_Data.Q_Implicit)], - body)))) + FStar_Tactics_Effect.lift_div_tac + (fun uu___3 -> + ([(ty, + FStar_Reflection_Data.Q_Implicit)], + body)) | (body, FStar_Reflection_Data.Q_Explicit)::[] -> - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___3 -> ([], body)))) + FStar_Tactics_Effect.lift_div_tac + (fun uu___3 -> ([], body)) | uu___3 -> - Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "ill-formed exists_"))) + FStar_Tactics_Derived.fail + "ill-formed exists_") (fun uu___3 -> (fun uu___3 -> match uu___3 with @@ -805,35 +804,37 @@ let rec (solve_gen_elim : (tr, FStar_Reflection_Data.Q_Explicit)::[] -> - FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.ST.GenElim.Base.fsti" - (Prims.of_int (383)) - (Prims.of_int (15)) - (Prims.of_int (383)) - (Prims.of_int (42))) - (Prims.mk_range - "Steel.ST.GenElim.Base.fsti" - (Prims.of_int (383)) - (Prims.of_int (12)) - (Prims.of_int (396)) - (Prims.of_int (72))) - (Obj.magic - (term_has_head tl - (FStar_Reflection_Builtins.pack_ln - (FStar_Reflection_Data.Tv_FVar + Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.ST.GenElim.Base.fsti" + (Prims.of_int (383)) + (Prims.of_int (15)) + (Prims.of_int (383)) + (Prims.of_int (42))) + (Prims.mk_range + "Steel.ST.GenElim.Base.fsti" + (Prims.of_int (383)) + (Prims.of_int (12)) + (Prims.of_int (396)) + (Prims.of_int (72))) + (Obj.magic + (term_has_head + tl + (FStar_Reflection_Builtins.pack_ln ( - FStar_Reflection_Builtins.pack_fv + FStar_Reflection_Data.Tv_FVar + (FStar_Reflection_Builtins.pack_fv ["Steel"; "ST"; "Util"; "exists_"]))))) - (fun uu___5 -> (fun uu___5 -> - if uu___5 - then - Obj.magic - (FStar_Tactics_Effect.tac_bind + (fun uu___5 -> + if uu___5 + then + Obj.magic + (FStar_Tactics_Effect.tac_bind (Prims.mk_range "Steel.ST.GenElim.Base.fsti" (Prims.of_int (385)) @@ -958,9 +959,9 @@ let rec (solve_gen_elim : FStar_Reflection_Data.Q_Explicit)])))) uu___6))) uu___6)) - else - Obj.magic - (FStar_Tactics_Effect.tac_bind + else + Obj.magic + (FStar_Tactics_Effect.tac_bind (Prims.mk_range "Steel.ST.GenElim.Base.fsti" (Prims.of_int (394)) @@ -1017,10 +1018,11 @@ let rec (solve_gen_elim : (tr', FStar_Reflection_Data.Q_Explicit)])))) uu___7))) - uu___5) + uu___5)) | uu___5 -> - FStar_Tactics_Derived.fail - "ill-formed star")) + Obj.repr + (FStar_Tactics_Derived.fail + "ill-formed star"))) else Obj.magic (Obj.repr @@ -1782,78 +1784,70 @@ let (solve_gen_elim_prop : FStar_Tactics_Effect.lift_div_tac (fun uu___3 -> Prims.op_Negation uu___2)))) (fun uu___2 -> - (fun uu___2 -> - if uu___2 - then - Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "not a squash goal")) - else - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___4 -> ())))) uu___2))) + if uu___2 + then + FStar_Tactics_Derived.fail + "not a squash goal" + else + FStar_Tactics_Effect.lift_div_tac + (fun uu___4 -> ())))) (fun uu___2 -> (fun uu___2 -> match tl with | (body1, FStar_Reflection_Data.Q_Explicit)::[] -> Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.ST.GenElim.Base.fsti" - (Prims.of_int (571)) - (Prims.of_int (21)) - (Prims.of_int (571)) - (Prims.of_int (40))) - (Prims.mk_range - "Steel.ST.GenElim.Base.fsti" - (Prims.of_int (571)) - (Prims.of_int (4)) - (Prims.of_int (603)) - (Prims.of_int (7))) - (FStar_Tactics_Effect.lift_div_tac - (fun uu___3 -> - FStar_Reflection_Derived.collect_app - body1)) - (fun uu___3 -> + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.ST.GenElim.Base.fsti" + (Prims.of_int (571)) + (Prims.of_int (21)) + (Prims.of_int (571)) + (Prims.of_int (40))) + (Prims.mk_range + "Steel.ST.GenElim.Base.fsti" + (Prims.of_int (571)) + (Prims.of_int (4)) + (Prims.of_int (603)) + (Prims.of_int (7))) + (FStar_Tactics_Effect.lift_div_tac + (fun uu___3 -> + FStar_Reflection_Derived.collect_app + body1)) (fun uu___3 -> - match uu___3 with - | (hd1, tl1) -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.ST.GenElim.Base.fsti" - (Prims.of_int (572)) - (Prims.of_int (4)) - (Prims.of_int (573)) - (Prims.of_int (42))) - (Prims.mk_range - "Steel.ST.GenElim.Base.fsti" - (Prims.of_int (574)) - (Prims.of_int (10)) - (Prims.of_int (602)) - (Prims.of_int (44))) - (if - Prims.op_Negation - (is_fvar hd1 - "Steel.ST.GenElim.Base.gen_elim_prop") - then - Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "not a gen_elim_prop goal")) - else - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___5 -> - ())))) - (fun uu___4 -> + (fun uu___3 -> + match uu___3 with + | (hd1, tl1) -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.ST.GenElim.Base.fsti" + (Prims.of_int (572)) + (Prims.of_int (4)) + (Prims.of_int (573)) + (Prims.of_int (42))) + (Prims.mk_range + "Steel.ST.GenElim.Base.fsti" + (Prims.of_int (574)) + (Prims.of_int (10)) + (Prims.of_int (602)) + (Prims.of_int (44))) + (if + Prims.op_Negation + (is_fvar hd1 + "Steel.ST.GenElim.Base.gen_elim_prop") + then + FStar_Tactics_Derived.fail + "not a gen_elim_prop goal" + else + FStar_Tactics_Effect.lift_div_tac + (fun uu___5 -> ())) (fun uu___4 -> - match FStar_List_Tot_Base.filter - (fun uu___5 - -> + (fun uu___4 -> + match FStar_List_Tot_Base.filter + ( + fun + uu___5 -> match uu___5 with | @@ -1861,35 +1855,37 @@ let (solve_gen_elim_prop : x) -> FStar_Reflection_Data.uu___is_Q_Explicit x) tl1 - with - | (enable_nondep_opt_tm, - uu___5):: - (p, uu___6):: - (a, uu___7):: - (q, uu___8):: - (post, uu___9)::[] - -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range + with + | (enable_nondep_opt_tm, + uu___5):: + (p, uu___6):: + (a, uu___7):: + (q, uu___8):: + (post, + uu___9)::[] + -> + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.ST.GenElim.Base.fsti" (Prims.of_int (576)) (Prims.of_int (30)) (Prims.of_int (576)) (Prims.of_int (74))) - (Prims.mk_range + (Prims.mk_range "Steel.ST.GenElim.Base.fsti" (Prims.of_int (577)) (Prims.of_int (6)) (Prims.of_int (601)) (Prims.of_int (44))) - (Obj.magic + (Obj.magic (FStar_Tactics_Builtins.term_eq_old enable_nondep_opt_tm (FStar_Reflection_Builtins.pack_ln (FStar_Reflection_Data.Tv_Const FStar_Reflection_Data.C_True)))) - (fun + (fun uu___10 -> (fun @@ -2416,16 +2412,19 @@ let (solve_gen_elim_prop : uu___10))) uu___10))) uu___10))) - uu___10)) - | uu___5 -> - Obj.magic - (FStar_Tactics_Derived.fail - "ill formed gen_elim_prop")) - uu___4))) uu___3)) + uu___10))) + | uu___5 -> + Obj.magic + (Obj.repr + (FStar_Tactics_Derived.fail + "ill formed gen_elim_prop"))) + uu___4))) uu___3))) | uu___3 -> Obj.magic - (FStar_Tactics_Derived.fail - "ill-formed squash")) uu___2))) uu___1) + (Obj.repr + (FStar_Tactics_Derived.fail + "ill-formed squash"))) uu___2))) + uu___1) let (solve_gen_elim_prop_placeholder : unit -> (Prims.bool, unit) FStar_Tactics_Effect.tac_repr) = fun uu___ -> @@ -2477,78 +2476,70 @@ let (solve_gen_elim_prop_placeholder : FStar_Tactics_Effect.lift_div_tac (fun uu___3 -> Prims.op_Negation uu___2)))) (fun uu___2 -> - (fun uu___2 -> - if uu___2 - then - Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "not a squash goal")) - else - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___4 -> ())))) uu___2))) + if uu___2 + then + FStar_Tactics_Derived.fail + "not a squash goal" + else + FStar_Tactics_Effect.lift_div_tac + (fun uu___4 -> ())))) (fun uu___2 -> (fun uu___2 -> match tl with | (body1, FStar_Reflection_Data.Q_Explicit)::[] -> Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.ST.GenElim.Base.fsti" - (Prims.of_int (615)) - (Prims.of_int (21)) - (Prims.of_int (615)) - (Prims.of_int (40))) - (Prims.mk_range - "Steel.ST.GenElim.Base.fsti" - (Prims.of_int (615)) - (Prims.of_int (4)) - (Prims.of_int (643)) - (Prims.of_int (7))) - (FStar_Tactics_Effect.lift_div_tac - (fun uu___3 -> - FStar_Reflection_Derived.collect_app - body1)) - (fun uu___3 -> + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.ST.GenElim.Base.fsti" + (Prims.of_int (615)) + (Prims.of_int (21)) + (Prims.of_int (615)) + (Prims.of_int (40))) + (Prims.mk_range + "Steel.ST.GenElim.Base.fsti" + (Prims.of_int (615)) + (Prims.of_int (4)) + (Prims.of_int (643)) + (Prims.of_int (7))) + (FStar_Tactics_Effect.lift_div_tac + (fun uu___3 -> + FStar_Reflection_Derived.collect_app + body1)) (fun uu___3 -> - match uu___3 with - | (hd1, tl1) -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.ST.GenElim.Base.fsti" - (Prims.of_int (616)) - (Prims.of_int (4)) - (Prims.of_int (617)) - (Prims.of_int (54))) - (Prims.mk_range - "Steel.ST.GenElim.Base.fsti" - (Prims.of_int (618)) - (Prims.of_int (10)) - (Prims.of_int (642)) - (Prims.of_int (56))) - (if - Prims.op_Negation - (is_fvar hd1 - "Steel.ST.GenElim.Base.gen_elim_prop_placeholder") - then - Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "not a gen_elim_prop_placeholder goal")) - else - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___5 -> - ())))) - (fun uu___4 -> + (fun uu___3 -> + match uu___3 with + | (hd1, tl1) -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.ST.GenElim.Base.fsti" + (Prims.of_int (616)) + (Prims.of_int (4)) + (Prims.of_int (617)) + (Prims.of_int (54))) + (Prims.mk_range + "Steel.ST.GenElim.Base.fsti" + (Prims.of_int (618)) + (Prims.of_int (10)) + (Prims.of_int (642)) + (Prims.of_int (56))) + (if + Prims.op_Negation + (is_fvar hd1 + "Steel.ST.GenElim.Base.gen_elim_prop_placeholder") + then + FStar_Tactics_Derived.fail + "not a gen_elim_prop_placeholder goal" + else + FStar_Tactics_Effect.lift_div_tac + (fun uu___5 -> ())) (fun uu___4 -> - match FStar_List_Tot_Base.filter - (fun uu___5 - -> + (fun uu___4 -> + match FStar_List_Tot_Base.filter + ( + fun + uu___5 -> match uu___5 with | @@ -2556,29 +2547,31 @@ let (solve_gen_elim_prop_placeholder : x) -> FStar_Reflection_Data.uu___is_Q_Explicit x) tl1 - with - | (enable_nondep_opt_tm, - uu___5):: - (p, uu___6):: - (a, uu___7):: - (q, uu___8):: - (post, uu___9)::[] - -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range + with + | (enable_nondep_opt_tm, + uu___5):: + (p, uu___6):: + (a, uu___7):: + (q, uu___8):: + (post, + uu___9)::[] + -> + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.ST.GenElim.Base.fsti" (Prims.of_int (620)) (Prims.of_int (6)) (Prims.of_int (621)) (Prims.of_int (47))) - (Prims.mk_range + (Prims.mk_range "Steel.ST.GenElim.Base.fsti" (Prims.of_int (622)) (Prims.of_int (6)) (Prims.of_int (641)) (Prims.of_int (10))) - (Obj.magic + (Obj.magic (FStar_Tactics_Effect.tac_bind (Prims.mk_range "Steel.ST.GenElim.Base.fsti" @@ -2622,25 +2615,17 @@ let (solve_gen_elim_prop_placeholder : (fun uu___10 -> - (fun - uu___10 - -> if uu___10 then - Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "pre-resource not solved yet")) + FStar_Tactics_Derived.fail + "pre-resource not solved yet" else - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac + FStar_Tactics_Effect.lift_div_tac (fun uu___12 -> ())))) - uu___10))) - (fun + (fun uu___10 -> (fun @@ -2806,17 +2791,13 @@ let (solve_gen_elim_prop_placeholder : && post_is_uvar) then - Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "gen_elim_prop_placeholder is already solved")) + FStar_Tactics_Derived.fail + "gen_elim_prop_placeholder is already solved" else - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac + FStar_Tactics_Effect.lift_div_tac (fun uu___12 - -> ())))) + -> ())) (fun uu___11 -> @@ -3215,16 +3196,19 @@ let (solve_gen_elim_prop_placeholder : uu___11))) uu___11))) uu___11))) - uu___10)) - | uu___5 -> - Obj.magic - (FStar_Tactics_Derived.fail - "ill-formed gen_elim_prop_placeholder")) - uu___4))) uu___3)) + uu___10))) + | uu___5 -> + Obj.magic + (Obj.repr + (FStar_Tactics_Derived.fail + "ill-formed gen_elim_prop_placeholder"))) + uu___4))) uu___3))) | uu___3 -> Obj.magic - (FStar_Tactics_Derived.fail - "ill-formed squash")) uu___2))) uu___1) + (Obj.repr + (FStar_Tactics_Derived.fail + "ill-formed squash"))) uu___2))) + uu___1) let (init_resolve_tac : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = fun uu___ -> Steel_Effect_Common.init_resolve_tac' From 290ed4fd4f3b735082fb6cc9246d814b6808c009 Mon Sep 17 00:00:00 2001 From: Aseem Rastogi Date: Mon, 20 Mar 2023 17:06:47 +0530 Subject: [PATCH 21/48] Revert "fixing an off by 1 error" This reverts commit ff030ad48d84fd5446d086b65f82ff4a90e10c31. --- ocaml/fstar-lib/FStar_Parser_ParseIt.ml | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/ocaml/fstar-lib/FStar_Parser_ParseIt.ml b/ocaml/fstar-lib/FStar_Parser_ParseIt.ml index 9a91d7fc285..9bc10ea9f5f 100644 --- a/ocaml/fstar-lib/FStar_Parser_ParseIt.ml +++ b/ocaml/fstar-lib/FStar_Parser_ParseIt.ml @@ -147,7 +147,7 @@ let parse fn = let contents_at = let lines = U.splitlines contents in fun (start_pos:Lexing.position) (end_pos:Lexing.position) -> - let suffix = FStar_Compiler_Util.nth_tail (Z.of_int start_pos.pos_lnum) lines in + let suffix = FStar_Compiler_Util.nth_tail (Z.of_int (start_pos.pos_lnum - 1)) lines in let text, _ = FStar_Compiler_Util.first_N (Z.of_int (end_pos.pos_lnum - start_pos.pos_lnum)) suffix in let range = range_of_positions start_pos end_pos in { range; From e00948ab643a663af84dd8a1076d09a03d5ccfa6 Mon Sep 17 00:00:00 2001 From: Aseem Rastogi Date: Mon, 20 Mar 2023 17:06:58 +0530 Subject: [PATCH 22/48] Revert "snap" This reverts commit 64eb234539de816eb48f80d377b417a7b769f695. --- .../FStar_InteractiveHelpers_Base.ml | 4 +- .../FStar_InteractiveHelpers_Effectful.ml | 21 +- .../FStar_InteractiveHelpers_ExploreTerm.ml | 265 +- .../FStar_InteractiveHelpers_PostProcess.ml | 109 +- ocaml/fstar-lib/generated/FStar_Tactics_BV.ml | 8 +- .../generated/FStar_Tactics_Canon.ml | 8 +- .../FStar_Tactics_CanonCommMonoid.ml | 68 +- .../FStar_Tactics_CanonCommMonoidSimple.ml | 188 +- ...tar_Tactics_CanonCommMonoidSimple_Equiv.ml | 99 +- .../FStar_Tactics_CanonCommSemiring.ml | 56 +- .../generated/FStar_Tactics_CanonMonoid.ml | 148 +- .../generated/FStar_Tactics_Derived.ml | 1062 ++++--- .../generated/FStar_Tactics_Effect.ml | 4 +- .../generated/FStar_Tactics_Logic.ml | 225 +- .../FStar_Tactics_PatternMatching.ml | 678 ++--- .../generated/FStar_Tactics_Simplifier.ml | 324 +- .../generated/FStar_Tactics_SyntaxHelpers.ml | 6 +- .../generated/FStar_Tactics_Typeclasses.ml | 73 +- .../generated/Steel_Effect_Common.ml | 2688 ++++++++--------- .../generated/Steel_ST_GenElim_Base.ml | 498 +-- 20 files changed, 3143 insertions(+), 3389 deletions(-) diff --git a/ocaml/fstar-lib/generated/FStar_InteractiveHelpers_Base.ml b/ocaml/fstar-lib/generated/FStar_InteractiveHelpers_Base.ml index e8b1b42a385..bea6412242a 100644 --- a/ocaml/fstar-lib/generated/FStar_InteractiveHelpers_Base.ml +++ b/ocaml/fstar-lib/generated/FStar_InteractiveHelpers_Base.ml @@ -946,9 +946,7 @@ let (__proj__MetaAnalysis__item__uu___ : Prims.exn -> Prims.string) = fun projectee -> match projectee with | MetaAnalysis uu___ -> uu___ let mfail : 'uuuuu . Prims.string -> ('uuuuu, unit) FStar_Tactics_Effect.tac_repr = - fun uu___ -> - (fun str -> Obj.magic (FStar_Tactics_Effect.raise (MetaAnalysis str))) - uu___ + fun str -> FStar_Tactics_Effect.raise (MetaAnalysis str) let (print_dbg : Prims.bool -> Prims.string -> (unit, unit) FStar_Tactics_Effect.tac_repr) = fun uu___1 -> diff --git a/ocaml/fstar-lib/generated/FStar_InteractiveHelpers_Effectful.ml b/ocaml/fstar-lib/generated/FStar_InteractiveHelpers_Effectful.ml index b6e37006c5d..d113c2ed518 100644 --- a/ocaml/fstar-lib/generated/FStar_InteractiveHelpers_Effectful.ml +++ b/ocaml/fstar-lib/generated/FStar_InteractiveHelpers_Effectful.ml @@ -1431,20 +1431,13 @@ let (compute_eterm_info : parameters)))) uu___2))) uu___2)) (fun uu___1 -> - (fun uu___1 -> - match uu___1 with - | FStar_Tactics_Common.TacticFailure msg -> - Obj.magic - (Obj.repr - (FStar_InteractiveHelpers_Base.mfail - (Prims.strcat - "compute_eterm_info: failure: '" - (Prims.strcat msg "'")))) - | e1 -> - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.raise e1))) - uu___1))) uu___) + match uu___1 with + | FStar_Tactics_Common.TacticFailure msg -> + FStar_InteractiveHelpers_Base.mfail + (Prims.strcat + "compute_eterm_info: failure: '" + (Prims.strcat msg "'")) + | e1 -> FStar_Tactics_Effect.raise e1))) uu___) let (has_refinement : FStar_InteractiveHelpers_ExploreTerm.type_info -> Prims.bool) = fun ty -> diff --git a/ocaml/fstar-lib/generated/FStar_InteractiveHelpers_ExploreTerm.ml b/ocaml/fstar-lib/generated/FStar_InteractiveHelpers_ExploreTerm.ml index fa2b8a6733f..7428ba9e9f6 100644 --- a/ocaml/fstar-lib/generated/FStar_InteractiveHelpers_ExploreTerm.ml +++ b/ocaml/fstar-lib/generated/FStar_InteractiveHelpers_ExploreTerm.ml @@ -1556,21 +1556,13 @@ let rec (inst_comp : match () with | () -> inst_comp_once e c t) (fun uu___ -> - (fun uu___ -> - match uu___ with - | FStar_InteractiveHelpers_Base.MetaAnalysis - msg -> - Obj.magic - (Obj.repr - (FStar_InteractiveHelpers_Base.mfail - (Prims.strcat - "inst_comp: error: " - msg))) - | err -> - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.raise - err))) uu___))) + match uu___ with + | FStar_InteractiveHelpers_Base.MetaAnalysis + msg -> + FStar_InteractiveHelpers_Base.mfail + (Prims.strcat "inst_comp: error: " + msg) + | err -> FStar_Tactics_Effect.raise err))) (fun uu___ -> (fun c' -> Obj.magic (inst_comp e c' tl')) uu___)))) uu___2 uu___1 uu___ @@ -1669,68 +1661,54 @@ let (_abs_update_typ : "_abs_update_typ: inconsistent state")) uu___1))) uu___1)) (fun uu___ -> - (fun uu___ -> - match uu___ with - | FStar_InteractiveHelpers_Base.MetaAnalysis msg -> - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.InteractiveHelpers.ExploreTerm.fst" - (Prims.of_int (303)) (Prims.of_int (10)) - (Prims.of_int (303)) (Prims.of_int (93))) - (Prims.mk_range - "FStar.InteractiveHelpers.ExploreTerm.fst" - (Prims.of_int (303)) (Prims.of_int (4)) - (Prims.of_int (303)) (Prims.of_int (93))) - (Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.InteractiveHelpers.ExploreTerm.fst" - (Prims.of_int (303)) - (Prims.of_int (61)) - (Prims.of_int (303)) - (Prims.of_int (92))) - (Prims.mk_range "prims.fst" - (Prims.of_int (606)) - (Prims.of_int (19)) - (Prims.of_int (606)) - (Prims.of_int (31))) - (Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.InteractiveHelpers.ExploreTerm.fst" - (Prims.of_int (303)) - (Prims.of_int (61)) - (Prims.of_int (303)) - (Prims.of_int (78))) - (Prims.mk_range "prims.fst" - (Prims.of_int (606)) - (Prims.of_int (19)) - (Prims.of_int (606)) - (Prims.of_int (31))) - (Obj.magic - (FStar_Tactics_Builtins.term_to_string - ty)) - (fun uu___1 -> - FStar_Tactics_Effect.lift_div_tac - (fun uu___2 -> - Prims.strcat uu___1 - (Prims.strcat ":\n" msg))))) - (fun uu___1 -> - FStar_Tactics_Effect.lift_div_tac - (fun uu___2 -> - Prims.strcat - "_abs_update_typ: could not find an arrow in: " - uu___1)))) - (fun uu___1 -> + match uu___ with + | FStar_InteractiveHelpers_Base.MetaAnalysis msg -> + FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.InteractiveHelpers.ExploreTerm.fst" + (Prims.of_int (303)) (Prims.of_int (10)) + (Prims.of_int (303)) (Prims.of_int (93))) + (Prims.mk_range + "FStar.InteractiveHelpers.ExploreTerm.fst" + (Prims.of_int (303)) (Prims.of_int (4)) + (Prims.of_int (303)) (Prims.of_int (93))) + (Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.InteractiveHelpers.ExploreTerm.fst" + (Prims.of_int (303)) (Prims.of_int (61)) + (Prims.of_int (303)) (Prims.of_int (92))) + (Prims.mk_range "prims.fst" (Prims.of_int (606)) + (Prims.of_int (19)) (Prims.of_int (606)) + (Prims.of_int (31))) + (Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.InteractiveHelpers.ExploreTerm.fst" + (Prims.of_int (303)) (Prims.of_int (61)) + (Prims.of_int (303)) (Prims.of_int (78))) + (Prims.mk_range "prims.fst" + (Prims.of_int (606)) (Prims.of_int (19)) + (Prims.of_int (606)) (Prims.of_int (31))) + (Obj.magic + (FStar_Tactics_Builtins.term_to_string ty)) (fun uu___1 -> - Obj.magic - (FStar_InteractiveHelpers_Base.mfail - uu___1)) uu___1))) - | err -> - Obj.magic (Obj.repr (FStar_Tactics_Effect.raise err))) - uu___) + FStar_Tactics_Effect.lift_div_tac + (fun uu___2 -> + Prims.strcat uu___1 + (Prims.strcat ":\n" msg))))) + (fun uu___1 -> + FStar_Tactics_Effect.lift_div_tac + (fun uu___2 -> + Prims.strcat + "_abs_update_typ: could not find an arrow in: " + uu___1)))) + (fun uu___1 -> + (fun uu___1 -> + Obj.magic + (FStar_InteractiveHelpers_Base.mfail uu___1)) + uu___1) + | err -> FStar_Tactics_Effect.raise err) let (abs_update_typ_or_comp : FStar_Reflection_Types.binder -> typ_or_comp -> @@ -1802,12 +1780,14 @@ let (abs_update_opt_typ_or_comp : | FStar_InteractiveHelpers_Base.MetaAnalysis msg -> Obj.magic - (FStar_Tactics_Effect.lift_div_tac - (fun uu___1 -> - FStar_Pervasives_Native.None)) + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___1 -> + FStar_Pervasives_Native.None))) | err -> Obj.magic - (FStar_Tactics_Effect.raise err)) + (Obj.repr + (FStar_Tactics_Effect.raise err))) uu___)))) uu___2 uu___1 uu___ let rec (_flush_typ_or_comp_comp : Prims.bool -> @@ -2335,77 +2315,62 @@ let (flush_typ_or_comp : uu___1) | TC_Comp (c, pl, n) -> flush_comp pl n c)) (fun uu___ -> - (fun uu___ -> - match uu___ with - | FStar_InteractiveHelpers_Base.MetaAnalysis msg - -> - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.InteractiveHelpers.ExploreTerm.fst" - (Prims.of_int (379)) - (Prims.of_int (15)) - (Prims.of_int (379)) - (Prims.of_int (90))) - (Prims.mk_range - "FStar.InteractiveHelpers.ExploreTerm.fst" - (Prims.of_int (379)) - (Prims.of_int (9)) - (Prims.of_int (379)) - (Prims.of_int (90))) - (Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.InteractiveHelpers.ExploreTerm.fst" - (Prims.of_int (379)) - (Prims.of_int (50)) - (Prims.of_int (379)) - (Prims.of_int (89))) - (Prims.mk_range "prims.fst" - (Prims.of_int (606)) - (Prims.of_int (19)) - (Prims.of_int (606)) - (Prims.of_int (31))) - (Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.InteractiveHelpers.ExploreTerm.fst" - (Prims.of_int (379)) - (Prims.of_int (50)) - (Prims.of_int (379)) - (Prims.of_int (75))) - (Prims.mk_range - "prims.fst" - (Prims.of_int (606)) - (Prims.of_int (19)) - (Prims.of_int (606)) - (Prims.of_int (31))) - (Obj.magic - (typ_or_comp_to_string - tyc)) - (fun uu___1 -> - FStar_Tactics_Effect.lift_div_tac - (fun uu___2 -> - Prims.strcat - uu___1 - (Prims.strcat - ":\n" msg))))) - (fun uu___1 -> - FStar_Tactics_Effect.lift_div_tac - (fun uu___2 -> - Prims.strcat - "flush_typ_or_comp failed on: " - uu___1)))) - (fun uu___1 -> + match uu___ with + | FStar_InteractiveHelpers_Base.MetaAnalysis msg -> + FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.InteractiveHelpers.ExploreTerm.fst" + (Prims.of_int (379)) (Prims.of_int (15)) + (Prims.of_int (379)) (Prims.of_int (90))) + (Prims.mk_range + "FStar.InteractiveHelpers.ExploreTerm.fst" + (Prims.of_int (379)) (Prims.of_int (9)) + (Prims.of_int (379)) (Prims.of_int (90))) + (Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.InteractiveHelpers.ExploreTerm.fst" + (Prims.of_int (379)) + (Prims.of_int (50)) + (Prims.of_int (379)) + (Prims.of_int (89))) + (Prims.mk_range "prims.fst" + (Prims.of_int (606)) + (Prims.of_int (19)) + (Prims.of_int (606)) + (Prims.of_int (31))) + (Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.InteractiveHelpers.ExploreTerm.fst" + (Prims.of_int (379)) + (Prims.of_int (50)) + (Prims.of_int (379)) + (Prims.of_int (75))) + (Prims.mk_range "prims.fst" + (Prims.of_int (606)) + (Prims.of_int (19)) + (Prims.of_int (606)) + (Prims.of_int (31))) + (Obj.magic + (typ_or_comp_to_string tyc)) (fun uu___1 -> - Obj.magic - (FStar_InteractiveHelpers_Base.mfail - uu___1)) uu___1))) - | err -> - Obj.magic - (Obj.repr (FStar_Tactics_Effect.raise err))) - uu___))) uu___) + FStar_Tactics_Effect.lift_div_tac + (fun uu___2 -> + Prims.strcat uu___1 + (Prims.strcat ":\n" msg))))) + (fun uu___1 -> + FStar_Tactics_Effect.lift_div_tac + (fun uu___2 -> + Prims.strcat + "flush_typ_or_comp failed on: " + uu___1)))) + (fun uu___1 -> + (fun uu___1 -> + Obj.magic + (FStar_InteractiveHelpers_Base.mfail + uu___1)) uu___1) + | err -> FStar_Tactics_Effect.raise err))) uu___) let (safe_arg_typ_or_comp : Prims.bool -> FStar_Reflection_Types.env -> diff --git a/ocaml/fstar-lib/generated/FStar_InteractiveHelpers_PostProcess.ml b/ocaml/fstar-lib/generated/FStar_InteractiveHelpers_PostProcess.ml index a2430d7f851..a3e884e6b41 100644 --- a/ocaml/fstar-lib/generated/FStar_InteractiveHelpers_PostProcess.ml +++ b/ocaml/fstar-lib/generated/FStar_InteractiveHelpers_PostProcess.ml @@ -2146,29 +2146,22 @@ let (pp_analyze_effectful_term : (fun uu___2 -> Obj.magic (end_proof ())) uu___2))) uu___2)) (fun uu___1 -> - (fun uu___1 -> - match uu___1 with - | FStar_InteractiveHelpers_Base.MetaAnalysis msg -> - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.InteractiveHelpers.PostProcess.fst" - (Prims.of_int (256)) (Prims.of_int (29)) - (Prims.of_int (256)) (Prims.of_int (49))) - (Prims.mk_range - "FStar.InteractiveHelpers.PostProcess.fst" - (Prims.of_int (256)) (Prims.of_int (51)) - (Prims.of_int (256)) (Prims.of_int (63))) - (Obj.magic - (FStar_InteractiveHelpers_Output.printout_failure - msg)) - (fun uu___2 -> - (fun uu___2 -> Obj.magic (end_proof ())) - uu___2))) - | err -> - Obj.magic (Obj.repr (FStar_Tactics_Effect.raise err))) - uu___1) + match uu___1 with + | FStar_InteractiveHelpers_Base.MetaAnalysis msg -> + FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.InteractiveHelpers.PostProcess.fst" + (Prims.of_int (256)) (Prims.of_int (29)) + (Prims.of_int (256)) (Prims.of_int (49))) + (Prims.mk_range + "FStar.InteractiveHelpers.PostProcess.fst" + (Prims.of_int (256)) (Prims.of_int (51)) + (Prims.of_int (256)) (Prims.of_int (63))) + (Obj.magic + (FStar_InteractiveHelpers_Output.printout_failure msg)) + (fun uu___2 -> + (fun uu___2 -> Obj.magic (end_proof ())) uu___2) + | err -> FStar_Tactics_Effect.raise err) let _ = FStar_Tactics_Native.register_tactic "FStar.InteractiveHelpers.PostProcess.pp_analyze_effectful_term" @@ -2612,27 +2605,20 @@ let (pp_split_assert_conjs : (fun uu___2 -> Obj.magic (end_proof ())) uu___2))) uu___2)) (fun uu___1 -> - (fun uu___1 -> - match uu___1 with - | FStar_InteractiveHelpers_Base.MetaAnalysis msg -> - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.InteractiveHelpers.PostProcess.fst" - (Prims.of_int (356)) (Prims.of_int (29)) - (Prims.of_int (356)) (Prims.of_int (49))) - (Prims.mk_range - "FStar.InteractiveHelpers.PostProcess.fst" - (Prims.of_int (356)) (Prims.of_int (51)) - (Prims.of_int (356)) (Prims.of_int (63))) - (Obj.magic - (FStar_InteractiveHelpers_Output.printout_failure - msg)) - (fun uu___2 -> - (fun uu___2 -> Obj.magic (end_proof ())) uu___2))) - | err -> Obj.magic (Obj.repr (FStar_Tactics_Effect.raise err))) - uu___1) + match uu___1 with + | FStar_InteractiveHelpers_Base.MetaAnalysis msg -> + FStar_Tactics_Effect.tac_bind + (Prims.mk_range "FStar.InteractiveHelpers.PostProcess.fst" + (Prims.of_int (356)) (Prims.of_int (29)) + (Prims.of_int (356)) (Prims.of_int (49))) + (Prims.mk_range "FStar.InteractiveHelpers.PostProcess.fst" + (Prims.of_int (356)) (Prims.of_int (51)) + (Prims.of_int (356)) (Prims.of_int (63))) + (Obj.magic + (FStar_InteractiveHelpers_Output.printout_failure msg)) + (fun uu___2 -> + (fun uu___2 -> Obj.magic (end_proof ())) uu___2) + | err -> FStar_Tactics_Effect.raise err) let _ = FStar_Tactics_Native.register_tactic "FStar.InteractiveHelpers.PostProcess.pp_split_assert_conjs" @@ -6927,27 +6913,20 @@ let (pp_unfold_in_assert_or_assume : (fun uu___2 -> Obj.magic (end_proof ())) uu___2))) uu___2)) (fun uu___1 -> - (fun uu___1 -> - match uu___1 with - | FStar_InteractiveHelpers_Base.MetaAnalysis msg -> - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.InteractiveHelpers.PostProcess.fst" - (Prims.of_int (748)) (Prims.of_int (29)) - (Prims.of_int (748)) (Prims.of_int (49))) - (Prims.mk_range - "FStar.InteractiveHelpers.PostProcess.fst" - (Prims.of_int (748)) (Prims.of_int (51)) - (Prims.of_int (748)) (Prims.of_int (63))) - (Obj.magic - (FStar_InteractiveHelpers_Output.printout_failure - msg)) - (fun uu___2 -> - (fun uu___2 -> Obj.magic (end_proof ())) uu___2))) - | err -> Obj.magic (Obj.repr (FStar_Tactics_Effect.raise err))) - uu___1) + match uu___1 with + | FStar_InteractiveHelpers_Base.MetaAnalysis msg -> + FStar_Tactics_Effect.tac_bind + (Prims.mk_range "FStar.InteractiveHelpers.PostProcess.fst" + (Prims.of_int (748)) (Prims.of_int (29)) + (Prims.of_int (748)) (Prims.of_int (49))) + (Prims.mk_range "FStar.InteractiveHelpers.PostProcess.fst" + (Prims.of_int (748)) (Prims.of_int (51)) + (Prims.of_int (748)) (Prims.of_int (63))) + (Obj.magic + (FStar_InteractiveHelpers_Output.printout_failure msg)) + (fun uu___2 -> + (fun uu___2 -> Obj.magic (end_proof ())) uu___2) + | err -> FStar_Tactics_Effect.raise err) let _ = FStar_Tactics_Native.register_tactic "FStar.InteractiveHelpers.PostProcess.pp_unfold_in_assert_or_assume" diff --git a/ocaml/fstar-lib/generated/FStar_Tactics_BV.ml b/ocaml/fstar-lib/generated/FStar_Tactics_BV.ml index 496adf26e40..cc7174ea77a 100644 --- a/ocaml/fstar-lib/generated/FStar_Tactics_BV.ml +++ b/ocaml/fstar-lib/generated/FStar_Tactics_BV.ml @@ -892,9 +892,11 @@ let (arith_to_bv_tac : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = "arith_to_bv_tac: unexpected: " uu___4)))) (fun uu___4 -> - FStar_Tactics_Derived.fail - uu___4))) uu___3))) - uu___3))) uu___2)) + (fun uu___4 -> + Obj.magic + (FStar_Tactics_Derived.fail + uu___4)) uu___4))) + uu___3))) uu___3))) uu___2)) let (bv_tac : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = fun uu___ -> FStar_Tactics_Derived.focus diff --git a/ocaml/fstar-lib/generated/FStar_Tactics_Canon.ml b/ocaml/fstar-lib/generated/FStar_Tactics_Canon.ml index ef14886be48..93e57a67c65 100644 --- a/ocaml/fstar-lib/generated/FStar_Tactics_Canon.ml +++ b/ocaml/fstar-lib/generated/FStar_Tactics_Canon.ml @@ -1028,8 +1028,10 @@ let (canon_point_entry : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) "impossible: " uu___4)))) (fun uu___4 -> - FStar_Tactics_Derived.fail - uu___4))) uu___2))) uu___2))) - uu___1) + (fun uu___4 -> + Obj.magic + (FStar_Tactics_Derived.fail + uu___4)) uu___4))) + uu___2))) uu___2))) uu___1) let (canon : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = fun uu___ -> FStar_Tactics_Derived.pointwise canon_point_entry \ No newline at end of file diff --git a/ocaml/fstar-lib/generated/FStar_Tactics_CanonCommMonoid.ml b/ocaml/fstar-lib/generated/FStar_Tactics_CanonCommMonoid.ml index a6482cacc8d..d71fd4339c0 100644 --- a/ocaml/fstar-lib/generated/FStar_Tactics_CanonCommMonoid.ml +++ b/ocaml/fstar-lib/generated/FStar_Tactics_CanonCommMonoid.ml @@ -1669,30 +1669,28 @@ let canon_monoid_aux : t), t1, t2) -> Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.CanonCommMonoid.fst" - (Prims.of_int (340)) - (Prims.of_int (9)) - (Prims.of_int (340)) - (Prims.of_int (25))) - (Prims.mk_range - "FStar.Tactics.CanonCommMonoid.fst" - (Prims.of_int (340)) - (Prims.of_int (6)) - (Prims.of_int (411)) - (Prims.of_int (69))) - (Obj.magic - (FStar_Tactics_Builtins.term_eq_old - t ta)) + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.CanonCommMonoid.fst" + (Prims.of_int (340)) + (Prims.of_int (9)) + (Prims.of_int (340)) + (Prims.of_int (25))) + (Prims.mk_range + "FStar.Tactics.CanonCommMonoid.fst" + (Prims.of_int (340)) + (Prims.of_int (6)) + (Prims.of_int (411)) + (Prims.of_int (69))) + (Obj.magic + (FStar_Tactics_Builtins.term_eq_old + t ta)) + (fun uu___2 -> (fun uu___2 -> - (fun uu___2 -> - if uu___2 - then - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind + if uu___2 + then + Obj.magic + (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.CanonCommMonoid.fst" (Prims.of_int (341)) @@ -1733,7 +1731,6 @@ let canon_monoid_aux : (r1::r2::[], vm) -> Obj.magic - (Obj.repr (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.CanonCommMonoid.fst" @@ -2000,25 +1997,22 @@ let canon_monoid_aux : uu___4))) uu___4))) uu___4))) - uu___4))) + uu___4)) | uu___4 -> Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "Unexpected"))) - uu___3))) - else - Obj.magic - (Obj.repr (FStar_Tactics_Derived.fail - "Goal should be an equality at the right monoid type"))) - uu___2))) + "Unexpected")) + uu___3)) + else + Obj.magic + (FStar_Tactics_Derived.fail + "Goal should be an equality at the right monoid type")) + uu___2)) | uu___2 -> Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "Goal should be an equality"))) + (FStar_Tactics_Derived.fail + "Goal should be an equality")) uu___1))) uu___) let canon_monoid_with : 'b . diff --git a/ocaml/fstar-lib/generated/FStar_Tactics_CanonCommMonoidSimple.ml b/ocaml/fstar-lib/generated/FStar_Tactics_CanonCommMonoidSimple.ml index 16012c10b64..d4730d05f43 100644 --- a/ocaml/fstar-lib/generated/FStar_Tactics_CanonCommMonoidSimple.ml +++ b/ocaml/fstar-lib/generated/FStar_Tactics_CanonCommMonoidSimple.ml @@ -624,97 +624,89 @@ let canon_monoid : (FStar_Pervasives_Native.Some t), t1, t2) -> Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.CanonCommMonoidSimple.fst" - (Prims.of_int (255)) (Prims.of_int (9)) - (Prims.of_int (255)) - (Prims.of_int (28))) - (Prims.mk_range - "FStar.Tactics.CanonCommMonoidSimple.fst" - (Prims.of_int (255)) (Prims.of_int (6)) - (Prims.of_int (273)) - (Prims.of_int (69))) - (Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.CanonCommMonoidSimple.fst" - (Prims.of_int (255)) - (Prims.of_int (19)) - (Prims.of_int (255)) - (Prims.of_int (28))) - (Prims.mk_range - "FStar.Tactics.CanonCommMonoidSimple.fst" - (Prims.of_int (255)) - (Prims.of_int (9)) - (Prims.of_int (255)) - (Prims.of_int (28))) - (FStar_Tactics_Effect.lift_div_tac - (fun uu___2 -> - (fun uu___2 -> - Obj.magic - (failwith - "Cannot evaluate open quotation at runtime")) - uu___2)) + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.CanonCommMonoidSimple.fst" + (Prims.of_int (255)) (Prims.of_int (9)) + (Prims.of_int (255)) (Prims.of_int (28))) + (Prims.mk_range + "FStar.Tactics.CanonCommMonoidSimple.fst" + (Prims.of_int (255)) (Prims.of_int (6)) + (Prims.of_int (273)) (Prims.of_int (69))) + (Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.CanonCommMonoidSimple.fst" + (Prims.of_int (255)) + (Prims.of_int (19)) + (Prims.of_int (255)) + (Prims.of_int (28))) + (Prims.mk_range + "FStar.Tactics.CanonCommMonoidSimple.fst" + (Prims.of_int (255)) + (Prims.of_int (9)) + (Prims.of_int (255)) + (Prims.of_int (28))) + (FStar_Tactics_Effect.lift_div_tac (fun uu___2 -> (fun uu___2 -> - Obj.magic (term_eq t uu___2)) - uu___2))) - (fun uu___2 -> + Obj.magic + (failwith + "Cannot evaluate open quotation at runtime")) + uu___2)) (fun uu___2 -> - if uu___2 - then - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.CanonCommMonoidSimple.fst" - (Prims.of_int (256)) - (Prims.of_int (27)) - (Prims.of_int (256)) - (Prims.of_int (67))) - (Prims.mk_range - "FStar.Tactics.CanonCommMonoidSimple.fst" - (Prims.of_int (256)) - (Prims.of_int (8)) - (Prims.of_int (271)) - (Prims.of_int (22))) - (Obj.magic - (reification m [] - (const - (FStar_Algebra_CommMonoid.__proj__CM__item__unit - m)) t1)) - (fun uu___3 -> - (fun uu___3 -> - match uu___3 with - | (r1, ts, am) -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.CanonCommMonoidSimple.fst" - (Prims.of_int (257)) - (Prims.of_int (26)) - (Prims.of_int (257)) - (Prims.of_int (48))) - (Prims.mk_range - "FStar.Tactics.CanonCommMonoidSimple.fst" - (Prims.of_int (257)) - (Prims.of_int (8)) - (Prims.of_int (271)) - (Prims.of_int (22))) - (Obj.magic - (reification - m ts am - t2)) - (fun uu___4 - -> - (fun - uu___4 -> - match uu___4 - with - | - (r2, + (fun uu___2 -> + Obj.magic (term_eq t uu___2)) + uu___2))) + (fun uu___2 -> + (fun uu___2 -> + if uu___2 + then + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.CanonCommMonoidSimple.fst" + (Prims.of_int (256)) + (Prims.of_int (27)) + (Prims.of_int (256)) + (Prims.of_int (67))) + (Prims.mk_range + "FStar.Tactics.CanonCommMonoidSimple.fst" + (Prims.of_int (256)) + (Prims.of_int (8)) + (Prims.of_int (271)) + (Prims.of_int (22))) + (Obj.magic + (reification m [] + (const + (FStar_Algebra_CommMonoid.__proj__CM__item__unit + m)) t1)) + (fun uu___3 -> + (fun uu___3 -> + match uu___3 with + | (r1, ts, am) -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.CanonCommMonoidSimple.fst" + (Prims.of_int (257)) + (Prims.of_int (26)) + (Prims.of_int (257)) + (Prims.of_int (48))) + (Prims.mk_range + "FStar.Tactics.CanonCommMonoidSimple.fst" + (Prims.of_int (257)) + (Prims.of_int (8)) + (Prims.of_int (271)) + (Prims.of_int (22))) + (Obj.magic + (reification m + ts am t2)) + (fun uu___4 -> + (fun uu___4 -> + match uu___4 + with + | (r2, uu___5, am1) -> Obj.magic @@ -908,17 +900,15 @@ let canon_monoid : uu___8))) uu___7))) uu___6))) - uu___4))) - uu___3))) - else - Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "Goal should be an equality at the right monoid type"))) - uu___2))) + uu___4))) + uu___3)) + else + Obj.magic + (FStar_Tactics_Derived.fail + "Goal should be an equality at the right monoid type")) + uu___2)) | uu___2 -> Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "Goal should be an equality"))) uu___1))) + (FStar_Tactics_Derived.fail + "Goal should be an equality")) uu___1))) uu___) \ No newline at end of file diff --git a/ocaml/fstar-lib/generated/FStar_Tactics_CanonCommMonoidSimple_Equiv.ml b/ocaml/fstar-lib/generated/FStar_Tactics_CanonCommMonoidSimple_Equiv.ml index 5d9bec7196b..6bd65b0452a 100644 --- a/ocaml/fstar-lib/generated/FStar_Tactics_CanonCommMonoidSimple_Equiv.ml +++ b/ocaml/fstar-lib/generated/FStar_Tactics_CanonCommMonoidSimple_Equiv.ml @@ -1075,74 +1075,67 @@ let (canon_monoid : (match rel_xy with | (rel_xy1, uu___2)::[] -> Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.CanonCommMonoidSimple.Equiv.fst" - (Prims.of_int (381)) - (Prims.of_int (21)) - (Prims.of_int (381)) - (Prims.of_int (43))) - (Prims.mk_range - "FStar.Tactics.CanonCommMonoidSimple.Equiv.fst" - (Prims.of_int (380)) - (Prims.of_int (21)) - (Prims.of_int (391)) - (Prims.of_int (6))) - (FStar_Tactics_Effect.lift_div_tac - (fun uu___3 -> - FStar_Reflection_Derived_Lemmas.collect_app_ref - rel_xy1)) + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.CanonCommMonoidSimple.Equiv.fst" + (Prims.of_int (381)) + (Prims.of_int (21)) + (Prims.of_int (381)) + (Prims.of_int (43))) + (Prims.mk_range + "FStar.Tactics.CanonCommMonoidSimple.Equiv.fst" + (Prims.of_int (380)) + (Prims.of_int (21)) + (Prims.of_int (391)) + (Prims.of_int (6))) + (FStar_Tactics_Effect.lift_div_tac (fun uu___3 -> - (fun uu___3 -> - match uu___3 with - | (rel, xy) -> - if - (FStar_List_Tot_Base.length - xy) - >= - (Prims.of_int (2)) - then - Obj.magic - (Obj.repr - (match - ((FStar_List_Tot_Base.index + FStar_Reflection_Derived_Lemmas.collect_app_ref + rel_xy1)) + (fun uu___3 -> + (fun uu___3 -> + match uu___3 with + | (rel, xy) -> + if + (FStar_List_Tot_Base.length + xy) + >= + (Prims.of_int (2)) + then + (match + ((FStar_List_Tot_Base.index xy ((FStar_List_Tot_Base.length xy) - (Prims.of_int (2)))), - (FStar_List_Tot_Base.index + (FStar_List_Tot_Base.index xy ((FStar_List_Tot_Base.length xy) - Prims.int_one))) - with - | - ((lhs, - FStar_Reflection_Data.Q_Explicit), - (rhs, - FStar_Reflection_Data.Q_Explicit)) - -> - Obj.repr + with + | ((lhs, + FStar_Reflection_Data.Q_Explicit), + (rhs, + FStar_Reflection_Data.Q_Explicit)) + -> + Obj.magic (canon_lhs_rhs eq m lhs rhs) - | - uu___4 -> - Obj.repr - (FStar_Tactics_Derived.fail - "Goal should have been an application of a binary relation to 2 explicit arguments"))) - else - Obj.magic - (Obj.repr + | uu___4 -> + Obj.magic (FStar_Tactics_Derived.fail - "Goal should have been an application of a binary relation to n implicit and 2 explicit arguments"))) - uu___3))) + "Goal should have been an application of a binary relation to 2 explicit arguments")) + else + Obj.magic + (FStar_Tactics_Derived.fail + "Goal should have been an application of a binary relation to n implicit and 2 explicit arguments")) + uu___3)) | uu___2 -> Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "Goal should be squash applied to a binary relation")))) + (FStar_Tactics_Derived.fail + "Goal should be squash applied to a binary relation"))) uu___1))) uu___1))) uu___) let _ = FStar_Tactics_Native.register_tactic diff --git a/ocaml/fstar-lib/generated/FStar_Tactics_CanonCommSemiring.ml b/ocaml/fstar-lib/generated/FStar_Tactics_CanonCommSemiring.ml index 547fce6978e..a7843079ede 100644 --- a/ocaml/fstar-lib/generated/FStar_Tactics_CanonCommSemiring.ml +++ b/ocaml/fstar-lib/generated/FStar_Tactics_CanonCommSemiring.ml @@ -2136,31 +2136,29 @@ let canon_semiring_aux : t), t1, t2) -> Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.CanonCommSemiring.fst" - (Prims.of_int (1634)) - (Prims.of_int (9)) - (Prims.of_int (1634)) - (Prims.of_int (21))) - (Prims.mk_range - "FStar.Tactics.CanonCommSemiring.fst" - (Prims.of_int (1634)) - (Prims.of_int (6)) - (Prims.of_int (1671)) - (Prims.of_int (73))) - (Obj.magic - (term_eq - t ta)) + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.CanonCommSemiring.fst" + (Prims.of_int (1634)) + (Prims.of_int (9)) + (Prims.of_int (1634)) + (Prims.of_int (21))) + (Prims.mk_range + "FStar.Tactics.CanonCommSemiring.fst" + (Prims.of_int (1634)) + (Prims.of_int (6)) + (Prims.of_int (1671)) + (Prims.of_int (73))) + (Obj.magic + (term_eq t + ta)) + (fun uu___3 + -> (fun - uu___3 -> - (fun uu___3 -> if uu___3 then Obj.magic - (Obj.repr (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.CanonCommSemiring.fst" @@ -2193,7 +2191,6 @@ let canon_semiring_aux : (e1::e2::[], vm) -> Obj.magic - (Obj.repr (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.CanonCommSemiring.fst" @@ -2432,25 +2429,22 @@ let canon_semiring_aux : uu___5))) uu___5))) uu___5))) - uu___5))) + uu___5)) | uu___5 -> Obj.magic - (Obj.repr (FStar_Tactics_Derived.fail - "Unexpected"))) - uu___4))) + "Unexpected")) + uu___4)) else Obj.magic - (Obj.repr (FStar_Tactics_Derived.fail - "Found equality, but terms do not have the expected type"))) - uu___3))) + "Found equality, but terms do not have the expected type")) + uu___3)) | uu___3 -> Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "Goal should be an equality"))) + (FStar_Tactics_Derived.fail + "Goal should be an equality")) uu___2))) uu___2))) uu___1)) let canon_semiring : 'a . 'a cr -> (unit, unit) FStar_Tactics_Effect.tac_repr diff --git a/ocaml/fstar-lib/generated/FStar_Tactics_CanonMonoid.ml b/ocaml/fstar-lib/generated/FStar_Tactics_CanonMonoid.ml index be918e0117b..648100aa857 100644 --- a/ocaml/fstar-lib/generated/FStar_Tactics_CanonMonoid.ml +++ b/ocaml/fstar-lib/generated/FStar_Tactics_CanonMonoid.ml @@ -448,89 +448,85 @@ let canon_monoid : me2) -> Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.CanonMonoid.fst" - (Prims.of_int (110)) - (Prims.of_int (9)) - (Prims.of_int (110)) - (Prims.of_int (32))) - (Prims.mk_range - "FStar.Tactics.CanonMonoid.fst" - (Prims.of_int (110)) - (Prims.of_int (6)) - (Prims.of_int (119)) - (Prims.of_int (69))) - (Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.CanonMonoid.fst" - (Prims.of_int (110)) - (Prims.of_int (23)) - (Prims.of_int (110)) - (Prims.of_int (32))) - (Prims.mk_range - "FStar.Tactics.CanonMonoid.fst" - (Prims.of_int (110)) - (Prims.of_int (9)) - (Prims.of_int (110)) - (Prims.of_int (32))) - (FStar_Tactics_Effect.lift_div_tac - (fun uu___2 -> - (fun uu___2 -> - Obj.magic - (failwith - "Cannot evaluate open quotation at runtime")) - uu___2)) + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.CanonMonoid.fst" + (Prims.of_int (110)) + (Prims.of_int (9)) + (Prims.of_int (110)) + (Prims.of_int (32))) + (Prims.mk_range + "FStar.Tactics.CanonMonoid.fst" + (Prims.of_int (110)) + (Prims.of_int (6)) + (Prims.of_int (119)) + (Prims.of_int (69))) + (Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.CanonMonoid.fst" + (Prims.of_int (110)) + (Prims.of_int (23)) + (Prims.of_int (110)) + (Prims.of_int (32))) + (Prims.mk_range + "FStar.Tactics.CanonMonoid.fst" + (Prims.of_int (110)) + (Prims.of_int (9)) + (Prims.of_int (110)) + (Prims.of_int (32))) + (FStar_Tactics_Effect.lift_div_tac (fun uu___2 -> (fun uu___2 -> Obj.magic - (FStar_Tactics_Builtins.term_eq_old - t uu___2)) - uu___2))) - (fun uu___2 -> + (failwith + "Cannot evaluate open quotation at runtime")) + uu___2)) (fun uu___2 -> - if uu___2 - then - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.CanonMonoid.fst" - (Prims.of_int (111)) - (Prims.of_int (17)) - (Prims.of_int (111)) - (Prims.of_int (34))) - (Prims.mk_range - "FStar.Tactics.CanonMonoid.fst" - (Prims.of_int (112)) - (Prims.of_int (8)) - (Prims.of_int (118)) - (Prims.of_int (51))) - (Obj.magic - (reification - m me1)) - (fun uu___3 -> - (fun r1 -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range + (fun uu___2 -> + Obj.magic + (FStar_Tactics_Builtins.term_eq_old + t uu___2)) uu___2))) + (fun uu___2 -> + (fun uu___2 -> + if uu___2 + then + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.CanonMonoid.fst" + (Prims.of_int (111)) + (Prims.of_int (17)) + (Prims.of_int (111)) + (Prims.of_int (34))) + (Prims.mk_range + "FStar.Tactics.CanonMonoid.fst" + (Prims.of_int (112)) + (Prims.of_int (8)) + (Prims.of_int (118)) + (Prims.of_int (51))) + (Obj.magic + (reification m me1)) + (fun uu___3 -> + (fun r1 -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "FStar.Tactics.CanonMonoid.fst" (Prims.of_int (112)) (Prims.of_int (17)) (Prims.of_int (112)) (Prims.of_int (34))) - (Prims.mk_range + (Prims.mk_range "FStar.Tactics.CanonMonoid.fst" (Prims.of_int (113)) (Prims.of_int (8)) (Prims.of_int (118)) (Prims.of_int (51))) - (Obj.magic + (Obj.magic (reification m me2)) - (fun + (fun uu___3 -> (fun r2 -> @@ -621,16 +617,14 @@ let canon_monoid : uu___4))) uu___3))) uu___3))) - uu___3))) - else - Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "Goal should be an equality at the right monoid type"))) - uu___2))) + uu___3)) + else + Obj.magic + (FStar_Tactics_Derived.fail + "Goal should be an equality at the right monoid type")) + uu___2)) | uu___2 -> Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "Goal should be an equality"))) + (FStar_Tactics_Derived.fail + "Goal should be an equality")) uu___1))) uu___1))) uu___) \ No newline at end of file diff --git a/ocaml/fstar-lib/generated/FStar_Tactics_Derived.ml b/ocaml/fstar-lib/generated/FStar_Tactics_Derived.ml index c8d79f468ed..4c8008fdce0 100644 --- a/ocaml/fstar-lib/generated/FStar_Tactics_Derived.ml +++ b/ocaml/fstar-lib/generated/FStar_Tactics_Derived.ml @@ -36,7 +36,7 @@ let (goals : (Prims.of_int (42)) (Prims.of_int (40)) (Prims.of_int (50))) (Prims.mk_range "FStar.Tactics.Derived.fst" (Prims.of_int (40)) (Prims.of_int (33)) (Prims.of_int (40)) (Prims.of_int (50))) - (FStar_Tactics_Effect.get ()) + (Obj.magic (FStar_Tactics_Effect.get ())) (fun uu___1 -> FStar_Tactics_Effect.lift_div_tac (fun uu___2 -> FStar_Tactics_Types.goals_of uu___1)) @@ -50,11 +50,11 @@ let (smt_goals : (Prims.of_int (50)) (Prims.of_int (41)) (Prims.of_int (58))) (Prims.mk_range "FStar.Tactics.Derived.fst" (Prims.of_int (41)) (Prims.of_int (37)) (Prims.of_int (41)) (Prims.of_int (58))) - (FStar_Tactics_Effect.get ()) + (Obj.magic (FStar_Tactics_Effect.get ())) (fun uu___1 -> FStar_Tactics_Effect.lift_div_tac (fun uu___2 -> FStar_Tactics_Types.smt_goals_of uu___1)) -let fail : 'a . Prims.string -> ('a, Obj.t) FStar_Tactics_Effect.tac_repr = +let fail : 'a . Prims.string -> ('a, unit) FStar_Tactics_Effect.tac_repr = fun m -> FStar_Tactics_Effect.raise (FStar_Tactics_Common.TacticFailure m) let fail_silently : 'a . Prims.string -> ('a, unit) FStar_Tactics_Effect.tac_repr = @@ -66,7 +66,10 @@ let fail_silently : (Prims.of_int (4)) (Prims.of_int (50)) (Prims.of_int (30))) (Obj.magic (FStar_Tactics_Builtins.set_urgency Prims.int_zero)) (fun uu___ -> - FStar_Tactics_Effect.raise (FStar_Tactics_Common.TacticFailure m)) + (fun uu___ -> + Obj.magic + (FStar_Tactics_Effect.raise + (FStar_Tactics_Common.TacticFailure m))) uu___) let (_cur_goal : unit -> (FStar_Tactics_Types.goal, unit) FStar_Tactics_Effect.tac_repr) = fun uu___ -> @@ -77,9 +80,14 @@ let (_cur_goal : (Prims.of_int (4)) (Prims.of_int (56)) (Prims.of_int (15))) (Obj.magic (goals ())) (fun uu___1 -> - match uu___1 with - | [] -> fail "no more goals" - | g::uu___2 -> FStar_Tactics_Effect.lift_div_tac (fun uu___3 -> g)) + (fun uu___1 -> + match uu___1 with + | [] -> Obj.magic (Obj.repr (fail "no more goals")) + | g::uu___2 -> + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac (fun uu___3 -> g)))) + uu___1) let (cur_env : unit -> (FStar_Reflection_Types.env, unit) FStar_Tactics_Effect.tac_repr) = fun uu___ -> @@ -131,7 +139,7 @@ let (cur_goal_safe : (Prims.of_int (18)) (Prims.of_int (72)) (Prims.of_int (26))) (Prims.mk_range "FStar.Tactics.Derived.fst" (Prims.of_int (72)) (Prims.of_int (9)) (Prims.of_int (72)) (Prims.of_int (26))) - (FStar_Tactics_Effect.get ()) + (Obj.magic (FStar_Tactics_Effect.get ())) (fun uu___1 -> FStar_Tactics_Effect.lift_div_tac (fun uu___2 -> FStar_Tactics_Types.goals_of uu___1)))) @@ -267,16 +275,14 @@ let (trivial : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = match uu___2 with | FStar_Reflection_Formula.True_ -> Obj.magic - (Obj.repr - (exact - (FStar_Reflection_Builtins.pack_ln - (FStar_Reflection_Data.Tv_Const - FStar_Reflection_Data.C_Unit)))) + (exact + (FStar_Reflection_Builtins.pack_ln + (FStar_Reflection_Data.Tv_Const + FStar_Reflection_Data.C_Unit))) | uu___3 -> Obj.magic - (Obj.repr - (FStar_Tactics_Effect.raise - Goal_not_trivial))) uu___2))) + (FStar_Tactics_Effect.raise + Goal_not_trivial)) uu___2))) uu___2))) uu___1) let (dismiss : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = fun uu___ -> @@ -289,9 +295,8 @@ let (dismiss : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = (fun uu___1 -> (fun uu___1 -> match uu___1 with - | [] -> Obj.magic (Obj.repr (fail "dismiss: no more goals")) - | uu___2::gs -> - Obj.magic (Obj.repr (FStar_Tactics_Builtins.set_goals gs))) + | [] -> Obj.magic (fail "dismiss: no more goals") + | uu___2::gs -> Obj.magic (FStar_Tactics_Builtins.set_goals gs)) uu___1) let (flip : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = fun uu___ -> @@ -315,17 +320,13 @@ let (flip : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = (fun uu___1 -> (fun uu___1 -> match uu___1 with - | [] -> - Obj.magic - (Obj.repr (fail "flip: less than two goals")) + | [] -> Obj.magic (fail "flip: less than two goals") | uu___2::[] -> - Obj.magic - (Obj.repr (fail "flip: less than two goals")) + Obj.magic (fail "flip: less than two goals") | g1::g2::gs1 -> Obj.magic - (Obj.repr - (FStar_Tactics_Builtins.set_goals (g2 :: g1 - :: gs1)))) uu___1))) uu___1) + (FStar_Tactics_Builtins.set_goals (g2 :: g1 :: + gs1))) uu___1))) uu___1) let (qed : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = fun uu___ -> FStar_Tactics_Effect.tac_bind @@ -335,9 +336,13 @@ let (qed : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = (Prims.of_int (4)) (Prims.of_int (130)) (Prims.of_int (32))) (Obj.magic (goals ())) (fun uu___1 -> - match uu___1 with - | [] -> FStar_Tactics_Effect.lift_div_tac (fun uu___2 -> ()) - | uu___2 -> fail "qed: not done!") + (fun uu___1 -> + match uu___1 with + | [] -> + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac (fun uu___2 -> ()))) + | uu___2 -> Obj.magic (Obj.repr (fail "qed: not done!"))) uu___1) let (debug : Prims.string -> (unit, unit) FStar_Tactics_Effect.tac_repr) = fun m -> FStar_Tactics_Effect.tac_bind @@ -386,24 +391,22 @@ let (smt : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = (fun uu___1 -> (fun uu___1 -> match uu___1 with - | ([], uu___2) -> - Obj.magic (Obj.repr (fail "smt: no active goals")) + | ([], uu___2) -> Obj.magic (fail "smt: no active goals") | (g::gs, gs') -> Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "FStar.Tactics.Derived.fst" - (Prims.of_int (147)) (Prims.of_int (8)) - (Prims.of_int (147)) (Prims.of_int (20))) - (Prims.mk_range "FStar.Tactics.Derived.fst" - (Prims.of_int (148)) (Prims.of_int (8)) - (Prims.of_int (148)) (Prims.of_int (32))) - (Obj.magic (FStar_Tactics_Builtins.set_goals gs)) + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "FStar.Tactics.Derived.fst" + (Prims.of_int (147)) (Prims.of_int (8)) + (Prims.of_int (147)) (Prims.of_int (20))) + (Prims.mk_range "FStar.Tactics.Derived.fst" + (Prims.of_int (148)) (Prims.of_int (8)) + (Prims.of_int (148)) (Prims.of_int (32))) + (Obj.magic (FStar_Tactics_Builtins.set_goals gs)) + (fun uu___2 -> (fun uu___2 -> - (fun uu___2 -> - Obj.magic - (FStar_Tactics_Builtins.set_smt_goals (g :: - gs'))) uu___2)))) uu___1) + Obj.magic + (FStar_Tactics_Builtins.set_smt_goals (g :: gs'))) + uu___2))) uu___1) let (idtac : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = fun uu___ -> (fun uu___ -> @@ -422,10 +425,9 @@ let (later : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = match uu___1 with | g::gs -> Obj.magic - (Obj.repr - (FStar_Tactics_Builtins.set_goals - (FStar_List_Tot_Base.op_At gs [g]))) - | uu___2 -> Obj.magic (Obj.repr (fail "later: no goals"))) uu___1) + (FStar_Tactics_Builtins.set_goals + (FStar_List_Tot_Base.op_At gs [g])) + | uu___2 -> Obj.magic (fail "later: no goals")) uu___1) let (apply : FStar_Reflection_Types.term -> (unit, unit) FStar_Tactics_Effect.tac_repr) = fun t -> FStar_Tactics_Builtins.t_apply true false false t @@ -534,21 +536,29 @@ let (topdown_rewrite : (Prims.of_int (254)) (Prims.of_int (10))) (match i with | uu___2 when uu___2 = Prims.int_zero -> - FStar_Tactics_Effect.lift_div_tac - (fun uu___3 -> - FStar_Tactics_Types.Continue) + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___3 -> + FStar_Tactics_Types.Continue))) | uu___2 when uu___2 = Prims.int_one -> - FStar_Tactics_Effect.lift_div_tac - (fun uu___3 -> - FStar_Tactics_Types.Skip) + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___3 -> + FStar_Tactics_Types.Skip))) | uu___2 when uu___2 = (Prims.of_int (2)) -> - FStar_Tactics_Effect.lift_div_tac - (fun uu___3 -> - FStar_Tactics_Types.Abort) + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___3 -> + FStar_Tactics_Types.Abort))) | uu___2 -> - fail - "topdown_rewrite: bad value from ctrl") + Obj.magic + (Obj.repr + (fail + "topdown_rewrite: bad value from ctrl"))) (fun f -> FStar_Tactics_Effect.lift_div_tac (fun uu___2 -> (b, f))))) uu___1))) @@ -671,8 +681,11 @@ let divide : (Prims.mk_range "FStar.Tactics.Derived.fst" (Prims.of_int (289)) (Prims.of_int (4)) (Prims.of_int (301)) (Prims.of_int (10))) (if n < Prims.int_zero - then fail "divide: negative n" - else FStar_Tactics_Effect.lift_div_tac (fun uu___1 -> ())) + then Obj.magic (Obj.repr (fail "divide: negative n")) + else + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac (fun uu___1 -> ())))) (fun uu___ -> (fun uu___ -> Obj.magic @@ -1093,91 +1106,84 @@ let focus : (fun uu___ -> (fun uu___ -> match uu___ with - | [] -> Obj.magic (Obj.repr (fail "focus: no goals")) + | [] -> Obj.magic (fail "focus: no goals") | g::gs -> Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "FStar.Tactics.Derived.fst" - (Prims.of_int (314)) (Prims.of_int (18)) - (Prims.of_int (314)) (Prims.of_int (30))) - (Prims.mk_range "FStar.Tactics.Derived.fst" - (Prims.of_int (315)) (Prims.of_int (8)) - (Prims.of_int (318)) (Prims.of_int (9))) - (Obj.magic (smt_goals ())) - (fun uu___1 -> - (fun sgs -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.Derived.fst" - (Prims.of_int (315)) (Prims.of_int (8)) - (Prims.of_int (315)) - (Prims.of_int (21))) - (Prims.mk_range - "FStar.Tactics.Derived.fst" - (Prims.of_int (315)) - (Prims.of_int (23)) - (Prims.of_int (318)) (Prims.of_int (9))) - (Obj.magic - (FStar_Tactics_Builtins.set_goals [g])) + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "FStar.Tactics.Derived.fst" + (Prims.of_int (314)) (Prims.of_int (18)) + (Prims.of_int (314)) (Prims.of_int (30))) + (Prims.mk_range "FStar.Tactics.Derived.fst" + (Prims.of_int (315)) (Prims.of_int (8)) + (Prims.of_int (318)) (Prims.of_int (9))) + (Obj.magic (smt_goals ())) + (fun uu___1 -> + (fun sgs -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "FStar.Tactics.Derived.fst" + (Prims.of_int (315)) (Prims.of_int (8)) + (Prims.of_int (315)) (Prims.of_int (21))) + (Prims.mk_range "FStar.Tactics.Derived.fst" + (Prims.of_int (315)) (Prims.of_int (23)) + (Prims.of_int (318)) (Prims.of_int (9))) + (Obj.magic + (FStar_Tactics_Builtins.set_goals [g])) + (fun uu___1 -> (fun uu___1 -> - (fun uu___1 -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.Derived.fst" - (Prims.of_int (315)) - (Prims.of_int (23)) - (Prims.of_int (315)) - (Prims.of_int (39))) - (Prims.mk_range - "FStar.Tactics.Derived.fst" - (Prims.of_int (316)) - (Prims.of_int (8)) - (Prims.of_int (318)) - (Prims.of_int (9))) - (Obj.magic - (FStar_Tactics_Builtins.set_smt_goals - [])) + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.Derived.fst" + (Prims.of_int (315)) + (Prims.of_int (23)) + (Prims.of_int (315)) + (Prims.of_int (39))) + (Prims.mk_range + "FStar.Tactics.Derived.fst" + (Prims.of_int (316)) + (Prims.of_int (8)) + (Prims.of_int (318)) + (Prims.of_int (9))) + (Obj.magic + (FStar_Tactics_Builtins.set_smt_goals + [])) + (fun uu___2 -> (fun uu___2 -> - (fun uu___2 -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.Derived.fst" - (Prims.of_int (316)) - (Prims.of_int (16)) - (Prims.of_int (316)) - (Prims.of_int (20))) - (Prims.mk_range - "FStar.Tactics.Derived.fst" - (Prims.of_int (317)) - (Prims.of_int (8)) - (Prims.of_int (318)) - (Prims.of_int (9))) - (Obj.magic (t ())) - (fun uu___3 -> - (fun x -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - ( - Prims.mk_range + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.Derived.fst" + (Prims.of_int (316)) + (Prims.of_int (16)) + (Prims.of_int (316)) + (Prims.of_int (20))) + (Prims.mk_range + "FStar.Tactics.Derived.fst" + (Prims.of_int (317)) + (Prims.of_int (8)) + (Prims.of_int (318)) + (Prims.of_int (9))) + (Obj.magic (t ())) + (fun uu___3 -> + (fun x -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "FStar.Tactics.Derived.fst" (Prims.of_int (317)) (Prims.of_int (8)) (Prims.of_int (317)) (Prims.of_int (33))) - ( - Prims.mk_range + (Prims.mk_range "FStar.Tactics.Derived.fst" (Prims.of_int (317)) (Prims.of_int (35)) (Prims.of_int (318)) (Prims.of_int (9))) + (Obj.magic ( - Obj.magic - (FStar_Tactics_Effect.tac_bind + FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.Derived.fst" (Prims.of_int (317)) @@ -1221,9 +1227,8 @@ let focus : (FStar_Tactics_Builtins.set_goals uu___3)) uu___3))) - ( - fun - uu___3 -> + (fun uu___3 + -> (fun uu___3 -> Obj.magic @@ -1294,9 +1299,8 @@ let focus : uu___5 -> x)))) uu___3))) - uu___3))) - uu___2))) uu___1))) uu___1)))) - uu___) + uu___3))) uu___2))) + uu___1))) uu___1))) uu___) let (dump1 : Prims.string -> (unit, unit) FStar_Tactics_Effect.tac_repr) = fun m -> focus (fun uu___ -> FStar_Tactics_Builtins.dump m) let rec mapAll : @@ -1782,8 +1786,10 @@ let (guard : Prims.bool -> (unit, unit) FStar_Tactics_Effect.tac_repr) = fun uu___ -> (fun b -> if Prims.op_Negation b - then Obj.magic (fail "guard failed") - else Obj.magic (FStar_Tactics_Effect.lift_div_tac (fun uu___1 -> ()))) + then Obj.magic (Obj.repr (fail "guard failed")) + else + Obj.magic + (Obj.repr (FStar_Tactics_Effect.lift_div_tac (fun uu___1 -> ())))) uu___ let try_with : 'a . @@ -1857,8 +1863,7 @@ let first : = fun ts -> FStar_List_Tot_Base.fold_right op_Less_Bar_Greater ts - (fun uu___ -> (fun uu___ -> Obj.magic (fail "no tactics to try")) uu___) - () + (fun uu___ -> fail "no tactics to try") () let rec repeat : 'a . (unit -> ('a, unit) FStar_Tactics_Effect.tac_repr) -> @@ -2136,9 +2141,8 @@ let (skip_guard : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = (Obj.magic (is_guard ())) (fun uu___1 -> (fun uu___1 -> - if uu___1 - then Obj.magic (Obj.repr (smt ())) - else Obj.magic (Obj.repr (fail ""))) uu___1) + if uu___1 then Obj.magic (smt ()) else Obj.magic (fail "")) + uu___1) let (guards_to_smt : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = fun uu___ -> FStar_Tactics_Effect.tac_bind @@ -2383,57 +2387,50 @@ let rec (__assumption_aux : FStar_Reflection_Types.binders -> (unit, unit) FStar_Tactics_Effect.tac_repr) = - fun uu___ -> - (fun bs -> - match bs with - | [] -> Obj.magic (Obj.repr (fail "no assumption matches goal")) - | b::bs1 -> - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "FStar.Tactics.Derived.fst" - (Prims.of_int (539)) (Prims.of_int (16)) - (Prims.of_int (539)) (Prims.of_int (32))) - (Prims.mk_range "FStar.Tactics.Derived.fst" - (Prims.of_int (540)) (Prims.of_int (8)) - (Prims.of_int (543)) (Prims.of_int (27))) - (Obj.magic (binder_to_term b)) - (fun uu___ -> - (fun t -> - Obj.magic - (try_with - (fun uu___ -> match () with | () -> exact t) - (fun uu___ -> - try_with - (fun uu___1 -> - match () with - | () -> - FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.Derived.fst" - (Prims.of_int (541)) - (Prims.of_int (13)) - (Prims.of_int (541)) - (Prims.of_int (48))) - (Prims.mk_range - "FStar.Tactics.Derived.fst" - (Prims.of_int (542)) - (Prims.of_int (13)) - (Prims.of_int (542)) - (Prims.of_int (20))) - (Obj.magic - (apply - (FStar_Reflection_Builtins.pack_ln - (FStar_Reflection_Data.Tv_FVar - (FStar_Reflection_Builtins.pack_fv - ["FStar"; - "Squash"; - "return_squash"]))))) - (fun uu___2 -> - (fun uu___2 -> - Obj.magic (exact t)) uu___2)) - (fun uu___1 -> __assumption_aux bs1)))) - uu___)))) uu___ + fun bs -> + match bs with + | [] -> fail "no assumption matches goal" + | b::bs1 -> + FStar_Tactics_Effect.tac_bind + (Prims.mk_range "FStar.Tactics.Derived.fst" (Prims.of_int (539)) + (Prims.of_int (16)) (Prims.of_int (539)) (Prims.of_int (32))) + (Prims.mk_range "FStar.Tactics.Derived.fst" (Prims.of_int (540)) + (Prims.of_int (8)) (Prims.of_int (543)) (Prims.of_int (27))) + (Obj.magic (binder_to_term b)) + (fun uu___ -> + (fun t -> + Obj.magic + (try_with (fun uu___ -> match () with | () -> exact t) + (fun uu___ -> + try_with + (fun uu___1 -> + match () with + | () -> + FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.Derived.fst" + (Prims.of_int (541)) + (Prims.of_int (13)) + (Prims.of_int (541)) + (Prims.of_int (48))) + (Prims.mk_range + "FStar.Tactics.Derived.fst" + (Prims.of_int (542)) + (Prims.of_int (13)) + (Prims.of_int (542)) + (Prims.of_int (20))) + (Obj.magic + (apply + (FStar_Reflection_Builtins.pack_ln + (FStar_Reflection_Data.Tv_FVar + (FStar_Reflection_Builtins.pack_fv + ["FStar"; + "Squash"; + "return_squash"]))))) + (fun uu___2 -> + (fun uu___2 -> Obj.magic (exact t)) + uu___2)) + (fun uu___1 -> __assumption_aux bs1)))) uu___) let (assumption : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = fun uu___ -> FStar_Tactics_Effect.tac_bind @@ -2525,8 +2522,7 @@ let (rewrite' : (fun uu___2 -> Obj.magic (FStar_Tactics_Builtins.rewrite b)) uu___2))) uu___1))) - (fun uu___ -> (fun uu___ -> Obj.magic (fail "rewrite' failed")) uu___) - () + (fun uu___ -> fail "rewrite' failed") () let rec (try_rewrite_equality : FStar_Reflection_Types.term -> FStar_Reflection_Types.binders -> @@ -2642,26 +2638,23 @@ let (unfold_def : match uu___ with | FStar_Reflection_Data.Tv_FVar fv -> Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "FStar.Tactics.Derived.fst" - (Prims.of_int (601)) (Prims.of_int (16)) - (Prims.of_int (601)) (Prims.of_int (42))) - (Prims.mk_range "FStar.Tactics.Derived.fst" - (Prims.of_int (602)) (Prims.of_int (8)) - (Prims.of_int (602)) (Prims.of_int (30))) - (FStar_Tactics_Effect.lift_div_tac - (fun uu___1 -> - FStar_Reflection_Builtins.implode_qn - (FStar_Reflection_Builtins.inspect_fv fv))) + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "FStar.Tactics.Derived.fst" + (Prims.of_int (601)) (Prims.of_int (16)) + (Prims.of_int (601)) (Prims.of_int (42))) + (Prims.mk_range "FStar.Tactics.Derived.fst" + (Prims.of_int (602)) (Prims.of_int (8)) + (Prims.of_int (602)) (Prims.of_int (30))) + (FStar_Tactics_Effect.lift_div_tac (fun uu___1 -> - (fun n -> - Obj.magic - (FStar_Tactics_Builtins.norm - [FStar_Pervasives.delta_fully [n]])) - uu___1))) - | uu___1 -> - Obj.magic (Obj.repr (fail "unfold_def: term is not a fv"))) + FStar_Reflection_Builtins.implode_qn + (FStar_Reflection_Builtins.inspect_fv fv))) + (fun uu___1 -> + (fun n -> + Obj.magic + (FStar_Tactics_Builtins.norm + [FStar_Pervasives.delta_fully [n]])) uu___1)) + | uu___1 -> Obj.magic (fail "unfold_def: term is not a fv")) uu___) let (l_to_r : FStar_Reflection_Types.term Prims.list -> @@ -2892,83 +2885,79 @@ let (grewrite_eq : | FStar_Reflection_Formula.Comp (FStar_Reflection_Formula.Eq uu___3, l, r) -> Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.Derived.fst" - (Prims.of_int (655)) - (Prims.of_int (6)) - (Prims.of_int (655)) - (Prims.of_int (18))) - (Prims.mk_range - "FStar.Tactics.Derived.fst" - (Prims.of_int (656)) - (Prims.of_int (6)) - (Prims.of_int (657)) - (Prims.of_int (56))) - (Obj.magic (grewrite l r)) + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.Derived.fst" + (Prims.of_int (655)) + (Prims.of_int (6)) + (Prims.of_int (655)) + (Prims.of_int (18))) + (Prims.mk_range + "FStar.Tactics.Derived.fst" + (Prims.of_int (656)) + (Prims.of_int (6)) + (Prims.of_int (657)) + (Prims.of_int (56))) + (Obj.magic (grewrite l r)) + (fun uu___4 -> (fun uu___4 -> - (fun uu___4 -> - Obj.magic - (iseq - [idtac; - (fun uu___5 -> - FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.Derived.fst" - (Prims.of_int (656)) - (Prims.of_int (30)) - (Prims.of_int (656)) - (Prims.of_int (55))) - (Prims.mk_range - "FStar.Tactics.Derived.fst" - (Prims.of_int (657)) - (Prims.of_int (30)) - (Prims.of_int (657)) - (Prims.of_int (54))) - (Obj.magic - (apply_lemma - (FStar_Reflection_Builtins.pack_ln - (FStar_Reflection_Data.Tv_FVar - (FStar_Reflection_Builtins.pack_fv - ["FStar"; - "Tactics"; - "Derived"; - "__un_sq_eq"]))))) + Obj.magic + (iseq + [idtac; + (fun uu___5 -> + FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.Derived.fst" + (Prims.of_int (656)) + (Prims.of_int (30)) + (Prims.of_int (656)) + (Prims.of_int (55))) + (Prims.mk_range + "FStar.Tactics.Derived.fst" + (Prims.of_int (657)) + (Prims.of_int (30)) + (Prims.of_int (657)) + (Prims.of_int (54))) + (Obj.magic + (apply_lemma + (FStar_Reflection_Builtins.pack_ln + (FStar_Reflection_Data.Tv_FVar + (FStar_Reflection_Builtins.pack_fv + ["FStar"; + "Tactics"; + "Derived"; + "__un_sq_eq"]))))) + (fun uu___6 -> (fun uu___6 -> - (fun uu___6 -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.Derived.fst" - (Prims.of_int (657)) - (Prims.of_int (36)) - (Prims.of_int (657)) - (Prims.of_int (54))) - (Prims.mk_range - "FStar.Tactics.Derived.fst" - (Prims.of_int (657)) - (Prims.of_int (30)) - (Prims.of_int (657)) - (Prims.of_int (54))) - (Obj.magic - (binder_to_term + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.Derived.fst" + (Prims.of_int (657)) + (Prims.of_int (36)) + (Prims.of_int (657)) + (Prims.of_int (54))) + (Prims.mk_range + "FStar.Tactics.Derived.fst" + (Prims.of_int (657)) + (Prims.of_int (30)) + (Prims.of_int (657)) + (Prims.of_int (54))) + (Obj.magic + (binder_to_term b)) + (fun uu___7 -> (fun uu___7 -> - (fun - uu___7 -> Obj.magic (exact uu___7)) uu___7))) - uu___6))])) - uu___4))) + uu___6))])) uu___4)) | uu___3 -> Obj.magic - (Obj.repr - (fail - "grewrite_eq: binder type is not an equality"))) + (fail + "grewrite_eq: binder type is not an equality")) uu___2))) uu___) let rec (apply_squash_or_lem : Prims.nat -> @@ -3000,121 +2989,101 @@ let rec (apply_squash_or_lem : (fun uu___1 -> try_with (fun uu___2 -> match () with | () -> apply_lemma t) (fun uu___2 -> - (fun uu___2 -> - if d <= Prims.int_zero - then - Obj.magic (Obj.repr (fail "mapply: out of fuel")) - else - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "FStar.Tactics.Derived.fst" - (Prims.of_int (688)) (Prims.of_int (13)) - (Prims.of_int (688)) (Prims.of_int (30))) - (Prims.mk_range "FStar.Tactics.Derived.fst" - (Prims.of_int (689)) (Prims.of_int (4)) - (Prims.of_int (737)) (Prims.of_int (41))) - (Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.Derived.fst" - (Prims.of_int (688)) - (Prims.of_int (16)) - (Prims.of_int (688)) - (Prims.of_int (28))) - (Prims.mk_range - "FStar.Tactics.Derived.fst" - (Prims.of_int (688)) - (Prims.of_int (13)) - (Prims.of_int (688)) - (Prims.of_int (30))) - (Obj.magic (cur_env ())) - (fun uu___4 -> - (fun uu___4 -> - Obj.magic - (FStar_Tactics_Builtins.tc - uu___4 t)) uu___4))) + if d <= Prims.int_zero + then fail "mapply: out of fuel" + else + FStar_Tactics_Effect.tac_bind + (Prims.mk_range "FStar.Tactics.Derived.fst" + (Prims.of_int (688)) (Prims.of_int (13)) + (Prims.of_int (688)) (Prims.of_int (30))) + (Prims.mk_range "FStar.Tactics.Derived.fst" + (Prims.of_int (689)) (Prims.of_int (4)) + (Prims.of_int (737)) (Prims.of_int (41))) + (Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "FStar.Tactics.Derived.fst" + (Prims.of_int (688)) (Prims.of_int (16)) + (Prims.of_int (688)) (Prims.of_int (28))) + (Prims.mk_range "FStar.Tactics.Derived.fst" + (Prims.of_int (688)) (Prims.of_int (13)) + (Prims.of_int (688)) (Prims.of_int (30))) + (Obj.magic (cur_env ())) + (fun uu___4 -> (fun uu___4 -> - (fun ty -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.Derived.fst" - (Prims.of_int (689)) - (Prims.of_int (17)) - (Prims.of_int (689)) - (Prims.of_int (31))) - (Prims.mk_range - "FStar.Tactics.Derived.fst" - (Prims.of_int (689)) - (Prims.of_int (4)) - (Prims.of_int (737)) - (Prims.of_int (41))) - (Obj.magic - (FStar_Tactics_SyntaxHelpers.collect_arr - ty)) - (fun uu___4 -> - (fun uu___4 -> - match uu___4 with - | (tys, c) -> - (match FStar_Reflection_Builtins.inspect_comp - c - with - | FStar_Reflection_Data.C_Lemma - (pre, post, - uu___5) - -> - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - ( - Prims.mk_range - "FStar.Tactics.Derived.fst" - (Prims.of_int (693)) - (Prims.of_int (18)) - (Prims.of_int (693)) - (Prims.of_int (32))) - ( - Prims.mk_range - "FStar.Tactics.Derived.fst" - (Prims.of_int (694)) - (Prims.of_int (7)) - (Prims.of_int (702)) - (Prims.of_int (41))) - ( - FStar_Tactics_Effect.lift_div_tac - (fun - uu___6 -> - FStar_Reflection_Builtins.pack_ln - (FStar_Reflection_Data.Tv_App - (post, + Obj.magic + (FStar_Tactics_Builtins.tc uu___4 t)) + uu___4))) + (fun uu___4 -> + (fun ty -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.Derived.fst" + (Prims.of_int (689)) + (Prims.of_int (17)) + (Prims.of_int (689)) + (Prims.of_int (31))) + (Prims.mk_range + "FStar.Tactics.Derived.fst" + (Prims.of_int (689)) + (Prims.of_int (4)) + (Prims.of_int (737)) + (Prims.of_int (41))) + (Obj.magic + (FStar_Tactics_SyntaxHelpers.collect_arr + ty)) + (fun uu___4 -> + (fun uu___4 -> + match uu___4 with + | (tys, c) -> + (match FStar_Reflection_Builtins.inspect_comp + c + with + | FStar_Reflection_Data.C_Lemma + (pre, post, uu___5) -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.Derived.fst" + (Prims.of_int (693)) + (Prims.of_int (18)) + (Prims.of_int (693)) + (Prims.of_int (32))) + (Prims.mk_range + "FStar.Tactics.Derived.fst" + (Prims.of_int (694)) + (Prims.of_int (7)) + (Prims.of_int (702)) + (Prims.of_int (41))) + (FStar_Tactics_Effect.lift_div_tac + (fun uu___6 -> + FStar_Reflection_Builtins.pack_ln + (FStar_Reflection_Data.Tv_App + (post, ((FStar_Reflection_Builtins.pack_ln (FStar_Reflection_Data.Tv_Const FStar_Reflection_Data.C_Unit)), FStar_Reflection_Data.Q_Explicit))))) - ( - fun - uu___6 -> - (fun - post1 -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range + (fun uu___6 -> + (fun post1 -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "FStar.Tactics.Derived.fst" (Prims.of_int (694)) (Prims.of_int (18)) (Prims.of_int (694)) (Prims.of_int (35))) - (Prims.mk_range + (Prims.mk_range "FStar.Tactics.Derived.fst" (Prims.of_int (696)) (Prims.of_int (7)) (Prims.of_int (702)) (Prims.of_int (41))) - (Obj.magic + (Obj.magic (norm_term [] post1)) - (fun + (fun uu___6 -> (fun post2 -> @@ -3145,7 +3114,6 @@ let rec (apply_squash_or_lem : FStar_Reflection_Formula.Implies (p, q) -> Obj.magic - (Obj.repr (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.Derived.fst" @@ -3177,47 +3145,42 @@ let rec (apply_squash_or_lem : (d - Prims.int_one) t)) - uu___7))) + uu___7)) | uu___7 -> Obj.magic - (Obj.repr (fail - "mapply: can't apply (1)"))) - uu___6))) + "mapply: can't apply (1)")) uu___6))) uu___6))) - | FStar_Reflection_Data.C_Total - rt -> - Obj.magic - (Obj.repr - (match - FStar_Reflection_Derived.unsquash_term - rt - with - | FStar_Pervasives_Native.Some - rt1 -> - FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.Derived.fst" - (Prims.of_int (710)) - (Prims.of_int (18)) - (Prims.of_int (710)) - (Prims.of_int (33))) - (Prims.mk_range - "FStar.Tactics.Derived.fst" - (Prims.of_int (712)) - (Prims.of_int (9)) - (Prims.of_int (718)) - (Prims.of_int (43))) - (Obj.magic - (norm_term - [] rt1)) - (fun - uu___5 -> - (fun rt2 - -> - Obj.magic + uu___6)) + | FStar_Reflection_Data.C_Total + rt -> + (match FStar_Reflection_Derived.unsquash_term + rt + with + | FStar_Pervasives_Native.Some + rt1 -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.Derived.fst" + (Prims.of_int (710)) + (Prims.of_int (18)) + (Prims.of_int (710)) + (Prims.of_int (33))) + (Prims.mk_range + "FStar.Tactics.Derived.fst" + (Prims.of_int (712)) + (Prims.of_int (9)) + (Prims.of_int (718)) + (Prims.of_int (43))) + (Obj.magic + (norm_term [] + rt1)) + (fun uu___5 -> + (fun rt2 -> + Obj.magic (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.Derived.fst" @@ -3244,7 +3207,6 @@ let rec (apply_squash_or_lem : FStar_Reflection_Formula.Implies (p, q) -> Obj.magic - (Obj.repr (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.Derived.fst" @@ -3276,38 +3238,36 @@ let rec (apply_squash_or_lem : (d - Prims.int_one) t)) - uu___6))) + uu___6)) | uu___6 -> Obj.magic - (Obj.repr (fail - "mapply: can't apply (1)"))) + "mapply: can't apply (1)")) uu___5))) - uu___5) - | FStar_Pervasives_Native.None - -> - FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.Derived.fst" - (Prims.of_int (725)) - (Prims.of_int (18)) - (Prims.of_int (725)) - (Prims.of_int (33))) - (Prims.mk_range - "FStar.Tactics.Derived.fst" - (Prims.of_int (727)) - (Prims.of_int (9)) - (Prims.of_int (734)) - (Prims.of_int (20))) - (Obj.magic - (norm_term - [] rt)) - (fun - uu___5 -> - (fun rt1 - -> - Obj.magic + uu___5)) + | FStar_Pervasives_Native.None + -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.Derived.fst" + (Prims.of_int (725)) + (Prims.of_int (18)) + (Prims.of_int (725)) + (Prims.of_int (33))) + (Prims.mk_range + "FStar.Tactics.Derived.fst" + (Prims.of_int (727)) + (Prims.of_int (9)) + (Prims.of_int (734)) + (Prims.of_int (20))) + (Obj.magic + (norm_term [] + rt)) + (fun uu___5 -> + (fun rt1 -> + Obj.magic (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.Derived.fst" @@ -3398,13 +3358,12 @@ let rec (apply_squash_or_lem : (apply t)) uu___7))) uu___5))) - uu___5))) - | uu___5 -> - Obj.magic - (Obj.repr - (fail - "mapply: can't apply (2)")))) - uu___4))) uu___4)))) uu___2))) + uu___5))) + | uu___5 -> + Obj.magic + (fail + "mapply: can't apply (2)"))) + uu___4))) uu___4)))) let (mapply : FStar_Reflection_Types.term -> (unit, unit) FStar_Tactics_Effect.tac_repr) = fun t -> apply_squash_or_lem (Prims.of_int (10)) t @@ -3518,10 +3477,7 @@ let finish_by : (Prims.of_int (773)) (Prims.of_int (9))) (Obj.magic (or_else qed - (fun uu___ -> - (fun uu___ -> - Obj.magic (fail "finish_by: not finished")) - uu___))) + (fun uu___ -> fail "finish_by: not finished"))) (fun uu___ -> FStar_Tactics_Effect.lift_div_tac (fun uu___1 -> x)))) uu___) @@ -3676,12 +3632,11 @@ let (tlabel : Prims.string -> (unit, unit) FStar_Tactics_Effect.tac_repr) = (fun uu___ -> (fun uu___ -> match uu___ with - | [] -> Obj.magic (Obj.repr (fail "tlabel: no goals")) + | [] -> Obj.magic (fail "tlabel: no goals") | h::t -> Obj.magic - (Obj.repr - (FStar_Tactics_Builtins.set_goals - ((FStar_Tactics_Types.set_label l h) :: t)))) uu___) + (FStar_Tactics_Builtins.set_goals + ((FStar_Tactics_Types.set_label l h) :: t))) uu___) let (tlabel' : Prims.string -> (unit, unit) FStar_Tactics_Effect.tac_repr) = fun l -> FStar_Tactics_Effect.tac_bind @@ -3693,27 +3648,26 @@ let (tlabel' : Prims.string -> (unit, unit) FStar_Tactics_Effect.tac_repr) = (fun uu___ -> (fun uu___ -> match uu___ with - | [] -> Obj.magic (Obj.repr (fail "tlabel': no goals")) + | [] -> Obj.magic (fail "tlabel': no goals") | h::t -> Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "FStar.Tactics.Derived.fst" - (Prims.of_int (821)) (Prims.of_int (16)) - (Prims.of_int (821)) (Prims.of_int (45))) - (Prims.mk_range "FStar.Tactics.Derived.fst" - (Prims.of_int (822)) (Prims.of_int (8)) - (Prims.of_int (822)) (Prims.of_int (26))) - (FStar_Tactics_Effect.lift_div_tac - (fun uu___1 -> - FStar_Tactics_Types.set_label - (Prims.strcat l - (FStar_Tactics_Types.get_label h)) h)) + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "FStar.Tactics.Derived.fst" + (Prims.of_int (821)) (Prims.of_int (16)) + (Prims.of_int (821)) (Prims.of_int (45))) + (Prims.mk_range "FStar.Tactics.Derived.fst" + (Prims.of_int (822)) (Prims.of_int (8)) + (Prims.of_int (822)) (Prims.of_int (26))) + (FStar_Tactics_Effect.lift_div_tac (fun uu___1 -> - (fun h1 -> - Obj.magic - (FStar_Tactics_Builtins.set_goals (h1 :: t))) - uu___1)))) uu___) + FStar_Tactics_Types.set_label + (Prims.strcat l + (FStar_Tactics_Types.get_label h)) h)) + (fun uu___1 -> + (fun h1 -> + Obj.magic + (FStar_Tactics_Builtins.set_goals (h1 :: t))) + uu___1))) uu___) let (focus_all : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = fun uu___ -> FStar_Tactics_Effect.tac_bind @@ -3795,11 +3749,9 @@ let (bump_nth : Prims.pos -> (unit, unit) FStar_Tactics_Effect.tac_repr) = (fun uu___ -> match uu___ with | FStar_Pervasives_Native.None -> - Obj.magic (Obj.repr (fail "bump_nth: not that many goals")) + Obj.magic (fail "bump_nth: not that many goals") | FStar_Pervasives_Native.Some (h, t) -> - Obj.magic - (Obj.repr (FStar_Tactics_Builtins.set_goals (h :: t)))) - uu___) + Obj.magic (FStar_Tactics_Builtins.set_goals (h :: t))) uu___) let rec (destruct_list : FStar_Reflection_Types.term -> (FStar_Reflection_Types.term Prims.list, unit) @@ -3829,22 +3781,20 @@ let rec (destruct_list : (FStar_Reflection_Builtins.inspect_fv fv) = FStar_Reflection_Const.cons_qn then - Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "FStar.Tactics.Derived.fst" - (Prims.of_int (851)) (Prims.of_int (17)) - (Prims.of_int (851)) (Prims.of_int (33))) - (Prims.mk_range "FStar.Tactics.Derived.fst" - (Prims.of_int (851)) (Prims.of_int (14)) - (Prims.of_int (851)) (Prims.of_int (16))) - (Obj.magic (destruct_list a2)) - (fun uu___1 -> - FStar_Tactics_Effect.lift_div_tac - (fun uu___2 -> a1 :: uu___1))) + FStar_Tactics_Effect.tac_bind + (Prims.mk_range "FStar.Tactics.Derived.fst" + (Prims.of_int (851)) (Prims.of_int (17)) + (Prims.of_int (851)) (Prims.of_int (33))) + (Prims.mk_range "FStar.Tactics.Derived.fst" + (Prims.of_int (851)) (Prims.of_int (14)) + (Prims.of_int (851)) (Prims.of_int (16))) + (Obj.magic (destruct_list a2)) + (fun uu___1 -> + FStar_Tactics_Effect.lift_div_tac + (fun uu___2 -> a1 :: uu___1)) else - Obj.repr - (FStar_Tactics_Effect.raise - FStar_Tactics_Common.NotAListLiteral))) + FStar_Tactics_Effect.raise + FStar_Tactics_Common.NotAListLiteral)) | (FStar_Reflection_Data.Tv_FVar fv, (uu___1, FStar_Reflection_Data.Q_Implicit)::(a1, FStar_Reflection_Data.Q_Explicit):: @@ -3855,22 +3805,20 @@ let rec (destruct_list : (FStar_Reflection_Builtins.inspect_fv fv) = FStar_Reflection_Const.cons_qn then - Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "FStar.Tactics.Derived.fst" - (Prims.of_int (851)) (Prims.of_int (17)) - (Prims.of_int (851)) (Prims.of_int (33))) - (Prims.mk_range "FStar.Tactics.Derived.fst" - (Prims.of_int (851)) (Prims.of_int (14)) - (Prims.of_int (851)) (Prims.of_int (16))) - (Obj.magic (destruct_list a2)) - (fun uu___2 -> - FStar_Tactics_Effect.lift_div_tac - (fun uu___3 -> a1 :: uu___2))) + FStar_Tactics_Effect.tac_bind + (Prims.mk_range "FStar.Tactics.Derived.fst" + (Prims.of_int (851)) (Prims.of_int (17)) + (Prims.of_int (851)) (Prims.of_int (33))) + (Prims.mk_range "FStar.Tactics.Derived.fst" + (Prims.of_int (851)) (Prims.of_int (14)) + (Prims.of_int (851)) (Prims.of_int (16))) + (Obj.magic (destruct_list a2)) + (fun uu___2 -> + FStar_Tactics_Effect.lift_div_tac + (fun uu___3 -> a1 :: uu___2)) else - Obj.repr - (FStar_Tactics_Effect.raise - FStar_Tactics_Common.NotAListLiteral))) + FStar_Tactics_Effect.raise + FStar_Tactics_Common.NotAListLiteral)) | (FStar_Reflection_Data.Tv_FVar fv, uu___1) -> Obj.magic (Obj.repr @@ -3878,11 +3826,13 @@ let rec (destruct_list : (FStar_Reflection_Builtins.inspect_fv fv) = FStar_Reflection_Const.nil_qn then - FStar_Tactics_Effect.lift_div_tac - (fun uu___2 -> []) + Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___2 -> [])) else - FStar_Tactics_Effect.raise - FStar_Tactics_Common.NotAListLiteral)) + Obj.repr + (FStar_Tactics_Effect.raise + FStar_Tactics_Common.NotAListLiteral))) | uu___1 -> Obj.magic (Obj.repr @@ -3910,25 +3860,30 @@ let (get_match_body : (fun uu___1 -> (fun uu___1 -> match uu___1 with - | FStar_Pervasives_Native.None -> Obj.magic (Obj.repr (fail "")) + | FStar_Pervasives_Native.None -> Obj.magic (fail "") | FStar_Pervasives_Native.Some t -> Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "FStar.Tactics.Derived.fst" - (Prims.of_int (863)) (Prims.of_int (20)) - (Prims.of_int (863)) (Prims.of_int (39))) - (Prims.mk_range "FStar.Tactics.Derived.fst" - (Prims.of_int (863)) (Prims.of_int (14)) - (Prims.of_int (865)) (Prims.of_int (46))) - (Obj.magic (inspect_unascribe t)) + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "FStar.Tactics.Derived.fst" + (Prims.of_int (863)) (Prims.of_int (20)) + (Prims.of_int (863)) (Prims.of_int (39))) + (Prims.mk_range "FStar.Tactics.Derived.fst" + (Prims.of_int (863)) (Prims.of_int (14)) + (Prims.of_int (865)) (Prims.of_int (46))) + (Obj.magic (inspect_unascribe t)) + (fun uu___2 -> (fun uu___2 -> match uu___2 with | FStar_Reflection_Data.Tv_Match (sc, uu___3, uu___4) -> - FStar_Tactics_Effect.lift_div_tac - (fun uu___5 -> sc) - | uu___3 -> fail "Goal is not a match")))) uu___1) + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___5 -> sc))) + | uu___3 -> + Obj.magic + (Obj.repr (fail "Goal is not a match"))) + uu___2))) uu___1) let rec last : 'a . 'a Prims.list -> ('a, unit) FStar_Tactics_Effect.tac_repr = fun uu___ -> @@ -4061,17 +4016,26 @@ let (nth_binder : (Prims.of_int (896)) (Prims.of_int (2)) (Prims.of_int (898)) (Prims.of_int (15))) (if k < Prims.int_zero - then fail "not enough binders" + then + Obj.magic + (Obj.repr (fail "not enough binders")) else - FStar_Tactics_Effect.lift_div_tac - (fun uu___1 -> k)) - (fun k1 -> - match FStar_List_Tot_Base.nth bs k1 with - | FStar_Pervasives_Native.None -> - fail "not enough binders" - | FStar_Pervasives_Native.Some b -> - FStar_Tactics_Effect.lift_div_tac - (fun uu___ -> b)))) uu___))) uu___) + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___1 -> k)))) + (fun uu___ -> + (fun k1 -> + match FStar_List_Tot_Base.nth bs k1 with + | FStar_Pervasives_Native.None -> + Obj.magic + (Obj.repr (fail "not enough binders")) + | FStar_Pervasives_Native.Some b -> + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___ -> b)))) uu___))) + uu___))) uu___) exception Appears let (uu___is_Appears : Prims.exn -> Prims.bool) = fun projectee -> match projectee with | Appears -> true | uu___ -> false @@ -4113,10 +4077,15 @@ let (name_appears_in : (if (FStar_Reflection_Builtins.inspect_fv fv) = nm - then FStar_Tactics_Effect.raise Appears + then + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.raise Appears)) else - FStar_Tactics_Effect.lift_div_tac - (fun uu___3 -> ())) + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___3 -> ())))) (fun uu___2 -> FStar_Tactics_Effect.lift_div_tac (fun uu___3 -> t1))) @@ -4157,9 +4126,12 @@ let (name_appears_in : match uu___ with | Appears -> Obj.magic - (FStar_Tactics_Effect.lift_div_tac - (fun uu___1 -> true)) - | e -> Obj.magic (FStar_Tactics_Effect.raise e)) + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___1 -> true))) + | e -> + Obj.magic + (Obj.repr (FStar_Tactics_Effect.raise e))) uu___))) uu___) let rec (mk_abs : FStar_Reflection_Types.binder Prims.list -> diff --git a/ocaml/fstar-lib/generated/FStar_Tactics_Effect.ml b/ocaml/fstar-lib/generated/FStar_Tactics_Effect.ml index c20d9427b13..226dda9374b 100644 --- a/ocaml/fstar-lib/generated/FStar_Tactics_Effect.ml +++ b/ocaml/fstar-lib/generated/FStar_Tactics_Effect.ml @@ -49,9 +49,9 @@ let __proj__TAC__item__bind = tac_bind type ('a, 'wp, 'uuuuu, 'uuuuu1) lift_div_tac_wp = 'wp let lift_div_tac : 'a 'wp . (unit -> 'a) -> ('a, 'wp) tac_repr = fun f -> fun ps -> FStar_Tactics_Result.Success ((f ()), ps) -let (get : unit -> (FStar_Tactics_Types.proofstate, Obj.t) tac_repr) = +let (get : unit -> (FStar_Tactics_Types.proofstate, unit) tac_repr) = fun uu___ -> fun ps -> FStar_Tactics_Result.Success (ps, ps) -let raise : 'a . Prims.exn -> ('a, Obj.t) tac_repr = +let raise : 'a . Prims.exn -> ('a, unit) tac_repr = fun e -> fun ps -> FStar_Tactics_Result.Failed (e, ps) type ('uuuuu, 'p) with_tactic = 'p let (rewrite_with_tactic : diff --git a/ocaml/fstar-lib/generated/FStar_Tactics_Logic.ml b/ocaml/fstar-lib/generated/FStar_Tactics_Logic.ml index 453aaf00628..9848dd07536 100644 --- a/ocaml/fstar-lib/generated/FStar_Tactics_Logic.ml +++ b/ocaml/fstar-lib/generated/FStar_Tactics_Logic.ml @@ -105,10 +105,7 @@ let (split : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = (FStar_Reflection_Data.Tv_FVar (FStar_Reflection_Builtins.pack_fv ["FStar"; "Tactics"; "Logic"; "split_lem"])))) - (fun uu___1 -> - (fun uu___1 -> - Obj.magic (FStar_Tactics_Derived.fail "Could not split goal")) - uu___1) + (fun uu___1 -> FStar_Tactics_Derived.fail "Could not split goal") let (implies_intro : unit -> (FStar_Reflection_Types.binder, unit) FStar_Tactics_Effect.tac_repr) = @@ -222,9 +219,12 @@ let (pose_lemma : (Prims.of_int (126)) (Prims.of_int (5))) (match FStar_Reflection_Builtins.inspect_comp c with | FStar_Reflection_Data.C_Lemma (pre, post, uu___) -> - FStar_Tactics_Effect.lift_div_tac - (fun uu___1 -> (pre, post)) - | uu___ -> FStar_Tactics_Derived.fail "") + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___1 -> (pre, post)))) + | uu___ -> + Obj.magic (Obj.repr (FStar_Tactics_Derived.fail ""))) (fun uu___ -> (fun uu___ -> match uu___ with @@ -729,80 +729,74 @@ let rec (simplify_eq_implication : match r with | FStar_Pervasives_Native.None -> Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "Not an equality implication")) + (FStar_Tactics_Derived.fail + "Not an equality implication") | FStar_Pervasives_Native.Some (uu___1, rhs) -> Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.Logic.fst" - (Prims.of_int (160)) - (Prims.of_int (19)) - (Prims.of_int (160)) - (Prims.of_int (35))) - (Prims.mk_range - "FStar.Tactics.Logic.fst" - (Prims.of_int (161)) - (Prims.of_int (8)) - (Prims.of_int (163)) - (Prims.of_int (37))) - (Obj.magic (implies_intro ())) - (fun uu___2 -> - (fun eq_h -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.Logic.fst" - (Prims.of_int (161)) - (Prims.of_int (8)) - (Prims.of_int (161)) - (Prims.of_int (20))) - (Prims.mk_range - "FStar.Tactics.Logic.fst" - (Prims.of_int (162)) - (Prims.of_int (8)) - (Prims.of_int (163)) - (Prims.of_int (37))) - (Obj.magic - (FStar_Tactics_Builtins.rewrite - eq_h)) + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.Logic.fst" + (Prims.of_int (160)) + (Prims.of_int (19)) + (Prims.of_int (160)) + (Prims.of_int (35))) + (Prims.mk_range + "FStar.Tactics.Logic.fst" + (Prims.of_int (161)) + (Prims.of_int (8)) + (Prims.of_int (163)) + (Prims.of_int (37))) + (Obj.magic (implies_intro ())) + (fun uu___2 -> + (fun eq_h -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.Logic.fst" + (Prims.of_int (161)) + (Prims.of_int (8)) + (Prims.of_int (161)) + (Prims.of_int (20))) + (Prims.mk_range + "FStar.Tactics.Logic.fst" + (Prims.of_int (162)) + (Prims.of_int (8)) + (Prims.of_int (163)) + (Prims.of_int (37))) + (Obj.magic + (FStar_Tactics_Builtins.rewrite + eq_h)) + (fun uu___2 -> (fun uu___2 -> - (fun uu___2 -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - ( - Prims.mk_range + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "FStar.Tactics.Logic.fst" (Prims.of_int (162)) (Prims.of_int (8)) (Prims.of_int (162)) (Prims.of_int (20))) - ( - Prims.mk_range + (Prims.mk_range "FStar.Tactics.Logic.fst" (Prims.of_int (163)) (Prims.of_int (8)) (Prims.of_int (163)) (Prims.of_int (37))) + (Obj.magic ( - Obj.magic - (FStar_Tactics_Builtins.clear_top + FStar_Tactics_Builtins.clear_top ())) - ( - fun - uu___3 -> + (fun uu___3 + -> (fun uu___3 -> Obj.magic (visit simplify_eq_implication)) uu___3))) - uu___2))) - uu___2)))) uu___1))) - uu___1))) uu___1) + uu___2))) uu___2))) + uu___1))) uu___1))) uu___1) let (rewrite_all_equalities : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = fun uu___ -> visit simplify_eq_implication @@ -862,52 +856,49 @@ let rec (unfold_definition_and_simplify_eq : match r with | FStar_Pervasives_Native.None -> Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "Not an equality implication")) + (FStar_Tactics_Derived.fail + "Not an equality implication") | FStar_Pervasives_Native.Some (uu___2, rhs) -> Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.Logic.fst" - (Prims.of_int (180)) - (Prims.of_int (23)) - (Prims.of_int (180)) - (Prims.of_int (39))) - (Prims.mk_range - "FStar.Tactics.Logic.fst" - (Prims.of_int (181)) - (Prims.of_int (12)) - (Prims.of_int (183)) - (Prims.of_int (66))) - (Obj.magic - (implies_intro ())) - (fun uu___3 -> - (fun eq_h -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.Logic.fst" - (Prims.of_int (181)) - (Prims.of_int (12)) - (Prims.of_int (181)) - (Prims.of_int (24))) - (Prims.mk_range - "FStar.Tactics.Logic.fst" - (Prims.of_int (182)) - (Prims.of_int (12)) - (Prims.of_int (183)) - (Prims.of_int (66))) - (Obj.magic - (FStar_Tactics_Builtins.rewrite - eq_h)) + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.Logic.fst" + (Prims.of_int (180)) + (Prims.of_int (23)) + (Prims.of_int (180)) + (Prims.of_int (39))) + (Prims.mk_range + "FStar.Tactics.Logic.fst" + (Prims.of_int (181)) + (Prims.of_int (12)) + (Prims.of_int (183)) + (Prims.of_int (66))) + (Obj.magic + (implies_intro ())) + (fun uu___3 -> + (fun eq_h -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.Logic.fst" + (Prims.of_int (181)) + (Prims.of_int (12)) + (Prims.of_int (181)) + (Prims.of_int (24))) + (Prims.mk_range + "FStar.Tactics.Logic.fst" + (Prims.of_int (182)) + (Prims.of_int (12)) + (Prims.of_int (183)) + (Prims.of_int (66))) + (Obj.magic + (FStar_Tactics_Builtins.rewrite + eq_h)) + (fun uu___3 -> (fun uu___3 -> - (fun - uu___3 -> - Obj.magic + Obj.magic (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.Logic.fst" @@ -935,8 +926,8 @@ let rec (unfold_definition_and_simplify_eq : unfold_definition_and_simplify_eq tm))) uu___4))) - uu___3))) - uu___3)))) uu___2)))) + uu___3))) + uu___3))) uu___2)))) uu___))) uu___) let (unsquash : FStar_Reflection_Types.term -> @@ -1247,10 +1238,7 @@ let (instantiate : (fa, FStar_Reflection_Data.Q_Explicit)))), (x, FStar_Reflection_Data.Q_Explicit))))) (fun uu___1 -> - (fun uu___1 -> - Obj.magic - (FStar_Tactics_Derived.fail "could not instantiate")) - uu___1)) + FStar_Tactics_Derived.fail "could not instantiate")) let (instantiate_as : FStar_Reflection_Types.term -> FStar_Reflection_Types.term -> @@ -1373,12 +1361,19 @@ let rec (sk_binder' : (fun uu___4 -> uu___3 <> Prims.int_one)))) (fun uu___3 -> - if uu___3 - then - FStar_Tactics_Derived.fail "no" - else - FStar_Tactics_Effect.lift_div_tac - (fun uu___5 -> ())))) + (fun uu___3 -> + if uu___3 + then + Obj.magic + (Obj.repr + (FStar_Tactics_Derived.fail + "no")) + else + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___5 -> ())))) + uu___3))) (fun uu___3 -> (fun uu___3 -> Obj.magic @@ -1573,7 +1568,5 @@ let (using_lemma : "lem3_fa"]))), (t, FStar_Reflection_Data.Q_Explicit))))) (fun uu___2 -> - (fun uu___2 -> - Obj.magic - (FStar_Tactics_Derived.fail - "using_lemma: failed to instantiate")) uu___2))) \ No newline at end of file + FStar_Tactics_Derived.fail + "using_lemma: failed to instantiate"))) \ No newline at end of file diff --git a/ocaml/fstar-lib/generated/FStar_Tactics_PatternMatching.ml b/ocaml/fstar-lib/generated/FStar_Tactics_PatternMatching.ml index 976c932a28e..de9f29a87f5 100644 --- a/ocaml/fstar-lib/generated/FStar_Tactics_PatternMatching.ml +++ b/ocaml/fstar-lib/generated/FStar_Tactics_PatternMatching.ml @@ -28,79 +28,73 @@ let (fetch_eq_side : | FStar_Reflection_Data.Tv_App (squash, (g1, uu___2)) -> Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.PatternMatching.fst" - (Prims.of_int (66)) (Prims.of_int (11)) - (Prims.of_int (66)) (Prims.of_int (25))) - (Prims.mk_range - "FStar.Tactics.PatternMatching.fst" - (Prims.of_int (66)) (Prims.of_int (4)) - (Prims.of_int (87)) (Prims.of_int (51))) - (Obj.magic - (FStar_Tactics_Builtins.inspect squash)) + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.PatternMatching.fst" + (Prims.of_int (66)) (Prims.of_int (11)) + (Prims.of_int (66)) (Prims.of_int (25))) + (Prims.mk_range + "FStar.Tactics.PatternMatching.fst" + (Prims.of_int (66)) (Prims.of_int (4)) + (Prims.of_int (87)) (Prims.of_int (51))) + (Obj.magic + (FStar_Tactics_Builtins.inspect squash)) + (fun uu___3 -> (fun uu___3 -> - (fun uu___3 -> - match uu___3 with - | FStar_Reflection_Data.Tv_UInst - (squash1, uu___4) -> - Obj.magic - (Obj.repr - (if - (FStar_Reflection_Derived.fv_to_string - squash1) - = - (FStar_Reflection_Derived.flatten_name - FStar_Reflection_Const.squash_qn) - then - Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.PatternMatching.fst" - (Prims.of_int (70)) - (Prims.of_int (16)) - (Prims.of_int (70)) - (Prims.of_int (25))) - (Prims.mk_range - "FStar.Tactics.PatternMatching.fst" - (Prims.of_int (70)) - (Prims.of_int (9)) - (Prims.of_int (85)) - (Prims.of_int (48))) - (Obj.magic - (FStar_Tactics_Builtins.inspect - g1)) - (fun uu___5 -> - (fun uu___5 -> - match uu___5 - with - | FStar_Reflection_Data.Tv_App - (eq_type_x, - (y, - uu___6)) - -> - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.PatternMatching.fst" - (Prims.of_int (72)) - (Prims.of_int (19)) - (Prims.of_int (72)) - (Prims.of_int (36))) - (Prims.mk_range - "FStar.Tactics.PatternMatching.fst" - (Prims.of_int (72)) - (Prims.of_int (12)) - (Prims.of_int (84)) - (Prims.of_int (39))) - (Obj.magic - (FStar_Tactics_Builtins.inspect + match uu___3 with + | FStar_Reflection_Data.Tv_UInst + (squash1, uu___4) -> + if + (FStar_Reflection_Derived.fv_to_string + squash1) + = + (FStar_Reflection_Derived.flatten_name + FStar_Reflection_Const.squash_qn) + then + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.PatternMatching.fst" + (Prims.of_int (70)) + (Prims.of_int (16)) + (Prims.of_int (70)) + (Prims.of_int (25))) + (Prims.mk_range + "FStar.Tactics.PatternMatching.fst" + (Prims.of_int (70)) + (Prims.of_int (9)) + (Prims.of_int (85)) + (Prims.of_int (48))) + (Obj.magic + (FStar_Tactics_Builtins.inspect + g1)) + (fun uu___5 -> + (fun uu___5 -> + match uu___5 with + | FStar_Reflection_Data.Tv_App + (eq_type_x, + (y, uu___6)) + -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.PatternMatching.fst" + (Prims.of_int (72)) + (Prims.of_int (19)) + (Prims.of_int (72)) + (Prims.of_int (36))) + (Prims.mk_range + "FStar.Tactics.PatternMatching.fst" + (Prims.of_int (72)) + (Prims.of_int (12)) + (Prims.of_int (84)) + (Prims.of_int (39))) + (Obj.magic + (FStar_Tactics_Builtins.inspect eq_type_x)) - (fun - uu___7 -> - (fun + (fun uu___7 + -> + (fun uu___7 -> match uu___7 with @@ -111,7 +105,6 @@ let (fetch_eq_side : uu___8)) -> Obj.magic - (Obj.repr (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.PatternMatching.fst" @@ -141,7 +134,6 @@ let (fetch_eq_side : uu___10)) -> Obj.magic - (Obj.repr (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.PatternMatching.fst" @@ -161,6 +153,9 @@ let (fetch_eq_side : (fun uu___11 -> + (fun + uu___11 + -> match uu___11 with | @@ -168,123 +163,126 @@ let (fetch_eq_side : (eq1, uu___12) -> - if + Obj.magic + (Obj.repr + (if (FStar_Reflection_Derived.fv_to_string eq1) = (FStar_Reflection_Derived.flatten_name FStar_Reflection_Const.eq2_qn) then - FStar_Tactics_Effect.lift_div_tac + Obj.repr + (FStar_Tactics_Effect.lift_div_tac (fun uu___13 - -> (x, y)) + -> (x, y))) else - FStar_Tactics_Derived.fail - "not an equality" + Obj.repr + (FStar_Tactics_Derived.fail + "not an equality"))) | FStar_Reflection_Data.Tv_FVar eq1 -> - if + Obj.magic + (Obj.repr + (if (FStar_Reflection_Derived.fv_to_string eq1) = (FStar_Reflection_Derived.flatten_name FStar_Reflection_Const.eq2_qn) then - FStar_Tactics_Effect.lift_div_tac + Obj.repr + (FStar_Tactics_Effect.lift_div_tac (fun uu___12 - -> (x, y)) + -> (x, y))) else - FStar_Tactics_Derived.fail - "not an equality" + Obj.repr + (FStar_Tactics_Derived.fail + "not an equality"))) | uu___12 -> - FStar_Tactics_Derived.fail + Obj.magic + (Obj.repr + (FStar_Tactics_Derived.fail "not an app2 of fvar: "))) + uu___11)) | uu___10 -> Obj.magic - (Obj.repr (FStar_Tactics_Derived.fail - "not an app3"))) - uu___9))) + "not an app3")) + uu___9)) | uu___8 -> Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "not an app2"))) - uu___7))) - | uu___6 -> - Obj.magic - (Obj.repr (FStar_Tactics_Derived.fail - "not an app under squash"))) - uu___5)) - else - Obj.repr - (FStar_Tactics_Derived.fail - "not a squash"))) - | FStar_Reflection_Data.Tv_FVar - squash1 -> - Obj.magic - (Obj.repr - (if - (FStar_Reflection_Derived.fv_to_string - squash1) - = - (FStar_Reflection_Derived.flatten_name - FStar_Reflection_Const.squash_qn) - then - Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.PatternMatching.fst" - (Prims.of_int (70)) - (Prims.of_int (16)) - (Prims.of_int (70)) - (Prims.of_int (25))) - (Prims.mk_range - "FStar.Tactics.PatternMatching.fst" - (Prims.of_int (70)) - (Prims.of_int (9)) - (Prims.of_int (85)) - (Prims.of_int (48))) - (Obj.magic - (FStar_Tactics_Builtins.inspect - g1)) - (fun uu___4 -> - (fun uu___4 -> - match uu___4 - with - | FStar_Reflection_Data.Tv_App - (eq_type_x, - (y, - uu___5)) - -> - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.PatternMatching.fst" - (Prims.of_int (72)) - (Prims.of_int (19)) - (Prims.of_int (72)) - (Prims.of_int (36))) - (Prims.mk_range - "FStar.Tactics.PatternMatching.fst" - (Prims.of_int (72)) - (Prims.of_int (12)) - (Prims.of_int (84)) - (Prims.of_int (39))) - (Obj.magic - (FStar_Tactics_Builtins.inspect + "not an app2")) + uu___7)) + | uu___6 -> + Obj.magic + (FStar_Tactics_Derived.fail + "not an app under squash")) + uu___5)) + else + Obj.magic + (FStar_Tactics_Derived.fail + "not a squash") + | FStar_Reflection_Data.Tv_FVar squash1 + -> + if + (FStar_Reflection_Derived.fv_to_string + squash1) + = + (FStar_Reflection_Derived.flatten_name + FStar_Reflection_Const.squash_qn) + then + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.PatternMatching.fst" + (Prims.of_int (70)) + (Prims.of_int (16)) + (Prims.of_int (70)) + (Prims.of_int (25))) + (Prims.mk_range + "FStar.Tactics.PatternMatching.fst" + (Prims.of_int (70)) + (Prims.of_int (9)) + (Prims.of_int (85)) + (Prims.of_int (48))) + (Obj.magic + (FStar_Tactics_Builtins.inspect + g1)) + (fun uu___4 -> + (fun uu___4 -> + match uu___4 with + | FStar_Reflection_Data.Tv_App + (eq_type_x, + (y, uu___5)) + -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.PatternMatching.fst" + (Prims.of_int (72)) + (Prims.of_int (19)) + (Prims.of_int (72)) + (Prims.of_int (36))) + (Prims.mk_range + "FStar.Tactics.PatternMatching.fst" + (Prims.of_int (72)) + (Prims.of_int (12)) + (Prims.of_int (84)) + (Prims.of_int (39))) + (Obj.magic + (FStar_Tactics_Builtins.inspect eq_type_x)) - (fun - uu___6 -> - (fun + (fun uu___6 + -> + (fun uu___6 -> match uu___6 with @@ -295,7 +293,6 @@ let (fetch_eq_side : uu___7)) -> Obj.magic - (Obj.repr (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.PatternMatching.fst" @@ -325,7 +322,6 @@ let (fetch_eq_side : uu___9)) -> Obj.magic - (Obj.repr (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.PatternMatching.fst" @@ -345,6 +341,9 @@ let (fetch_eq_side : (fun uu___10 -> + (fun + uu___10 + -> match uu___10 with | @@ -352,76 +351,81 @@ let (fetch_eq_side : (eq1, uu___11) -> - if + Obj.magic + (Obj.repr + (if (FStar_Reflection_Derived.fv_to_string eq1) = (FStar_Reflection_Derived.flatten_name FStar_Reflection_Const.eq2_qn) then - FStar_Tactics_Effect.lift_div_tac + Obj.repr + (FStar_Tactics_Effect.lift_div_tac (fun uu___12 - -> (x, y)) + -> (x, y))) else - FStar_Tactics_Derived.fail - "not an equality" + Obj.repr + (FStar_Tactics_Derived.fail + "not an equality"))) | FStar_Reflection_Data.Tv_FVar eq1 -> - if + Obj.magic + (Obj.repr + (if (FStar_Reflection_Derived.fv_to_string eq1) = (FStar_Reflection_Derived.flatten_name FStar_Reflection_Const.eq2_qn) then - FStar_Tactics_Effect.lift_div_tac + Obj.repr + (FStar_Tactics_Effect.lift_div_tac (fun uu___11 - -> (x, y)) + -> (x, y))) else - FStar_Tactics_Derived.fail - "not an equality" + Obj.repr + (FStar_Tactics_Derived.fail + "not an equality"))) | uu___11 -> - FStar_Tactics_Derived.fail + Obj.magic + (Obj.repr + (FStar_Tactics_Derived.fail "not an app2 of fvar: "))) + uu___10)) | uu___9 -> Obj.magic - (Obj.repr (FStar_Tactics_Derived.fail - "not an app3"))) - uu___8))) + "not an app3")) + uu___8)) | uu___7 -> Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "not an app2"))) - uu___6))) - | uu___5 -> - Obj.magic - (Obj.repr (FStar_Tactics_Derived.fail - "not an app under squash"))) - uu___4)) - else - Obj.repr - (FStar_Tactics_Derived.fail - "not a squash"))) - | uu___4 -> - Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "not an app of fvar at top level"))) - uu___3))) + "not an app2")) + uu___6)) + | uu___5 -> + Obj.magic + (FStar_Tactics_Derived.fail + "not an app under squash")) + uu___4)) + else + Obj.magic + (FStar_Tactics_Derived.fail + "not a squash") + | uu___4 -> + Obj.magic + (FStar_Tactics_Derived.fail + "not an app of fvar at top level")) + uu___3)) | uu___2 -> Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "not an app at top level"))) uu___1))) - uu___1) + (FStar_Tactics_Derived.fail + "not an app at top level")) uu___1))) uu___1) let mustfail : 'a . (unit -> ('a, unit) FStar_Tactics_Effect.tac_repr) -> @@ -437,11 +441,15 @@ let mustfail : (Prims.of_int (130)) (Prims.of_int (4)) (Prims.of_int (132)) (Prims.of_int (16))) (Obj.magic (FStar_Tactics_Derived.trytac t)) (fun uu___ -> - match uu___ with - | FStar_Pervasives_Native.Some uu___1 -> - FStar_Tactics_Derived.fail message - | FStar_Pervasives_Native.None -> - FStar_Tactics_Effect.lift_div_tac (fun uu___1 -> ())) + (fun uu___ -> + match uu___ with + | FStar_Pervasives_Native.Some uu___1 -> + Obj.magic (Obj.repr (FStar_Tactics_Derived.fail message)) + | FStar_Pervasives_Native.None -> + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac (fun uu___1 -> ())))) + uu___) let (implies_intro' : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = fun uu___ -> FStar_Tactics_Effect.tac_bind @@ -1037,8 +1045,10 @@ let lift_exn_tac : (Prims.of_int (268)) (Prims.of_int (18)) (Prims.of_int (268)) (Prims.of_int (61))) (Obj.magic (string_of_match_exception ex)) - (fun uu___ -> FStar_Tactics_Derived.fail uu___)))) - uu___1 uu___ + (fun uu___ -> + (fun uu___ -> + Obj.magic (FStar_Tactics_Derived.fail uu___)) + uu___)))) uu___1 uu___ let lift_exn_tactic : 'a 'b . ('a -> 'b match_res) -> 'a -> ('b, unit) FStar_Tactics_Effect.tac_repr @@ -1063,8 +1073,10 @@ let lift_exn_tactic : (Prims.of_int (273)) (Prims.of_int (18)) (Prims.of_int (273)) (Prims.of_int (61))) (Obj.magic (string_of_match_exception ex)) - (fun uu___ -> FStar_Tactics_Derived.fail uu___)))) - uu___1 uu___ + (fun uu___ -> + (fun uu___ -> + Obj.magic (FStar_Tactics_Derived.fail uu___)) + uu___)))) uu___1 uu___ type bindings = (varname * FStar_Reflection_Types.term) Prims.list let (string_of_bindings : bindings -> (Prims.string, unit) FStar_Tactics_Effect.tac_repr) = @@ -1484,8 +1496,10 @@ let (match_term : (Prims.of_int (338)) (Prims.of_int (20)) (Prims.of_int (338)) (Prims.of_int (63))) (Obj.magic (string_of_match_exception ex)) - (fun uu___1 -> FStar_Tactics_Derived.fail uu___1)))) - uu___) + (fun uu___1 -> + (fun uu___1 -> + Obj.magic (FStar_Tactics_Derived.fail uu___1)) + uu___1)))) uu___) let debug : 'uuuuu . 'uuuuu -> (unit, unit) FStar_Tactics_Effect.tac_repr = fun uu___ -> (fun msg -> @@ -1649,9 +1663,13 @@ let assoc_varname_fail : match FStar_List_Tot_Base.assoc key ls with | FStar_Pervasives_Native.None -> Obj.magic - (FStar_Tactics_Derived.fail (Prims.strcat "Not found: " key)) + (Obj.repr + (FStar_Tactics_Derived.fail + (Prims.strcat "Not found: " key))) | FStar_Pervasives_Native.Some x -> - Obj.magic (FStar_Tactics_Effect.lift_div_tac (fun uu___ -> x))) + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac (fun uu___ -> x)))) uu___1 uu___ let ms_locate_hyp : 'a . @@ -1694,123 +1712,98 @@ let rec solve_mp_for_single_hyp : (matching_solution -> ('a, unit) FStar_Tactics_Effect.tac_repr) -> matching_solution -> ('a, unit) FStar_Tactics_Effect.tac_repr = - fun uu___4 -> - fun uu___3 -> - fun uu___2 -> - fun uu___1 -> - fun uu___ -> - (fun name -> - fun pat -> - fun hypotheses -> - fun body -> - fun part_sol -> - match hypotheses with - | [] -> - Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "No matching hypothesis")) - | h::hs -> - Obj.magic - (Obj.repr - (FStar_Tactics_Derived.or_else - (fun uu___ -> - FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.PatternMatching.fst" - (Prims.of_int (448)) - (Prims.of_int (15)) - (Prims.of_int (448)) - (Prims.of_int (73))) - (Prims.mk_range - "FStar.Tactics.PatternMatching.fst" - (Prims.of_int (448)) - (Prims.of_int (9)) - (Prims.of_int (453)) - (Prims.of_int (73))) - (Obj.magic - (interp_pattern_aux pat - part_sol.ms_vars - (FStar_Reflection_Derived.type_of_binder - h))) - (fun uu___1 -> - (fun uu___1 -> - match uu___1 with - | Failure ex -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.PatternMatching.fst" - (Prims.of_int (450)) - (Prims.of_int (16)) - (Prims.of_int (450)) - (Prims.of_int (74))) - (Prims.mk_range - "FStar.Tactics.PatternMatching.fst" - (Prims.of_int (450)) - (Prims.of_int (11)) - (Prims.of_int (450)) - (Prims.of_int (74))) - (Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.PatternMatching.fst" - (Prims.of_int (450)) - (Prims.of_int (43)) - (Prims.of_int (450)) - (Prims.of_int (73))) - (Prims.mk_range - "prims.fst" - (Prims.of_int (606)) - (Prims.of_int (19)) - (Prims.of_int (606)) - (Prims.of_int (31))) - (Obj.magic - (string_of_match_exception - ex)) - (fun uu___2 -> - FStar_Tactics_Effect.lift_div_tac - (fun uu___3 - -> - Prims.strcat - "Failed to match hyp: " - uu___2)))) - (fun uu___2 -> - FStar_Tactics_Derived.fail - uu___2)) - | Success bindings1 -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.PatternMatching.fst" - (Prims.of_int (452)) - (Prims.of_int (35)) - (Prims.of_int (452)) - (Prims.of_int (37))) - (Prims.mk_range - "FStar.Tactics.PatternMatching.fst" - (Prims.of_int (453)) - (Prims.of_int (11)) - (Prims.of_int (453)) - (Prims.of_int (73))) - (FStar_Tactics_Effect.lift_div_tac - (fun uu___2 -> - (name, h) :: - (part_sol.ms_hyps))) - (fun uu___2 -> - (fun ms_hyps -> - Obj.magic - (body - { - ms_vars = - bindings1; - ms_hyps - })) uu___2))) - uu___1)) - (fun uu___ -> - solve_mp_for_single_hyp name pat hs - body part_sol)))) uu___4 uu___3 - uu___2 uu___1 uu___ + fun name -> + fun pat -> + fun hypotheses -> + fun body -> + fun part_sol -> + match hypotheses with + | [] -> FStar_Tactics_Derived.fail "No matching hypothesis" + | h::hs -> + FStar_Tactics_Derived.or_else + (fun uu___ -> + FStar_Tactics_Effect.tac_bind + (Prims.mk_range "FStar.Tactics.PatternMatching.fst" + (Prims.of_int (448)) (Prims.of_int (15)) + (Prims.of_int (448)) (Prims.of_int (73))) + (Prims.mk_range "FStar.Tactics.PatternMatching.fst" + (Prims.of_int (448)) (Prims.of_int (9)) + (Prims.of_int (453)) (Prims.of_int (73))) + (Obj.magic + (interp_pattern_aux pat part_sol.ms_vars + (FStar_Reflection_Derived.type_of_binder h))) + (fun uu___1 -> + (fun uu___1 -> + match uu___1 with + | Failure ex -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.PatternMatching.fst" + (Prims.of_int (450)) + (Prims.of_int (16)) + (Prims.of_int (450)) + (Prims.of_int (74))) + (Prims.mk_range + "FStar.Tactics.PatternMatching.fst" + (Prims.of_int (450)) + (Prims.of_int (11)) + (Prims.of_int (450)) + (Prims.of_int (74))) + (Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.PatternMatching.fst" + (Prims.of_int (450)) + (Prims.of_int (43)) + (Prims.of_int (450)) + (Prims.of_int (73))) + (Prims.mk_range "prims.fst" + (Prims.of_int (606)) + (Prims.of_int (19)) + (Prims.of_int (606)) + (Prims.of_int (31))) + (Obj.magic + (string_of_match_exception ex)) + (fun uu___2 -> + FStar_Tactics_Effect.lift_div_tac + (fun uu___3 -> + Prims.strcat + "Failed to match hyp: " + uu___2)))) + (fun uu___2 -> + (fun uu___2 -> + Obj.magic + (FStar_Tactics_Derived.fail + uu___2)) uu___2)) + | Success bindings1 -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.PatternMatching.fst" + (Prims.of_int (452)) + (Prims.of_int (35)) + (Prims.of_int (452)) + (Prims.of_int (37))) + (Prims.mk_range + "FStar.Tactics.PatternMatching.fst" + (Prims.of_int (453)) + (Prims.of_int (11)) + (Prims.of_int (453)) + (Prims.of_int (73))) + (FStar_Tactics_Effect.lift_div_tac + (fun uu___2 -> (name, h) :: + (part_sol.ms_hyps))) + (fun uu___2 -> + (fun ms_hyps -> + Obj.magic + (body + { + ms_vars = bindings1; + ms_hyps + })) uu___2))) uu___1)) + (fun uu___ -> + solve_mp_for_single_hyp name pat hs body part_sol) let rec solve_mp_for_hyps : 'a . (varname * pattern) Prims.list -> @@ -1905,8 +1898,10 @@ let solve_mp : "Failed to match goal: " uu___1)))) (fun uu___1 -> - FStar_Tactics_Derived.fail - uu___1))) + (fun uu___1 -> + Obj.magic + (FStar_Tactics_Derived.fail + uu___1)) uu___1))) | Success bindings1 -> Obj.magic (Obj.repr @@ -2047,8 +2042,10 @@ let (pattern_of_term : (Prims.of_int (532)) (Prims.of_int (20)) (Prims.of_int (532)) (Prims.of_int (63))) (Obj.magic (string_of_match_exception ex)) - (fun uu___1 -> FStar_Tactics_Derived.fail uu___1)))) - uu___) + (fun uu___1 -> + (fun uu___1 -> + Obj.magic (FStar_Tactics_Derived.fail uu___1)) + uu___1)))) uu___) type 'a hyp = FStar_Reflection_Types.binder type 'a pm_goal = unit let (hyp_qn : Prims.string) = "FStar.Tactics.PatternMatching.hyp" @@ -2208,6 +2205,8 @@ let (classify_abspat_binder : ( fun uu___2 -> + (fun + uu___2 -> match uu___2 with | @@ -2215,25 +2214,32 @@ let (classify_abspat_binder : ((uu___3, goal_typ)::[]) -> - FStar_Tactics_Effect.lift_div_tac + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac (fun uu___4 -> (ABKGoal, - goal_typ)) + goal_typ)))) | Success uu___3 -> - FStar_Tactics_Derived.fail - "classifiy_abspat_binder: impossible (2)" + Obj.magic + (Obj.repr + (FStar_Tactics_Derived.fail + "classifiy_abspat_binder: impossible (2)")) | Failure uu___3 -> - FStar_Tactics_Effect.lift_div_tac + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac (fun uu___4 -> ((ABKVar typ), - typ)))))) + typ))))) + uu___2)))) uu___))) uu___))) uu___))) uu___))) uu___) let rec (binders_and_body_of_abs : diff --git a/ocaml/fstar-lib/generated/FStar_Tactics_Simplifier.ml b/ocaml/fstar-lib/generated/FStar_Tactics_Simplifier.ml index e6d04f4ca30..d3708421bb0 100644 --- a/ocaml/fstar-lib/generated/FStar_Tactics_Simplifier.ml +++ b/ocaml/fstar-lib/generated/FStar_Tactics_Simplifier.ml @@ -144,55 +144,45 @@ let (inhabit : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = match FStar_Reflection_Builtins.inspect_ln t with | FStar_Reflection_Data.Tv_FVar fv -> Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "FStar.Tactics.Simplifier.fst" - (Prims.of_int (200)) (Prims.of_int (17)) - (Prims.of_int (200)) (Prims.of_int (30))) - (Prims.mk_range "FStar.Tactics.Simplifier.fst" - (Prims.of_int (201)) (Prims.of_int (13)) - (Prims.of_int (204)) (Prims.of_int (20))) - (FStar_Tactics_Effect.lift_div_tac - (fun uu___1 -> - FStar_Reflection_Builtins.inspect_fv fv)) + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "FStar.Tactics.Simplifier.fst" + (Prims.of_int (200)) (Prims.of_int (17)) + (Prims.of_int (200)) (Prims.of_int (30))) + (Prims.mk_range "FStar.Tactics.Simplifier.fst" + (Prims.of_int (201)) (Prims.of_int (13)) + (Prims.of_int (204)) (Prims.of_int (20))) + (FStar_Tactics_Effect.lift_div_tac (fun uu___1 -> - (fun qn -> - if qn = FStar_Reflection_Const.int_lid - then - Obj.magic - (Obj.repr - (FStar_Tactics_Derived.exact - (FStar_Reflection_Builtins.pack_ln - (FStar_Reflection_Data.Tv_Const - (FStar_Reflection_Data.C_Int - (Prims.of_int (42))))))) - else - Obj.magic - (Obj.repr - (if qn = FStar_Reflection_Const.bool_lid - then - Obj.repr - (FStar_Tactics_Derived.exact - (FStar_Reflection_Builtins.pack_ln - (FStar_Reflection_Data.Tv_Const - FStar_Reflection_Data.C_True))) - else - Obj.repr - (if - qn = - FStar_Reflection_Const.unit_lid - then - Obj.repr - (FStar_Tactics_Derived.exact - (FStar_Reflection_Builtins.pack_ln - (FStar_Reflection_Data.Tv_Const - FStar_Reflection_Data.C_Unit))) - else - Obj.repr - (FStar_Tactics_Derived.fail ""))))) - uu___1))) - | uu___1 -> Obj.magic (Obj.repr (FStar_Tactics_Derived.fail ""))) - uu___1) + FStar_Reflection_Builtins.inspect_fv fv)) + (fun uu___1 -> + (fun qn -> + if qn = FStar_Reflection_Const.int_lid + then + Obj.magic + (FStar_Tactics_Derived.exact + (FStar_Reflection_Builtins.pack_ln + (FStar_Reflection_Data.Tv_Const + (FStar_Reflection_Data.C_Int + (Prims.of_int (42)))))) + else + if qn = FStar_Reflection_Const.bool_lid + then + Obj.magic + (FStar_Tactics_Derived.exact + (FStar_Reflection_Builtins.pack_ln + (FStar_Reflection_Data.Tv_Const + FStar_Reflection_Data.C_True))) + else + if qn = FStar_Reflection_Const.unit_lid + then + Obj.magic + (FStar_Tactics_Derived.exact + (FStar_Reflection_Builtins.pack_ln + (FStar_Reflection_Data.Tv_Const + FStar_Reflection_Data.C_Unit))) + else Obj.magic (FStar_Tactics_Derived.fail "")) + uu___1)) + | uu___1 -> Obj.magic (FStar_Tactics_Derived.fail "")) uu___1) let rec (simplify_point : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = fun uu___ -> @@ -249,31 +239,29 @@ let rec (simplify_point : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) | FStar_Reflection_Formula.Iff (l, r) -> Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.Simplifier.fst" - (Prims.of_int (217)) - (Prims.of_int (20)) - (Prims.of_int (217)) - (Prims.of_int (38))) - (Prims.mk_range - "FStar.Tactics.Simplifier.fst" - (Prims.of_int (217)) - (Prims.of_int (14)) - (Prims.of_int (264)) - (Prims.of_int (22))) - (Obj.magic - (FStar_Reflection_Formula.term_as_formula' - l)) + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.Simplifier.fst" + (Prims.of_int (217)) + (Prims.of_int (20)) + (Prims.of_int (217)) + (Prims.of_int (38))) + (Prims.mk_range + "FStar.Tactics.Simplifier.fst" + (Prims.of_int (217)) + (Prims.of_int (14)) + (Prims.of_int (264)) + (Prims.of_int (22))) + (Obj.magic + (FStar_Reflection_Formula.term_as_formula' + l)) + (fun uu___3 -> (fun uu___3 -> - (fun uu___3 -> - match uu___3 - with - | FStar_Reflection_Formula.And - (p, q) -> - Obj.magic - (FStar_Tactics_Effect.tac_bind + match uu___3 with + | FStar_Reflection_Formula.And + (p, q) -> + Obj.magic + (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.Simplifier.fst" (Prims.of_int (219)) @@ -413,10 +401,10 @@ let rec (simplify_point : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) uu___8))) uu___6))) uu___4)) - | FStar_Reflection_Formula.Or - (p, q) -> - Obj.magic - (FStar_Tactics_Effect.tac_bind + | FStar_Reflection_Formula.Or + (p, q) -> + Obj.magic + (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.Simplifier.fst" (Prims.of_int (226)) @@ -556,10 +544,10 @@ let rec (simplify_point : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) uu___8))) uu___6))) uu___4)) - | FStar_Reflection_Formula.Implies - (p, q) -> - Obj.magic - (FStar_Tactics_Effect.tac_bind + | FStar_Reflection_Formula.Implies + (p, q) -> + Obj.magic + (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.Simplifier.fst" (Prims.of_int (233)) @@ -662,10 +650,10 @@ let rec (simplify_point : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) uu___8))) uu___6))) uu___4)) - | FStar_Reflection_Formula.Forall - (b, p) -> - Obj.magic - (FStar_Tactics_Effect.tac_bind + | FStar_Reflection_Formula.Forall + (b, p) -> + Obj.magic + (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.Simplifier.fst" (Prims.of_int (239)) @@ -760,10 +748,10 @@ let rec (simplify_point : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) (tiff ())) uu___6))) uu___4)) - | FStar_Reflection_Formula.Exists - (b, p) -> - Obj.magic - (FStar_Tactics_Effect.tac_bind + | FStar_Reflection_Formula.Exists + (b, p) -> + Obj.magic + (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.Simplifier.fst" (Prims.of_int (244)) @@ -858,10 +846,10 @@ let rec (simplify_point : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) (tiff ())) uu___6))) uu___4)) - | FStar_Reflection_Formula.Not - p -> - Obj.magic - (FStar_Tactics_Effect.tac_bind + | FStar_Reflection_Formula.Not + p -> + Obj.magic + (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.Simplifier.fst" (Prims.of_int (249)) @@ -930,10 +918,10 @@ let rec (simplify_point : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) (tiff ())) uu___6))) uu___4)) - | FStar_Reflection_Formula.Iff - (p, q) -> - Obj.magic - (FStar_Tactics_Effect.tac_bind + | FStar_Reflection_Formula.Iff + (p, q) -> + Obj.magic + (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.Simplifier.fst" (Prims.of_int (256)) @@ -1116,15 +1104,14 @@ let rec (simplify_point : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) ())) uu___5))) uu___4)) - | uu___4 -> - Obj.magic - (tiff ())) - uu___3))) + | uu___4 -> + Obj.magic + (tiff ())) + uu___3)) | uu___3 -> Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "simplify_point: failed precondition: goal should be `g <==> ?u`"))) + (FStar_Tactics_Derived.fail + "simplify_point: failed precondition: goal should be `g <==> ?u`")) uu___3))) uu___3))) uu___2))) uu___1) and (recurse : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = @@ -1182,33 +1169,31 @@ and (recurse : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = | FStar_Reflection_Formula.Iff (l, r) -> Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.Simplifier.fst" - (Prims.of_int (275)) - (Prims.of_int (20)) - (Prims.of_int (275)) - (Prims.of_int (38))) - (Prims.mk_range - "FStar.Tactics.Simplifier.fst" - (Prims.of_int (275)) - (Prims.of_int (14)) - (Prims.of_int (302)) - (Prims.of_int (22))) - (Obj.magic - (FStar_Reflection_Formula.term_as_formula' - l)) + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.Simplifier.fst" + (Prims.of_int (275)) + (Prims.of_int (20)) + (Prims.of_int (275)) + (Prims.of_int (38))) + (Prims.mk_range + "FStar.Tactics.Simplifier.fst" + (Prims.of_int (275)) + (Prims.of_int (14)) + (Prims.of_int (302)) + (Prims.of_int (22))) + (Obj.magic + (FStar_Reflection_Formula.term_as_formula' + l)) + (fun uu___3 -> (fun uu___3 -> - (fun uu___3 -> - match uu___3 - with - | FStar_Reflection_Formula.And - (uu___4, - uu___5) - -> - Obj.magic - (FStar_Tactics_Derived.seq + match uu___3 with + | FStar_Reflection_Formula.And + (uu___4, + uu___5) + -> + Obj.magic + (FStar_Tactics_Derived.seq (fun uu___6 -> FStar_Tactics_Derived.apply_lemma @@ -1220,12 +1205,12 @@ and (recurse : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = "Simplifier"; "and_cong"])))) simplify_point) - | FStar_Reflection_Formula.Or - (uu___4, - uu___5) - -> - Obj.magic - (FStar_Tactics_Derived.seq + | FStar_Reflection_Formula.Or + (uu___4, + uu___5) + -> + Obj.magic + (FStar_Tactics_Derived.seq (fun uu___6 -> FStar_Tactics_Derived.apply_lemma @@ -1237,12 +1222,12 @@ and (recurse : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = "Simplifier"; "or_cong"])))) simplify_point) - | FStar_Reflection_Formula.Implies - (uu___4, - uu___5) - -> - Obj.magic - (FStar_Tactics_Derived.seq + | FStar_Reflection_Formula.Implies + (uu___4, + uu___5) + -> + Obj.magic + (FStar_Tactics_Derived.seq (fun uu___6 -> FStar_Tactics_Derived.apply_lemma @@ -1254,12 +1239,12 @@ and (recurse : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = "Simplifier"; "imp_cong"])))) simplify_point) - | FStar_Reflection_Formula.Forall - (uu___4, - uu___5) - -> - Obj.magic - (FStar_Tactics_Effect.tac_bind + | FStar_Reflection_Formula.Forall + (uu___4, + uu___5) + -> + Obj.magic + (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.Simplifier.fst" (Prims.of_int (286)) @@ -1311,12 +1296,12 @@ and (recurse : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = ())) uu___7))) uu___6)) - | FStar_Reflection_Formula.Exists - (uu___4, - uu___5) - -> - Obj.magic - (FStar_Tactics_Effect.tac_bind + | FStar_Reflection_Formula.Exists + (uu___4, + uu___5) + -> + Obj.magic + (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.Simplifier.fst" (Prims.of_int (291)) @@ -1368,10 +1353,10 @@ and (recurse : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = ())) uu___7))) uu___6)) - | FStar_Reflection_Formula.Not - uu___4 -> - Obj.magic - (FStar_Tactics_Effect.tac_bind + | FStar_Reflection_Formula.Not + uu___4 -> + Obj.magic + (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.Simplifier.fst" (Prims.of_int (296)) @@ -1401,12 +1386,12 @@ and (recurse : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = (simplify_point ())) uu___5)) - | FStar_Reflection_Formula.Iff - (uu___4, - uu___5) - -> - Obj.magic - (FStar_Tactics_Derived.seq + | FStar_Reflection_Formula.Iff + (uu___4, + uu___5) + -> + Obj.magic + (FStar_Tactics_Derived.seq (fun uu___6 -> FStar_Tactics_Derived.apply_lemma @@ -1418,15 +1403,14 @@ and (recurse : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = "Simplifier"; "iff_cong"])))) simplify_point) - | uu___4 -> - Obj.magic - (tiff ())) - uu___3))) + | uu___4 -> + Obj.magic + (tiff ())) + uu___3)) | uu___3 -> Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "recurse: failed precondition: goal should be `g <==> ?u`"))) + (FStar_Tactics_Derived.fail + "recurse: failed precondition: goal should be `g <==> ?u`")) uu___3))) uu___3))) uu___2))) uu___1) let (simplify : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = diff --git a/ocaml/fstar-lib/generated/FStar_Tactics_SyntaxHelpers.ml b/ocaml/fstar-lib/generated/FStar_Tactics_SyntaxHelpers.ml index 805016faf18..6f6a6fa5759 100644 --- a/ocaml/fstar-lib/generated/FStar_Tactics_SyntaxHelpers.ml +++ b/ocaml/fstar-lib/generated/FStar_Tactics_SyntaxHelpers.ml @@ -123,11 +123,7 @@ let (collect_abs : match uu___ with | (bs, t') -> ((FStar_List_Tot_Base.rev bs), t'))) let fail : 'a . Prims.string -> ('a, unit) FStar_Tactics_Effect.tac_repr = - fun uu___ -> - (fun m -> - Obj.magic - (FStar_Tactics_Effect.raise (FStar_Tactics_Common.TacticFailure m))) - uu___ + fun m -> FStar_Tactics_Effect.raise (FStar_Tactics_Common.TacticFailure m) let rec (mk_arr : FStar_Reflection_Types.binder Prims.list -> FStar_Reflection_Types.comp -> diff --git a/ocaml/fstar-lib/generated/FStar_Tactics_Typeclasses.ml b/ocaml/fstar-lib/generated/FStar_Tactics_Typeclasses.ml index b296a249006..df399efc6e2 100644 --- a/ocaml/fstar-lib/generated/FStar_Tactics_Typeclasses.ml +++ b/ocaml/fstar-lib/generated/FStar_Tactics_Typeclasses.ml @@ -4,18 +4,13 @@ let rec first : ('a -> ('b, unit) FStar_Tactics_Effect.tac_repr) -> 'a Prims.list -> ('b, unit) FStar_Tactics_Effect.tac_repr = - fun uu___1 -> - fun uu___ -> - (fun f -> - fun l -> - match l with - | [] -> - Obj.magic (Obj.repr (FStar_Tactics_Derived.fail "no cands")) - | x::xs -> - Obj.magic - (Obj.repr - (FStar_Tactics_Derived.or_else (fun uu___ -> f x) - (fun uu___ -> first f xs)))) uu___1 uu___ + fun f -> + fun l -> + match l with + | [] -> FStar_Tactics_Derived.fail "no cands" + | x::xs -> + FStar_Tactics_Derived.or_else (fun uu___ -> f x) + (fun uu___ -> first f xs) let rec (tcresolve' : FStar_Reflection_Types.term Prims.list -> Prims.int -> (unit, unit) FStar_Tactics_Effect.tac_repr) @@ -28,8 +23,10 @@ let rec (tcresolve' : (Prims.mk_range "FStar.Tactics.Typeclasses.fst" (Prims.of_int (45)) (Prims.of_int (4)) (Prims.of_int (50)) (Prims.of_int (137))) (if fuel <= Prims.int_zero - then FStar_Tactics_Derived.fail "out of fuel" - else FStar_Tactics_Effect.lift_div_tac (fun uu___1 -> ())) + then Obj.magic (Obj.repr (FStar_Tactics_Derived.fail "out of fuel")) + else + Obj.magic + (Obj.repr (FStar_Tactics_Effect.lift_div_tac (fun uu___1 -> ())))) (fun uu___ -> (fun uu___ -> Obj.magic @@ -75,10 +72,15 @@ let rec (tcresolve' : (FStar_Reflection_Builtins.term_eq g) seen then - FStar_Tactics_Derived.fail "loop" + Obj.magic + (Obj.repr + (FStar_Tactics_Derived.fail + "loop")) else - FStar_Tactics_Effect.lift_div_tac - (fun uu___3 -> ())) + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___3 -> ())))) (fun uu___2 -> (fun uu___2 -> Obj.magic @@ -153,7 +155,11 @@ let rec (tcresolve' : uu___5)))) (fun uu___5 -> - FStar_Tactics_Derived.fail + (fun + uu___5 -> + Obj.magic + (FStar_Tactics_Derived.fail + uu___5)) uu___5))))) uu___3))) uu___2))) uu___2))) uu___1))) uu___) @@ -354,15 +360,11 @@ let (tcresolve : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = (fun uu___2 -> match () with | () -> tcresolve' [] (Prims.of_int (16))) (fun uu___2 -> - (fun uu___2 -> - match uu___2 with - | FStar_Tactics_Common.TacticFailure s -> - Obj.magic - (FStar_Tactics_Derived.fail - (Prims.strcat "Typeclass resolution failed: " - s)) - | e -> Obj.magic (FStar_Tactics_Effect.raise e)) - uu___2))) uu___1) + match uu___2 with + | FStar_Tactics_Common.TacticFailure s -> + FStar_Tactics_Derived.fail + (Prims.strcat "Typeclass resolution failed: " s) + | e -> FStar_Tactics_Effect.raise e))) uu___1) let _ = FStar_Tactics_Native.register_tactic "FStar.Tactics.Typeclasses.tcresolve" (Prims.of_int (2)) @@ -1098,14 +1100,11 @@ let (mk_class : FStar_Pervasives_Native.None -> Obj.magic - (Obj.repr (FStar_Tactics_Derived.fail - "mk_class: proj not found?")) + "mk_class: proj not found?") | FStar_Pervasives_Native.Some se1 -> - Obj.magic - (Obj.repr (match FStar_Reflection_Builtins.inspect_sigelt se1 @@ -1114,7 +1113,7 @@ let (mk_class : FStar_Reflection_Data.Sg_Let (uu___10, lbs) -> - Obj.repr + Obj.magic (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.Typeclasses.fst" @@ -1155,9 +1154,9 @@ let (mk_class : | uu___10 -> - Obj.repr + Obj.magic (FStar_Tactics_Derived.fail - "mk_class: proj not Sg_Let?")))) + "mk_class: proj not Sg_Let?"))) uu___9))) (fun uu___9 -> @@ -1242,14 +1241,12 @@ let (mk_class : | [] -> Obj.magic - (Obj.repr (FStar_Tactics_Derived.fail - "mk_class: impossible, no binders")) + "mk_class: impossible, no binders") | b1::bs' -> Obj.magic - (Obj.repr (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.Typeclasses.fst" @@ -1314,7 +1311,7 @@ let (mk_class : :: bs')) cod2)) uu___12))) - uu___11))))) + uu___11)))) uu___10))) uu___9))) (fun diff --git a/ocaml/fstar-lib/generated/Steel_Effect_Common.ml b/ocaml/fstar-lib/generated/Steel_Effect_Common.ml index cdd2331bbcf..afb14f087b1 100644 --- a/ocaml/fstar-lib/generated/Steel_Effect_Common.ml +++ b/ocaml/fstar-lib/generated/Steel_Effect_Common.ml @@ -804,15 +804,26 @@ let (name_appears_in : (Prims.of_int (730)) (Prims.of_int (13))) (Obj.magic (FStar_Tactics_Builtins.inspect t1)) (fun uu___1 -> - match uu___1 with - | FStar_Reflection_Data.Tv_FVar fv -> - if (FStar_Reflection_Builtins.inspect_fv fv) = nm - then FStar_Tactics_Effect.raise Appears - else - FStar_Tactics_Effect.lift_div_tac - (fun uu___3 -> ()) - | t2 -> - FStar_Tactics_Effect.lift_div_tac (fun uu___2 -> ())))) + (fun uu___1 -> + match uu___1 with + | FStar_Reflection_Data.Tv_FVar fv -> + Obj.magic + (Obj.repr + (if + (FStar_Reflection_Builtins.inspect_fv fv) + = nm + then + Obj.repr + (FStar_Tactics_Effect.raise Appears) + else + Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___3 -> ())))) + | t2 -> + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___2 -> ())))) uu___1))) (fun uu___ -> (fun ff -> Obj.magic @@ -847,9 +858,12 @@ let (name_appears_in : match uu___ with | Appears -> Obj.magic - (FStar_Tactics_Effect.lift_div_tac - (fun uu___1 -> true)) - | e -> Obj.magic (FStar_Tactics_Effect.raise e)) + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___1 -> true))) + | e -> + Obj.magic + (Obj.repr (FStar_Tactics_Effect.raise e))) uu___))) uu___) let (term_appears_in : FStar_Reflection_Types.term -> @@ -1298,13 +1312,17 @@ let rec (try_candidates : (select t am) (select hd am))) (fun uu___3 -> - if uu___3 - then - FStar_Tactics_Effect.raise - Success - else - FStar_Tactics_Effect.raise - Failed)) + (fun uu___3 -> + if uu___3 + then + Obj.magic + (FStar_Tactics_Effect.raise + Success) + else + Obj.magic + (FStar_Tactics_Effect.raise + Failed)) + uu___3)) (fun uu___2 -> (fun uu___2 -> Obj.magic @@ -1366,8 +1384,9 @@ let rec (remove_from_list : (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (863)) (Prims.of_int (72)) (Prims.of_int (863)) (Prims.of_int (74))) - (FStar_Tactics_Derived.fail - "atom in remove_from_list not found: should not happen") + (Obj.magic + (FStar_Tactics_Derived.fail + "atom in remove_from_list not found: should not happen")) (fun uu___ -> FStar_Tactics_Effect.lift_div_tac (fun uu___1 -> [])))) @@ -1729,10 +1748,16 @@ let rec (try_unifying_remaining : (FStar_Tactics_Derived.unify u (select hd am))) (fun uu___1 -> - if uu___1 - then - FStar_Tactics_Effect.raise Success - else FStar_Tactics_Effect.raise Failed)) + (fun uu___1 -> + if uu___1 + then + Obj.magic + (FStar_Tactics_Effect.raise + Success) + else + Obj.magic + (FStar_Tactics_Effect.raise + Failed)) uu___1)) (fun uu___ -> match uu___ with | Success -> try_unifying_remaining tl u am @@ -1773,8 +1798,11 @@ let rec (try_unifying_remaining : "could not find candidate for scrutinee " uu___2)))) (fun uu___2 -> - FStar_Tactics_Derived.fail uu___2))))) - uu___2 uu___1 uu___ + (fun uu___2 -> + Obj.magic + (FStar_Tactics_Derived.fail + uu___2)) uu___2))))) uu___2 + uu___1 uu___ let (is_smt_binder : FStar_Reflection_Types.binder -> (Prims.bool, unit) FStar_Tactics_Effect.tac_repr) @@ -2057,11 +2085,13 @@ let rec (new_args_for_smt_attrs : with | FStar_Reflection_Data.C_Total ty2 -> - FStar_Tactics_Effect.lift_div_tac + Obj.magic ( - fun + Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___1 -> - ty2) + ty2))) | FStar_Reflection_Data.C_Eff (uu___1, eff_name, @@ -2069,22 +2099,30 @@ let rec (new_args_for_smt_attrs : uu___2, uu___3) -> - if + Obj.magic + ( + Obj.repr + (if eff_name = ["Prims"; "Tot"] - then - FStar_Tactics_Effect.lift_div_tac + then + Obj.repr + (FStar_Tactics_Effect.lift_div_tac (fun uu___4 -> - ty2) - else - FStar_Tactics_Derived.fail - "computation type not supported in definition of slprops" + ty2)) + else + Obj.repr + (FStar_Tactics_Derived.fail + "computation type not supported in definition of slprops"))) | uu___1 -> - FStar_Tactics_Derived.fail - "computation type not supported in definition of slprops") + Obj.magic + ( + Obj.repr + (FStar_Tactics_Derived.fail + "computation type not supported in definition of slprops"))) (fun uu___1 -> (fun ty2 -> Obj.magic @@ -2322,7 +2360,9 @@ let fail_atoms : (fun uu___1 -> Prims.strcat "could not find a solution for unifying\n" uu___)))) - (fun uu___ -> FStar_Tactics_Derived.fail uu___) + (fun uu___ -> + (fun uu___ -> Obj.magic (FStar_Tactics_Derived.fail uu___)) + uu___) let rec (equivalent_lists_fallback : Prims.nat -> atom Prims.list -> @@ -2431,7 +2471,11 @@ let rec (equivalent_lists_fallback : "could not find candidates for " uu___2)))) (fun uu___2 -> - FStar_Tactics_Derived.fail + (fun uu___2 + -> + Obj.magic + (FStar_Tactics_Derived.fail + uu___2)) uu___2)))) uu___)) | uu___ -> @@ -2473,8 +2517,10 @@ let rec (equivalent_lists_fallback : "could not find candidates for " uu___1)))) (fun uu___1 -> - FStar_Tactics_Derived.fail - uu___1)))) + (fun uu___1 -> + Obj.magic + (FStar_Tactics_Derived.fail + uu___1)) uu___1)))) | uu___ -> Obj.magic (Obj.repr @@ -2721,7 +2767,11 @@ let rec (equivalent_lists' : uu___2)))) (fun uu___2 -> - FStar_Tactics_Derived.fail + (fun + uu___2 -> + Obj.magic + (FStar_Tactics_Derived.fail + uu___2)) uu___2)))) uu___)) | uu___ -> @@ -2763,8 +2813,10 @@ let rec (equivalent_lists' : "could not find candidates for " uu___1)))) (fun uu___1 -> - FStar_Tactics_Derived.fail - uu___1)))) + (fun uu___1 -> + Obj.magic + (FStar_Tactics_Derived.fail + uu___1)) uu___1)))) | uu___ -> Obj.magic (Obj.repr @@ -3119,13 +3171,18 @@ let rec (unifies_with_all_uvars : t hd_t)) (fun uu___3 -> + (fun + uu___3 -> if uu___3 then - FStar_Tactics_Effect.raise - Success + Obj.magic + (FStar_Tactics_Effect.raise + Success) else - FStar_Tactics_Effect.raise + Obj.magic + (FStar_Tactics_Effect.raise Failed)) + uu___3)) (fun uu___2 -> (fun @@ -4286,93 +4343,96 @@ let rec (unify_pr_with_true : | (hd, tl) -> if is_and hd then - Obj.magic - (Obj.repr - (match tl with - | (pr_l, uu___1)::(pr_r, uu___2)::[] -> - Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (1795)) (Prims.of_int (6)) - (Prims.of_int (1795)) (Prims.of_int (29))) - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (1796)) (Prims.of_int (6)) - (Prims.of_int (1796)) (Prims.of_int (29))) - (Obj.magic (unify_pr_with_true pr_l)) - (fun uu___3 -> - (fun uu___3 -> - Obj.magic (unify_pr_with_true pr_r)) - uu___3)) - | uu___1 -> - Obj.repr - (FStar_Tactics_Derived.fail - "unify_pr_with_true: ill-formed /\\"))) + (match tl with + | (pr_l, uu___1)::(pr_r, uu___2)::[] -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (1795)) (Prims.of_int (6)) + (Prims.of_int (1795)) (Prims.of_int (29))) + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (1796)) (Prims.of_int (6)) + (Prims.of_int (1796)) (Prims.of_int (29))) + (Obj.magic (unify_pr_with_true pr_l)) + (fun uu___3 -> + (fun uu___3 -> + Obj.magic (unify_pr_with_true pr_r)) uu___3)) + | uu___1 -> + Obj.magic + (FStar_Tactics_Derived.fail + "unify_pr_with_true: ill-formed /\\")) else Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (1799)) (Prims.of_int (10)) - (Prims.of_int (1799)) (Prims.of_int (30))) - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (1799)) (Prims.of_int (4)) - (Prims.of_int (1811)) (Prims.of_int (73))) - (Obj.magic - (FStar_Tactics_Derived.inspect_unascribe hd)) + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (1799)) (Prims.of_int (10)) + (Prims.of_int (1799)) (Prims.of_int (30))) + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (1799)) (Prims.of_int (4)) + (Prims.of_int (1811)) (Prims.of_int (73))) + (Obj.magic + (FStar_Tactics_Derived.inspect_unascribe hd)) + (fun uu___2 -> (fun uu___2 -> - (fun uu___2 -> - match uu___2 with - | FStar_Reflection_Data.Tv_Uvar - (uu___3, uu___4) -> - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (1801)) - (Prims.of_int (9)) - (Prims.of_int (1801)) - (Prims.of_int (27))) - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (1801)) - (Prims.of_int (6)) - (Prims.of_int (1805)) - (Prims.of_int (9))) - (Obj.magic - (FStar_Tactics_Derived.unify - pr - (FStar_Reflection_Builtins.pack_ln - (FStar_Reflection_Data.Tv_FVar - (FStar_Reflection_Builtins.pack_fv - ["Steel"; - "Effect"; - "Common"; - "true_p"]))))) + match uu___2 with + | FStar_Reflection_Data.Tv_Uvar (uu___3, uu___4) + -> + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (1801)) + (Prims.of_int (9)) + (Prims.of_int (1801)) + (Prims.of_int (27))) + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (1801)) + (Prims.of_int (6)) + (Prims.of_int (1805)) + (Prims.of_int (9))) + (Obj.magic + (FStar_Tactics_Derived.unify pr + (FStar_Reflection_Builtins.pack_ln + (FStar_Reflection_Data.Tv_FVar + (FStar_Reflection_Builtins.pack_fv + ["Steel"; + "Effect"; + "Common"; + "true_p"]))))) + (fun uu___5 -> (fun uu___5 -> if uu___5 then - FStar_Tactics_Effect.lift_div_tac - (fun uu___6 -> ()) + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___6 -> ()))) else - FStar_Tactics_Derived.fail - "unify_pr_with_true: could not unify SMT prop with True"))) - | uu___3 -> - Obj.magic - (Obj.repr - (if - (FStar_List_Tot_Base.length - (FStar_Reflection_Builtins.free_uvars - pr)) - = Prims.int_zero - then - FStar_Tactics_Effect.lift_div_tac - (fun uu___4 -> ()) - else - FStar_Tactics_Effect.raise + Obj.magic + (Obj.repr + (FStar_Tactics_Derived.fail + "unify_pr_with_true: could not unify SMT prop with True"))) + uu___5))) + | uu___3 -> + Obj.magic + (Obj.repr + (if + (FStar_List_Tot_Base.length + (FStar_Reflection_Builtins.free_uvars + pr)) + = Prims.int_zero + then + Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___4 -> ())) + else + Obj.repr + (FStar_Tactics_Effect.raise (Postpone - "unify_pr_with_true: some uvars are still there")))) - uu___2)))) uu___) + "unify_pr_with_true: some uvars are still there"))))) + uu___2))) uu___) let rec (set_abduction_variable_term : FStar_Reflection_Types.term -> (FStar_Reflection_Types.term, unit) FStar_Tactics_Effect.tac_repr) @@ -4391,112 +4451,27 @@ let rec (set_abduction_variable_term : | (hd, tl) -> if is_and hd then - Obj.magic - (Obj.repr - (match tl with - | (pr_l, FStar_Reflection_Data.Q_Explicit)::(pr_r, - FStar_Reflection_Data.Q_Explicit)::[] - -> - Obj.repr - (if - (FStar_List_Tot_Base.length - (FStar_Reflection_Builtins.free_uvars - pr_r)) - = Prims.int_zero - then - Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (1838)) - (Prims.of_int (18)) - (Prims.of_int (1838)) - (Prims.of_int (50))) - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (1839)) - (Prims.of_int (8)) - (Prims.of_int (1839)) - (Prims.of_int (53))) - (Obj.magic - (set_abduction_variable_term pr_l)) - (fun arg -> - FStar_Tactics_Effect.lift_div_tac - (fun uu___1 -> - FStar_Reflection_Derived.mk_app - (FStar_Reflection_Builtins.pack_ln - (FStar_Reflection_Data.Tv_FVar - (FStar_Reflection_Builtins.pack_fv - ["Steel"; - "Effect"; - "Common"; - "elim_and_l_squash"]))) - [(arg, - FStar_Reflection_Data.Q_Explicit)]))) - else - Obj.repr - (if - (FStar_List_Tot_Base.length - (FStar_Reflection_Builtins.free_uvars - pr_l)) - = Prims.int_zero - then - Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (1842)) - (Prims.of_int (18)) - (Prims.of_int (1842)) - (Prims.of_int (50))) - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (1843)) - (Prims.of_int (8)) - (Prims.of_int (1843)) - (Prims.of_int (53))) - (Obj.magic - (set_abduction_variable_term - pr_r)) - (fun arg -> - FStar_Tactics_Effect.lift_div_tac - (fun uu___2 -> - FStar_Reflection_Derived.mk_app - (FStar_Reflection_Builtins.pack_ln - (FStar_Reflection_Data.Tv_FVar - (FStar_Reflection_Builtins.pack_fv - ["Steel"; - "Effect"; - "Common"; - "elim_and_r_squash"]))) - [(arg, - FStar_Reflection_Data.Q_Explicit)]))) - else - Obj.repr - (FStar_Tactics_Effect.raise - (Postpone - "set_abduction_variable_term: there are still uvars on both sides of l_and")))) - | uu___1 -> - Obj.repr - (FStar_Tactics_Derived.fail - "set_abduction_variable: ill-formed /\\"))) - else - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (1831)) (Prims.of_int (6)) - (Prims.of_int (1831)) (Prims.of_int (8))) - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (1849)) (Prims.of_int (4)) - (Prims.of_int (1852)) (Prims.of_int (54))) - (Obj.magic (FStar_Tactics_Builtins.inspect hd)) - (fun uu___2 -> - match uu___2 with - | FStar_Reflection_Data.Tv_Uvar (uu___3, uu___4) - -> + (match tl with + | (pr_l, FStar_Reflection_Data.Q_Explicit)::(pr_r, + FStar_Reflection_Data.Q_Explicit)::[] + -> + if + (FStar_List_Tot_Base.length + (FStar_Reflection_Builtins.free_uvars pr_r)) + = Prims.int_zero + then + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (1838)) (Prims.of_int (18)) + (Prims.of_int (1838)) (Prims.of_int (50))) + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (1839)) (Prims.of_int (8)) + (Prims.of_int (1839)) (Prims.of_int (53))) + (Obj.magic (set_abduction_variable_term pr_l)) + (fun arg -> FStar_Tactics_Effect.lift_div_tac - (fun uu___5 -> + (fun uu___1 -> FStar_Reflection_Derived.mk_app (FStar_Reflection_Builtins.pack_ln (FStar_Reflection_Data.Tv_FVar @@ -4504,15 +4479,83 @@ let rec (set_abduction_variable_term : ["Steel"; "Effect"; "Common"; - "_return_squash"]))) - [((FStar_Reflection_Builtins.pack_ln - (FStar_Reflection_Data.Tv_Const - FStar_Reflection_Data.C_Unit)), - FStar_Reflection_Data.Q_Explicit)]) + "elim_and_l_squash"]))) + [(arg, + FStar_Reflection_Data.Q_Explicit)]))) + else + if + (FStar_List_Tot_Base.length + (FStar_Reflection_Builtins.free_uvars pr_l)) + = Prims.int_zero + then + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (1842)) (Prims.of_int (18)) + (Prims.of_int (1842)) (Prims.of_int (50))) + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (1843)) (Prims.of_int (8)) + (Prims.of_int (1843)) (Prims.of_int (53))) + (Obj.magic (set_abduction_variable_term pr_r)) + (fun arg -> + FStar_Tactics_Effect.lift_div_tac + (fun uu___2 -> + FStar_Reflection_Derived.mk_app + (FStar_Reflection_Builtins.pack_ln + (FStar_Reflection_Data.Tv_FVar + (FStar_Reflection_Builtins.pack_fv + ["Steel"; + "Effect"; + "Common"; + "elim_and_r_squash"]))) + [(arg, + FStar_Reflection_Data.Q_Explicit)]))) + else + Obj.magic + (FStar_Tactics_Effect.raise + (Postpone + "set_abduction_variable_term: there are still uvars on both sides of l_and")) + | uu___1 -> + Obj.magic + (FStar_Tactics_Derived.fail + "set_abduction_variable: ill-formed /\\")) + else + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (1831)) (Prims.of_int (6)) + (Prims.of_int (1831)) (Prims.of_int (8))) + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (1849)) (Prims.of_int (4)) + (Prims.of_int (1852)) (Prims.of_int (54))) + (Obj.magic (FStar_Tactics_Builtins.inspect hd)) + (fun uu___2 -> + (fun uu___2 -> + match uu___2 with + | FStar_Reflection_Data.Tv_Uvar (uu___3, uu___4) + -> + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___5 -> + FStar_Reflection_Derived.mk_app + (FStar_Reflection_Builtins.pack_ln + (FStar_Reflection_Data.Tv_FVar + (FStar_Reflection_Builtins.pack_fv + ["Steel"; + "Effect"; + "Common"; + "_return_squash"]))) + [((FStar_Reflection_Builtins.pack_ln + (FStar_Reflection_Data.Tv_Const + FStar_Reflection_Data.C_Unit)), + FStar_Reflection_Data.Q_Explicit)]))) | uu___3 -> - FStar_Tactics_Derived.fail - "set_abduction_variable: cannot unify")))) - uu___) + Obj.magic + (Obj.repr + (FStar_Tactics_Derived.fail + "set_abduction_variable: cannot unify"))) + uu___2))) uu___) let (set_abduction_variable : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = fun uu___ -> @@ -4538,68 +4581,62 @@ let (set_abduction_variable : match uu___1 with | FStar_Reflection_Data.Tv_Arrow (b, uu___2) -> Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (1858)) - (Prims.of_int (18)) - (Prims.of_int (1858)) - (Prims.of_int (34))) - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (1858)) - (Prims.of_int (4)) - (Prims.of_int (1861)) - (Prims.of_int (42))) - (FStar_Tactics_Effect.lift_div_tac - (fun uu___3 -> - FStar_Reflection_Builtins.inspect_binder - b)) + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (1858)) (Prims.of_int (18)) + (Prims.of_int (1858)) (Prims.of_int (34))) + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (1858)) (Prims.of_int (4)) + (Prims.of_int (1861)) (Prims.of_int (42))) + (FStar_Tactics_Effect.lift_div_tac (fun uu___3 -> - (fun uu___3 -> - match uu___3 with - | (bv, uu___4) -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (1859)) - (Prims.of_int (13)) - (Prims.of_int (1859)) - (Prims.of_int (26))) - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (1860)) - (Prims.of_int (4)) - (Prims.of_int (1861)) - (Prims.of_int (42))) - (FStar_Tactics_Effect.lift_div_tac - (fun uu___5 -> - FStar_Reflection_Builtins.inspect_bv - bv)) + FStar_Reflection_Builtins.inspect_binder + b)) + (fun uu___3 -> + (fun uu___3 -> + match uu___3 with + | (bv, uu___4) -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (1859)) + (Prims.of_int (13)) + (Prims.of_int (1859)) + (Prims.of_int (26))) + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (1860)) + (Prims.of_int (4)) + (Prims.of_int (1861)) + (Prims.of_int (42))) + (FStar_Tactics_Effect.lift_div_tac (fun uu___5 -> - (fun bv1 -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (1860)) - (Prims.of_int (13)) - (Prims.of_int (1860)) - (Prims.of_int (23))) - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (1861)) - (Prims.of_int (4)) - (Prims.of_int (1861)) - (Prims.of_int (42))) - (FStar_Tactics_Effect.lift_div_tac - (fun uu___5 - -> - bv1.FStar_Reflection_Data.bv_sort)) + FStar_Reflection_Builtins.inspect_bv + bv)) + (fun uu___5 -> + (fun bv1 -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (1860)) + (Prims.of_int (13)) + (Prims.of_int (1860)) + (Prims.of_int (23))) + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (1861)) + (Prims.of_int (4)) + (Prims.of_int (1861)) + (Prims.of_int (42))) + (FStar_Tactics_Effect.lift_div_tac (fun uu___5 -> - (fun pr -> - Obj.magic - (FStar_Tactics_Effect.tac_bind + bv1.FStar_Reflection_Data.bv_sort)) + (fun uu___5 -> + (fun pr -> + Obj.magic + (FStar_Tactics_Effect.tac_bind (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (1861)) @@ -4623,13 +4660,12 @@ let (set_abduction_variable : (FStar_Tactics_Derived.exact uu___5)) uu___5))) - uu___5))) - uu___5))) uu___3))) + uu___5))) + uu___5))) uu___3)) | uu___2 -> Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "Not an arrow goal"))) uu___1))) uu___1) + (FStar_Tactics_Derived.fail "Not an arrow goal")) + uu___1))) uu___1) let (canon_l_r : Prims.bool -> FStar_Reflection_Types.term -> @@ -4814,11 +4850,15 @@ let (canon_l_r : (flatten r2_raw) am2)) + (fun + uu___4 -> (fun res -> - FStar_Tactics_Effect.raise + Obj.magic + (FStar_Tactics_Effect.raise (Result res))) + uu___4)) (fun uu___3 -> (fun @@ -4829,21 +4869,24 @@ let (canon_l_r : FStar_Tactics_Common.TacticFailure m1 -> Obj.magic + (Obj.repr (FStar_Tactics_Derived.fail - m1) + m1)) | Result res -> Obj.magic + (Obj.repr (FStar_Tactics_Effect.lift_div_tac (fun uu___4 -> - res)) + res))) | uu___4 -> Obj.magic + (Obj.repr (FStar_Tactics_Derived.fail - "uncaught exception in equivalent_lists")) + "uncaught exception in equivalent_lists"))) uu___3))) (fun uu___3 -> @@ -6534,25 +6577,19 @@ let (canon_l_r : (fun uu___11 -> - (fun - uu___11 - -> match uu___11 with | FStar_Tactics_Common.TacticFailure msg -> - Obj.magic - (FStar_Tactics_Derived.fail + FStar_Tactics_Derived.fail (Prims.strcat "Cannot unify pr with true: " - msg)) + msg) | e -> - Obj.magic - (FStar_Tactics_Effect.raise - e)) - uu___11))) + FStar_Tactics_Effect.raise + e))) | l -> Obj.magic @@ -6795,40 +6832,34 @@ let (canon_monoid : (match rel_xy with | (rel_xy1, uu___2)::[] -> Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2090)) - (Prims.of_int (21)) - (Prims.of_int (2090)) - (Prims.of_int (43))) - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2090)) - (Prims.of_int (7)) - (Prims.of_int (2100)) - (Prims.of_int (8))) - (FStar_Tactics_Effect.lift_div_tac - (fun uu___3 - -> - FStar_Reflection_Derived_Lemmas.collect_app_ref - rel_xy1)) + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2090)) + (Prims.of_int (21)) + (Prims.of_int (2090)) + (Prims.of_int (43))) + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2090)) + (Prims.of_int (7)) + (Prims.of_int (2100)) + (Prims.of_int (8))) + (FStar_Tactics_Effect.lift_div_tac (fun uu___3 -> - (fun uu___3 - -> - match uu___3 - with - | - (rel, xy) + FStar_Reflection_Derived_Lemmas.collect_app_ref + rel_xy1)) + (fun uu___3 -> + (fun uu___3 -> + match uu___3 + with + | (rel, xy) -> if (FStar_List_Tot_Base.length xy) >= (Prims.of_int (2)) then - Obj.magic - (Obj.repr (match ((FStar_List_Tot_Base.index xy @@ -6847,7 +6878,7 @@ let (canon_monoid : (rhs, FStar_Reflection_Data.Q_Explicit)) -> - Obj.repr + Obj.magic (canon_l_r use_smt carrier_t @@ -6857,20 +6888,18 @@ let (canon_monoid : rhs) | uu___4 -> - Obj.repr + Obj.magic (FStar_Tactics_Derived.fail - "Goal should have been an application of a binary relation to 2 explicit arguments"))) + "Goal should have been an application of a binary relation to 2 explicit arguments")) else Obj.magic - (Obj.repr (FStar_Tactics_Derived.fail - "Goal should have been an application of a binary relation to n implicit and 2 explicit arguments"))) - uu___3))) + "Goal should have been an application of a binary relation to n implicit and 2 explicit arguments")) + uu___3)) | uu___2 -> Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "Goal should be squash applied to a binary relation")))) + (FStar_Tactics_Derived.fail + "Goal should be squash applied to a binary relation"))) uu___1))) uu___1))) uu___) let (canon' : Prims.bool -> @@ -7736,15 +7765,13 @@ let rec (extract_contexts : FStar_Pervasives_Native.None -> Obj.magic - (Obj.repr (FStar_Tactics_Derived.fail - "no context on the right either")) + "no context on the right either") | FStar_Pervasives_Native.Some g -> Obj.magic - (Obj.repr - (g ()))) + (g ())) uu___5))))))) uu___1))) uu___1)) | uu___1 -> @@ -7991,9 +8018,15 @@ let (open_existentials : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) "Common"; "solve_can_be_split_lookup"]))) e) then - FStar_Tactics_Derived.fail - "Tactic disabled: no available lemmas in context" - else FStar_Tactics_Effect.lift_div_tac (fun uu___2 -> ())) + Obj.magic + (Obj.repr + (FStar_Tactics_Derived.fail + "Tactic disabled: no available lemmas in context")) + else + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___2 -> ())))) (fun uu___1 -> (fun uu___1 -> Obj.magic @@ -8036,44 +8069,40 @@ let (open_existentials : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) FStar_Reflection_Data.Q_Explicit)::[]) -> Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2309)) - (Prims.of_int (22)) - (Prims.of_int (2309)) - (Prims.of_int (36))) - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2309)) - (Prims.of_int (9)) - (Prims.of_int (2325)) - (Prims.of_int (60))) - (FStar_Tactics_Effect.lift_div_tac - (fun uu___4 -> - FStar_Reflection_Derived.collect_app - t1)) + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2309)) + (Prims.of_int (22)) + (Prims.of_int (2309)) + (Prims.of_int (36))) + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2309)) + (Prims.of_int (9)) + (Prims.of_int (2325)) + (Prims.of_int (60))) + (FStar_Tactics_Effect.lift_div_tac (fun uu___4 -> - (fun uu___4 -> - match uu___4 - with - | (hd, tl) -> - if - FStar_Reflection_Derived.is_fvar + FStar_Reflection_Derived.collect_app + t1)) + (fun uu___4 -> + (fun uu___4 -> + match uu___4 with + | (hd, tl) -> + if + FStar_Reflection_Derived.is_fvar hd "Steel.Effect.Common.can_be_split" - then - Obj.magic - (Obj.repr - (match tl - with - | - uu___5:: + then + (match tl + with + | + uu___5:: (rhs, FStar_Reflection_Data.Q_Explicit)::[] -> - Obj.repr + Obj.magic (FStar_Tactics_Effect.tac_bind (Prims.mk_range "Steel.Effect.Common.fsti" @@ -8100,14 +8129,12 @@ let (open_existentials : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) FStar_Pervasives_Native.None -> Obj.magic - (Obj.repr (FStar_Tactics_Derived.fail - "open_existentials: no context found")) + "open_existentials: no context found") | FStar_Pervasives_Native.Some f -> Obj.magic - (Obj.repr (FStar_Tactics_Effect.tac_bind (Prims.mk_range "Steel.Effect.Common.fsti" @@ -8205,24 +8232,23 @@ let (open_existentials : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) uu___10))) uu___9))) uu___8))) - uu___7)))) + uu___7))) uu___6)) - | - uu___5 -> - Obj.repr - (FStar_Tactics_Derived.fail - "open_existentials: ill-formed can_be_split"))) - else + | + uu___5 -> Obj.magic - (Obj.repr (FStar_Tactics_Derived.fail - "open_existentials: not a can_be_split goal"))) - uu___4))) + "open_existentials: ill-formed can_be_split")) + else + Obj.magic + ( + FStar_Tactics_Derived.fail + "open_existentials: not a can_be_split goal")) + uu___4)) | uu___3 -> Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "open_existentials: not a squash goal"))) + (FStar_Tactics_Derived.fail + "open_existentials: not a squash goal")) uu___3))) uu___2))) uu___1))) uu___1) let (try_open_existentials : @@ -8541,104 +8567,85 @@ let (solve_can_be_split_dep : FStar_Reflection_Data.argv Prims.list -> (Prims.bool, unit) FStar_Tactics_Effect.tac_repr) = - fun uu___ -> - (fun args -> - match args with - | (p, uu___)::(t1, uu___1)::(t2, uu___2)::[] -> - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2388)) (Prims.of_int (17)) - (Prims.of_int (2388)) (Prims.of_int (36))) - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2389)) (Prims.of_int (6)) - (Prims.of_int (2420)) (Prims.of_int (18))) - (Obj.magic (slterm_nbr_uvars t1)) - (fun uu___3 -> - (fun lnbr -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2389)) (Prims.of_int (17)) - (Prims.of_int (2389)) (Prims.of_int (36))) - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2390)) (Prims.of_int (6)) - (Prims.of_int (2420)) (Prims.of_int (18))) - (Obj.magic (slterm_nbr_uvars t2)) - (fun uu___3 -> - (fun rnbr -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2391)) - (Prims.of_int (8)) - (Prims.of_int (2393)) - (Prims.of_int (18))) - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2390)) - (Prims.of_int (6)) - (Prims.of_int (2420)) - (Prims.of_int (18))) - (if (lnbr + rnbr) <= Prims.int_one - then - Obj.magic - (Obj.repr (unfold_guard ())) - else - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___4 -> false)))) - (fun uu___3 -> - (fun uu___3 -> - if uu___3 - then - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2396)) - (Prims.of_int (8)) - (Prims.of_int (2416)) - (Prims.of_int (36))) - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2418)) - (Prims.of_int (8)) - (Prims.of_int (2418)) - (Prims.of_int (12))) - (Obj.magic - (FStar_Tactics_Derived.focus - (fun uu___4 -> - FStar_Tactics_Effect.tac_bind - ( - Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2397)) - (Prims.of_int (23)) - (Prims.of_int (2397)) - (Prims.of_int (39))) - ( - Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2398)) - (Prims.of_int (10)) - (Prims.of_int (2416)) - (Prims.of_int (35))) - ( - Obj.magic - (FStar_Tactics_Logic.implies_intro - ())) - ( - fun - uu___5 -> - (fun - p_bind -> - Obj.magic - (FStar_Tactics_Effect.tac_bind + fun args -> + match args with + | (p, uu___)::(t1, uu___1)::(t2, uu___2)::[] -> + FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2388)) + (Prims.of_int (17)) (Prims.of_int (2388)) (Prims.of_int (36))) + (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2389)) + (Prims.of_int (6)) (Prims.of_int (2420)) (Prims.of_int (18))) + (Obj.magic (slterm_nbr_uvars t1)) + (fun uu___3 -> + (fun lnbr -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2389)) (Prims.of_int (17)) + (Prims.of_int (2389)) (Prims.of_int (36))) + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2390)) (Prims.of_int (6)) + (Prims.of_int (2420)) (Prims.of_int (18))) + (Obj.magic (slterm_nbr_uvars t2)) + (fun uu___3 -> + (fun rnbr -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2391)) (Prims.of_int (8)) + (Prims.of_int (2393)) (Prims.of_int (18))) + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2390)) (Prims.of_int (6)) + (Prims.of_int (2420)) (Prims.of_int (18))) + (if (lnbr + rnbr) <= Prims.int_one + then Obj.magic (Obj.repr (unfold_guard ())) + else + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___4 -> false)))) + (fun uu___3 -> + (fun uu___3 -> + if uu___3 + then + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2396)) + (Prims.of_int (8)) + (Prims.of_int (2416)) + (Prims.of_int (36))) + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2418)) + (Prims.of_int (8)) + (Prims.of_int (2418)) + (Prims.of_int (12))) + (Obj.magic + (FStar_Tactics_Derived.focus + (fun uu___4 -> + FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2397)) + (Prims.of_int (23)) + (Prims.of_int (2397)) + (Prims.of_int (39))) + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2398)) + (Prims.of_int (10)) + (Prims.of_int (2416)) + (Prims.of_int (35))) + (Obj.magic + (FStar_Tactics_Logic.implies_intro + ())) + (fun uu___5 -> + (fun p_bind -> + Obj.magic + (FStar_Tactics_Effect.tac_bind (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2398)) @@ -8733,13 +8740,17 @@ let (solve_can_be_split_dep : Prims.op_Negation b then - FStar_Tactics_Derived.fail - "could not unify SMT prop with True" + Obj.magic + (Obj.repr + (FStar_Tactics_Derived.fail + "could not unify SMT prop with True")) else - FStar_Tactics_Effect.lift_div_tac + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac (fun uu___9 -> - ())) + ())))) (fun uu___8 -> (fun @@ -8901,144 +8912,121 @@ let (solve_can_be_split_dep : uu___8)))) uu___6))) uu___5))) - uu___5)))) - (fun uu___4 -> - FStar_Tactics_Effect.lift_div_tac - (fun uu___5 -> - true)))) - else - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___5 -> false)))) - uu___3))) uu___3))) uu___3))) - | uu___ -> - Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail "ill-formed can_be_split_dep"))) - uu___ + uu___5)))) + (fun uu___4 -> + FStar_Tactics_Effect.lift_div_tac + (fun uu___5 -> true)))) + else + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___5 -> false)))) + uu___3))) uu___3))) uu___3) + | uu___ -> FStar_Tactics_Derived.fail "ill-formed can_be_split_dep" let (solve_can_be_split_forall : FStar_Reflection_Data.argv Prims.list -> (Prims.bool, unit) FStar_Tactics_Effect.tac_repr) = - fun uu___ -> - (fun args -> - match args with - | uu___::(t1, uu___1)::(t2, uu___2)::[] -> - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2432)) (Prims.of_int (17)) - (Prims.of_int (2432)) (Prims.of_int (36))) - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2433)) (Prims.of_int (6)) - (Prims.of_int (2458)) (Prims.of_int (18))) - (Obj.magic (slterm_nbr_uvars t1)) - (fun uu___3 -> - (fun lnbr -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2433)) (Prims.of_int (17)) - (Prims.of_int (2433)) (Prims.of_int (36))) - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2434)) (Prims.of_int (6)) - (Prims.of_int (2458)) (Prims.of_int (18))) - (Obj.magic (slterm_nbr_uvars t2)) - (fun uu___3 -> - (fun rnbr -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2435)) - (Prims.of_int (8)) - (Prims.of_int (2437)) - (Prims.of_int (18))) - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2434)) - (Prims.of_int (6)) - (Prims.of_int (2458)) - (Prims.of_int (18))) - (if (lnbr + rnbr) <= Prims.int_one - then - Obj.magic - (Obj.repr (unfold_guard ())) - else - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___4 -> false)))) - (fun uu___3 -> - (fun uu___3 -> - if uu___3 - then - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2440)) - (Prims.of_int (8)) - (Prims.of_int (2456)) - (Prims.of_int (46))) - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2457)) - (Prims.of_int (8)) - (Prims.of_int (2457)) - (Prims.of_int (12))) - (Obj.magic - (FStar_Tactics_Derived.focus - (fun uu___4 -> - FStar_Tactics_Effect.tac_bind - ( - Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2441)) - (Prims.of_int (10)) - (Prims.of_int (2441)) - (Prims.of_int (33))) - ( - Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2442)) - (Prims.of_int (10)) - (Prims.of_int (2456)) - (Prims.of_int (45))) - ( - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range + fun args -> + match args with + | uu___::(t1, uu___1)::(t2, uu___2)::[] -> + FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2432)) + (Prims.of_int (17)) (Prims.of_int (2432)) (Prims.of_int (36))) + (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2433)) + (Prims.of_int (6)) (Prims.of_int (2458)) (Prims.of_int (18))) + (Obj.magic (slterm_nbr_uvars t1)) + (fun uu___3 -> + (fun lnbr -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2433)) (Prims.of_int (17)) + (Prims.of_int (2433)) (Prims.of_int (36))) + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2434)) (Prims.of_int (6)) + (Prims.of_int (2458)) (Prims.of_int (18))) + (Obj.magic (slterm_nbr_uvars t2)) + (fun uu___3 -> + (fun rnbr -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2435)) (Prims.of_int (8)) + (Prims.of_int (2437)) (Prims.of_int (18))) + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2434)) (Prims.of_int (6)) + (Prims.of_int (2458)) (Prims.of_int (18))) + (if (lnbr + rnbr) <= Prims.int_one + then Obj.magic (Obj.repr (unfold_guard ())) + else + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___4 -> false)))) + (fun uu___3 -> + (fun uu___3 -> + if uu___3 + then + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2440)) + (Prims.of_int (8)) + (Prims.of_int (2456)) + (Prims.of_int (46))) + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2457)) + (Prims.of_int (8)) + (Prims.of_int (2457)) + (Prims.of_int (12))) + (Obj.magic + (FStar_Tactics_Derived.focus + (fun uu___4 -> + FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2441)) + (Prims.of_int (10)) + (Prims.of_int (2441)) + (Prims.of_int (33))) + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2442)) + (Prims.of_int (10)) + (Prims.of_int (2456)) + (Prims.of_int (45))) + (Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2441)) (Prims.of_int (17)) (Prims.of_int (2441)) (Prims.of_int (33))) - (Prims.mk_range + (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2441)) (Prims.of_int (10)) (Prims.of_int (2441)) (Prims.of_int (33))) - (Obj.magic - (FStar_Tactics_Logic.forall_intro + (Obj.magic + ( + FStar_Tactics_Logic.forall_intro ())) - (fun - uu___5 -> + (fun uu___5 + -> FStar_Tactics_Effect.lift_div_tac (fun uu___6 -> ())))) - ( - fun - uu___5 -> - (fun - uu___5 -> - Obj.magic - (FStar_Tactics_Effect.tac_bind + (fun uu___5 -> + (fun uu___5 -> + Obj.magic + (FStar_Tactics_Effect.tac_bind (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2442)) @@ -9237,23 +9225,19 @@ let (solve_can_be_split_forall : uu___9)))) uu___7))) uu___6))) - uu___5)))) - (fun uu___4 -> - FStar_Tactics_Effect.lift_div_tac - (fun uu___5 -> - true)))) - else - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___5 -> false)))) - uu___3))) uu___3))) uu___3))) - | uu___ -> - Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "Ill-formed can_be_split_forall, should not happen"))) - uu___ + uu___5)))) + (fun uu___4 -> + FStar_Tactics_Effect.lift_div_tac + (fun uu___5 -> true)))) + else + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___5 -> false)))) + uu___3))) uu___3))) uu___3) + | uu___ -> + FStar_Tactics_Derived.fail + "Ill-formed can_be_split_forall, should not happen" let (extract_cbs_forall_dep_contexts : FStar_Reflection_Types.term -> ((unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) @@ -9314,9 +9298,15 @@ let (open_existentials_forall_dep : "solve_can_be_split_forall_dep_lookup"]))) e) then - FStar_Tactics_Derived.fail - "Tactic disabled: no available lemmas in context" - else FStar_Tactics_Effect.lift_div_tac (fun uu___2 -> ())) + Obj.magic + (Obj.repr + (FStar_Tactics_Derived.fail + "Tactic disabled: no available lemmas in context")) + else + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___2 -> ())))) (fun uu___1 -> (fun uu___1 -> Obj.magic @@ -9364,44 +9354,40 @@ let (open_existentials_forall_dep : FStar_Reflection_Data.Q_Explicit)::[]) -> Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2491)) - (Prims.of_int (17)) - (Prims.of_int (2491)) - (Prims.of_int (31))) - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2491)) - (Prims.of_int (4)) - (Prims.of_int (2513)) - (Prims.of_int (78))) - (FStar_Tactics_Effect.lift_div_tac - (fun uu___4 -> - FStar_Reflection_Derived.collect_app - t1)) + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2491)) + (Prims.of_int (17)) + (Prims.of_int (2491)) + (Prims.of_int (31))) + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2491)) + (Prims.of_int (4)) + (Prims.of_int (2513)) + (Prims.of_int (78))) + (FStar_Tactics_Effect.lift_div_tac (fun uu___4 -> - (fun uu___4 -> - match uu___4 - with - | (hd, tl) -> - if - FStar_Reflection_Derived.is_fvar + FStar_Reflection_Derived.collect_app + t1)) + (fun uu___4 -> + (fun uu___4 -> + match uu___4 with + | (hd, tl) -> + if + FStar_Reflection_Derived.is_fvar hd "Steel.Effect.Common.can_be_split_forall_dep" - then - Obj.magic - (Obj.repr - (match tl - with - | - uu___5::uu___6:: + then + (match tl + with + | + uu___5::uu___6:: (rhs, FStar_Reflection_Data.Q_Explicit)::[] -> - Obj.repr + Obj.magic (FStar_Tactics_Effect.tac_bind (Prims.mk_range "Steel.Effect.Common.fsti" @@ -9429,7 +9415,6 @@ let (open_existentials_forall_dep : (uu___8, body) -> Obj.magic - (Obj.repr (FStar_Tactics_Effect.tac_bind (Prims.mk_range "Steel.Effect.Common.fsti" @@ -9456,14 +9441,12 @@ let (open_existentials_forall_dep : FStar_Pervasives_Native.None -> Obj.magic - (Obj.repr (FStar_Tactics_Derived.fail - "open_existentials_forall_dep: no candidate")) + "open_existentials_forall_dep: no candidate") | FStar_Pervasives_Native.Some f -> Obj.magic - (Obj.repr (FStar_Tactics_Effect.tac_bind (Prims.mk_range "Steel.Effect.Common.fsti" @@ -9649,22 +9632,21 @@ let (open_existentials_forall_dep : uu___13))) uu___12))) uu___11))) - uu___10)))) - uu___9))) + uu___10))) + uu___9)) | uu___8 -> Obj.magic - (Obj.repr (FStar_Tactics_Derived.fail - "open_existentials_forall_dep : not an abstraction"))) + "open_existentials_forall_dep : not an abstraction")) uu___7)) - | - (uu___5, + | + (uu___5, FStar_Reflection_Data.Q_Implicit)::uu___6::uu___7:: (rhs, FStar_Reflection_Data.Q_Explicit)::[] -> - Obj.repr + Obj.magic (FStar_Tactics_Effect.tac_bind (Prims.mk_range "Steel.Effect.Common.fsti" @@ -9692,7 +9674,6 @@ let (open_existentials_forall_dep : (uu___9, body) -> Obj.magic - (Obj.repr (FStar_Tactics_Effect.tac_bind (Prims.mk_range "Steel.Effect.Common.fsti" @@ -9721,14 +9702,12 @@ let (open_existentials_forall_dep : FStar_Pervasives_Native.None -> Obj.magic - (Obj.repr (FStar_Tactics_Derived.fail - "open_existentials_forall_dep: no candidate")) + "open_existentials_forall_dep: no candidate") | FStar_Pervasives_Native.Some f -> Obj.magic - (Obj.repr (FStar_Tactics_Effect.tac_bind (Prims.mk_range "Steel.Effect.Common.fsti" @@ -9914,31 +9893,29 @@ let (open_existentials_forall_dep : uu___14))) uu___13))) uu___12))) - uu___11)))) - uu___10))) + uu___11))) + uu___10)) | uu___9 -> Obj.magic - (Obj.repr (FStar_Tactics_Derived.fail - "open_existentials_forall_dep : not an abstraction"))) + "open_existentials_forall_dep : not an abstraction")) uu___8)) - | - uu___5 -> - Obj.repr - (FStar_Tactics_Derived.fail - "open_existentials_forall_dep : wrong number of arguments to can_be_split_forall_dep"))) - else + | + uu___5 -> Obj.magic - (Obj.repr (FStar_Tactics_Derived.fail - "open_existentials_forall_dep : not a can_be_split_forall_dep goal"))) - uu___4))) + "open_existentials_forall_dep : wrong number of arguments to can_be_split_forall_dep")) + else + Obj.magic + ( + FStar_Tactics_Derived.fail + "open_existentials_forall_dep : not a can_be_split_forall_dep goal")) + uu___4)) | uu___3 -> Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "open_existentials_forall_dep : not a squash/auto_squash goal"))) + (FStar_Tactics_Derived.fail + "open_existentials_forall_dep : not a squash/auto_squash goal")) uu___3))) uu___2))) uu___1))) uu___1) let (try_open_existentials_forall_dep : @@ -9969,100 +9946,90 @@ let rec (solve_can_be_split_forall_dep : FStar_Reflection_Data.argv Prims.list -> (Prims.bool, unit) FStar_Tactics_Effect.tac_repr) = - fun uu___ -> - (fun args -> - match args with - | uu___::(pr, uu___1)::(t1, uu___2)::(t2, uu___3)::[] -> - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2530)) (Prims.of_int (17)) - (Prims.of_int (2530)) (Prims.of_int (36))) - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2531)) (Prims.of_int (6)) - (Prims.of_int (2575)) (Prims.of_int (18))) - (Obj.magic (slterm_nbr_uvars t1)) - (fun uu___4 -> - (fun lnbr -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2531)) (Prims.of_int (17)) - (Prims.of_int (2531)) (Prims.of_int (36))) - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2532)) (Prims.of_int (6)) - (Prims.of_int (2575)) (Prims.of_int (18))) - (Obj.magic (slterm_nbr_uvars t2)) - (fun uu___4 -> - (fun rnbr -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2533)) - (Prims.of_int (8)) - (Prims.of_int (2535)) - (Prims.of_int (18))) - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2532)) - (Prims.of_int (6)) - (Prims.of_int (2575)) - (Prims.of_int (18))) - (if (lnbr + rnbr) <= Prims.int_one - then - Obj.magic - (Obj.repr (unfold_guard ())) - else - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___5 -> false)))) - (fun uu___4 -> - (fun uu___4 -> - if uu___4 - then - Obj.magic - (Obj.repr - (FStar_Tactics_Derived.try_with - (fun uu___5 -> - match () with - | () -> - FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2539)) - (Prims.of_int (9)) - (Prims.of_int (2563)) - (Prims.of_int (37))) - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2565)) - (Prims.of_int (9)) - (Prims.of_int (2565)) - (Prims.of_int (13))) - (Obj.magic - (FStar_Tactics_Derived.focus - (fun - uu___6 -> - FStar_Tactics_Effect.tac_bind - (Prims.mk_range + fun args -> + match args with + | uu___::(pr, uu___1)::(t1, uu___2)::(t2, uu___3)::[] -> + FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2530)) + (Prims.of_int (17)) (Prims.of_int (2530)) (Prims.of_int (36))) + (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2531)) + (Prims.of_int (6)) (Prims.of_int (2575)) (Prims.of_int (18))) + (Obj.magic (slterm_nbr_uvars t1)) + (fun uu___4 -> + (fun lnbr -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2531)) (Prims.of_int (17)) + (Prims.of_int (2531)) (Prims.of_int (36))) + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2532)) (Prims.of_int (6)) + (Prims.of_int (2575)) (Prims.of_int (18))) + (Obj.magic (slterm_nbr_uvars t2)) + (fun uu___4 -> + (fun rnbr -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2533)) (Prims.of_int (8)) + (Prims.of_int (2535)) (Prims.of_int (18))) + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2532)) (Prims.of_int (6)) + (Prims.of_int (2575)) (Prims.of_int (18))) + (if (lnbr + rnbr) <= Prims.int_one + then Obj.magic (Obj.repr (unfold_guard ())) + else + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___5 -> false)))) + (fun uu___4 -> + (fun uu___4 -> + if uu___4 + then + Obj.magic + (Obj.repr + (FStar_Tactics_Derived.try_with + (fun uu___5 -> + match () with + | () -> + FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2539)) + (Prims.of_int (9)) + (Prims.of_int (2563)) + (Prims.of_int (37))) + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2565)) + (Prims.of_int (9)) + (Prims.of_int (2565)) + (Prims.of_int (13))) + (Obj.magic + (FStar_Tactics_Derived.focus + (fun uu___6 -> + FStar_Tactics_Effect.tac_bind + ( + Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2540)) (Prims.of_int (10)) (Prims.of_int (2540)) (Prims.of_int (17))) - (Prims.mk_range + ( + Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2541)) (Prims.of_int (10)) (Prims.of_int (2563)) (Prims.of_int (36))) - (Obj.magic + ( + Obj.magic (FStar_Tactics_Builtins.norm [])) - (fun + ( + fun uu___7 -> (fun uu___7 -> @@ -10323,13 +10290,17 @@ let rec (solve_can_be_split_forall_dep : Prims.op_Negation b then - FStar_Tactics_Derived.fail - "could not unify SMT prop with True" + Obj.magic + (Obj.repr + (FStar_Tactics_Derived.fail + "could not unify SMT prop with True")) else - FStar_Tactics_Effect.lift_div_tac + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac (fun uu___12 - -> ())) + -> ())))) (fun uu___11 -> @@ -10506,186 +10477,158 @@ let rec (solve_can_be_split_forall_dep : uu___8))) uu___8))) uu___7)))) + (fun uu___6 -> + FStar_Tactics_Effect.lift_div_tac + (fun uu___7 -> + true))) + (fun uu___5 -> + (fun uu___5 -> + match uu___5 with + | Postpone msg -> + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac (fun uu___6 - -> - FStar_Tactics_Effect.lift_div_tac - (fun - uu___7 -> - true))) - (fun uu___5 -> - (fun uu___5 -> - match uu___5 - with - | Postpone msg - -> - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun - uu___6 -> - false))) - | FStar_Tactics_Common.TacticFailure - msg -> - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range + -> false))) + | FStar_Tactics_Common.TacticFailure + msg -> + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2570)) (Prims.of_int (22)) (Prims.of_int (2570)) (Prims.of_int (57))) - (Prims.mk_range + (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2571)) (Prims.of_int (9)) (Prims.of_int (2573)) (Prims.of_int (22))) - (Obj.magic + (Obj.magic (try_open_existentials_forall_dep ())) - (fun - uu___6 -> + (fun uu___6 + -> (fun opened -> if opened then Obj.magic - (Obj.repr (solve_can_be_split_forall_dep - args)) + args) else Obj.magic - (Obj.repr (FStar_Tactics_Derived.fail - msg))) + msg)) uu___6))) - | uu___6 -> - Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "Unexpected exception in framing tactic"))) - uu___5))) - else - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___6 -> false)))) - uu___4))) uu___4))) uu___4))) - | uu___ -> - Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "Ill-formed can_be_split_forall_dep, should not happen"))) - uu___ + | uu___6 -> + Obj.magic + (Obj.repr + (FStar_Tactics_Derived.fail + "Unexpected exception in framing tactic"))) + uu___5))) + else + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___6 -> false)))) + uu___4))) uu___4))) uu___4) + | uu___ -> + FStar_Tactics_Derived.fail + "Ill-formed can_be_split_forall_dep, should not happen" let (solve_equiv_forall : FStar_Reflection_Data.argv Prims.list -> (Prims.bool, unit) FStar_Tactics_Effect.tac_repr) = - fun uu___ -> - (fun args -> - match args with - | uu___::(t1, uu___1)::(t2, uu___2)::[] -> - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2583)) (Prims.of_int (17)) - (Prims.of_int (2583)) (Prims.of_int (36))) - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2584)) (Prims.of_int (6)) - (Prims.of_int (2612)) (Prims.of_int (18))) - (Obj.magic (slterm_nbr_uvars t1)) - (fun uu___3 -> - (fun lnbr -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2584)) (Prims.of_int (17)) - (Prims.of_int (2584)) (Prims.of_int (36))) - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2585)) (Prims.of_int (6)) - (Prims.of_int (2612)) (Prims.of_int (18))) - (Obj.magic (slterm_nbr_uvars t2)) - (fun uu___3 -> - (fun rnbr -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2586)) - (Prims.of_int (8)) - (Prims.of_int (2588)) - (Prims.of_int (18))) - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2585)) - (Prims.of_int (6)) - (Prims.of_int (2612)) - (Prims.of_int (18))) - (if (lnbr + rnbr) <= Prims.int_one - then - Obj.magic - (Obj.repr (unfold_guard ())) - else - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___4 -> false)))) - (fun uu___3 -> - (fun uu___3 -> - if uu___3 - then - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2591)) - (Prims.of_int (8)) - (Prims.of_int (2610)) - (Prims.of_int (62))) - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2611)) - (Prims.of_int (8)) - (Prims.of_int (2611)) - (Prims.of_int (12))) - (Obj.magic - (FStar_Tactics_Derived.focus - (fun uu___4 -> - FStar_Tactics_Effect.tac_bind - ( - Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2591)) - (Prims.of_int (24)) - (Prims.of_int (2591)) - (Prims.of_int (56))) - ( - Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2592)) - (Prims.of_int (22)) - (Prims.of_int (2610)) - (Prims.of_int (61))) + fun args -> + match args with + | uu___::(t1, uu___1)::(t2, uu___2)::[] -> + FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2583)) + (Prims.of_int (17)) (Prims.of_int (2583)) (Prims.of_int (36))) + (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2584)) + (Prims.of_int (6)) (Prims.of_int (2612)) (Prims.of_int (18))) + (Obj.magic (slterm_nbr_uvars t1)) + (fun uu___3 -> + (fun lnbr -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2584)) (Prims.of_int (17)) + (Prims.of_int (2584)) (Prims.of_int (36))) + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2585)) (Prims.of_int (6)) + (Prims.of_int (2612)) (Prims.of_int (18))) + (Obj.magic (slterm_nbr_uvars t2)) + (fun uu___3 -> + (fun rnbr -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2586)) (Prims.of_int (8)) + (Prims.of_int (2588)) (Prims.of_int (18))) + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2585)) (Prims.of_int (6)) + (Prims.of_int (2612)) (Prims.of_int (18))) + (if (lnbr + rnbr) <= Prims.int_one + then Obj.magic (Obj.repr (unfold_guard ())) + else + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___4 -> false)))) + (fun uu___3 -> + (fun uu___3 -> + if uu___3 + then + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2591)) + (Prims.of_int (8)) + (Prims.of_int (2610)) + (Prims.of_int (62))) + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2611)) + (Prims.of_int (8)) + (Prims.of_int (2611)) + (Prims.of_int (12))) + (Obj.magic + (FStar_Tactics_Derived.focus + (fun uu___4 -> + FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2591)) + (Prims.of_int (24)) + (Prims.of_int (2591)) + (Prims.of_int (56))) + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2592)) + (Prims.of_int (22)) + (Prims.of_int (2610)) + (Prims.of_int (61))) + (Obj.magic + (FStar_Tactics_Derived.apply_lemma + (FStar_Reflection_Builtins.pack_ln ( - Obj.magic - (FStar_Tactics_Derived.apply_lemma - (FStar_Reflection_Builtins.pack_ln - (FStar_Reflection_Data.Tv_FVar + FStar_Reflection_Data.Tv_FVar (FStar_Reflection_Builtins.pack_fv ["Steel"; "Effect"; "Common"; "equiv_forall_elim"]))))) - ( - fun - uu___5 -> - (fun - uu___5 -> - Obj.magic - (FStar_Tactics_Effect.tac_bind + (fun uu___5 -> + (fun uu___5 -> + Obj.magic + (FStar_Tactics_Effect.tac_bind (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2592)) @@ -10937,134 +10880,114 @@ let (solve_equiv_forall : uu___9))) uu___8)))) uu___6))) - uu___5)))) - (fun uu___4 -> - FStar_Tactics_Effect.lift_div_tac - (fun uu___5 -> - true)))) - else - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___5 -> false)))) - uu___3))) uu___3))) uu___3))) - | uu___ -> - Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "Ill-formed equiv_forall, should not happen"))) uu___ + uu___5)))) + (fun uu___4 -> + FStar_Tactics_Effect.lift_div_tac + (fun uu___5 -> true)))) + else + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___5 -> false)))) + uu___3))) uu___3))) uu___3) + | uu___ -> + FStar_Tactics_Derived.fail + "Ill-formed equiv_forall, should not happen" let (solve_equiv : FStar_Reflection_Data.argv Prims.list -> (Prims.bool, unit) FStar_Tactics_Effect.tac_repr) = - fun uu___ -> - (fun args -> - match args with - | (t1, uu___)::(t2, uu___1)::[] -> - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2620)) (Prims.of_int (17)) - (Prims.of_int (2620)) (Prims.of_int (36))) - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2621)) (Prims.of_int (6)) - (Prims.of_int (2645)) (Prims.of_int (18))) - (Obj.magic (slterm_nbr_uvars t1)) - (fun uu___2 -> - (fun lnbr -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2621)) (Prims.of_int (17)) - (Prims.of_int (2621)) (Prims.of_int (36))) - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2622)) (Prims.of_int (6)) - (Prims.of_int (2645)) (Prims.of_int (18))) - (Obj.magic (slterm_nbr_uvars t2)) - (fun uu___2 -> - (fun rnbr -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2623)) - (Prims.of_int (8)) - (Prims.of_int (2625)) - (Prims.of_int (18))) - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2622)) - (Prims.of_int (6)) - (Prims.of_int (2645)) - (Prims.of_int (18))) - (if (lnbr + rnbr) <= Prims.int_one - then - Obj.magic - (Obj.repr (unfold_guard ())) - else - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___3 -> false)))) - (fun uu___2 -> - (fun uu___2 -> - if uu___2 - then - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2628)) - (Prims.of_int (8)) - (Prims.of_int (2642)) - (Prims.of_int (48))) - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2643)) - (Prims.of_int (8)) - (Prims.of_int (2643)) - (Prims.of_int (12))) - (Obj.magic - (FStar_Tactics_Derived.focus - (fun uu___3 -> - FStar_Tactics_Derived.or_else - ( - fun - uu___4 -> - FStar_Tactics_Derived.apply_lemma - (FStar_Reflection_Builtins.pack_ln - (FStar_Reflection_Data.Tv_FVar + fun args -> + match args with + | (t1, uu___)::(t2, uu___1)::[] -> + FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2620)) + (Prims.of_int (17)) (Prims.of_int (2620)) (Prims.of_int (36))) + (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2621)) + (Prims.of_int (6)) (Prims.of_int (2645)) (Prims.of_int (18))) + (Obj.magic (slterm_nbr_uvars t1)) + (fun uu___2 -> + (fun lnbr -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2621)) (Prims.of_int (17)) + (Prims.of_int (2621)) (Prims.of_int (36))) + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2622)) (Prims.of_int (6)) + (Prims.of_int (2645)) (Prims.of_int (18))) + (Obj.magic (slterm_nbr_uvars t2)) + (fun uu___2 -> + (fun rnbr -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2623)) (Prims.of_int (8)) + (Prims.of_int (2625)) (Prims.of_int (18))) + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2622)) (Prims.of_int (6)) + (Prims.of_int (2645)) (Prims.of_int (18))) + (if (lnbr + rnbr) <= Prims.int_one + then Obj.magic (Obj.repr (unfold_guard ())) + else + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___3 -> false)))) + (fun uu___2 -> + (fun uu___2 -> + if uu___2 + then + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2628)) + (Prims.of_int (8)) + (Prims.of_int (2642)) + (Prims.of_int (48))) + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2643)) + (Prims.of_int (8)) + (Prims.of_int (2643)) + (Prims.of_int (12))) + (Obj.magic + (FStar_Tactics_Derived.focus + (fun uu___3 -> + FStar_Tactics_Derived.or_else + (fun uu___4 -> + FStar_Tactics_Derived.apply_lemma + (FStar_Reflection_Builtins.pack_ln + (FStar_Reflection_Data.Tv_FVar (FStar_Reflection_Builtins.pack_fv ["Steel"; "Effect"; "Common"; "equiv_refl"])))) - ( - fun - uu___4 -> - FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2632)) - (Prims.of_int (14)) - (Prims.of_int (2632)) - (Prims.of_int (68))) - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2633)) - (Prims.of_int (14)) - (Prims.of_int (2642)) - (Prims.of_int (46))) - (if - (lnbr <> + (fun uu___4 -> + FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2632)) + (Prims.of_int (14)) + (Prims.of_int (2632)) + (Prims.of_int (68))) + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2633)) + (Prims.of_int (14)) + (Prims.of_int (2642)) + (Prims.of_int (46))) + (if + (lnbr <> Prims.int_zero) && (rnbr = Prims.int_zero) - then - Obj.magic + then + Obj.magic (Obj.repr (FStar_Tactics_Derived.apply_lemma (FStar_Reflection_Builtins.pack_ln @@ -11074,16 +10997,16 @@ let (solve_equiv : "Effect"; "Common"; "equiv_sym"]))))) - else - Obj.magic + else + Obj.magic (Obj.repr (FStar_Tactics_Effect.lift_div_tac (fun uu___6 -> ())))) - (fun - uu___5 -> - (fun + (fun uu___5 + -> + (fun uu___5 -> Obj.magic (FStar_Tactics_Effect.tac_bind @@ -11146,150 +11069,127 @@ let (solve_equiv : "FStar.Pervasives.Native.snd"]; FStar_Pervasives.delta_attr ["Steel.Effect.Common.__reduce__"]; - FStar_Pervasives.primops; - FStar_Pervasives.iota; - FStar_Pervasives.zeta])) - (fun - uu___7 -> - (fun - uu___7 -> - Obj.magic - (canon' - false - (FStar_Reflection_Builtins.pack_ln - (FStar_Reflection_Data.Tv_FVar - (FStar_Reflection_Builtins.pack_fv - ["Steel"; - "Effect"; - "Common"; - "true_p"]))) - (FStar_Reflection_Builtins.pack_ln - (FStar_Reflection_Data.Tv_FVar - (FStar_Reflection_Builtins.pack_fv - ["Steel"; - "Effect"; - "Common"; - "true_p"]))))) - uu___7))) - uu___6))) - uu___5))))) - (fun uu___3 -> - FStar_Tactics_Effect.lift_div_tac - (fun uu___4 -> - true)))) - else - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___4 -> false)))) - uu___2))) uu___2))) uu___2))) - | uu___ -> - Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "Ill-formed equiv, should not happen"))) uu___ -let (solve_can_be_split_post : - FStar_Reflection_Data.argv Prims.list -> - (Prims.bool, unit) FStar_Tactics_Effect.tac_repr) - = - fun uu___ -> - (fun args -> - match args with - | uu___::uu___1::(t1, uu___2)::(t2, uu___3)::[] -> - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2653)) (Prims.of_int (17)) - (Prims.of_int (2653)) (Prims.of_int (36))) - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2654)) (Prims.of_int (6)) - (Prims.of_int (2685)) (Prims.of_int (18))) - (Obj.magic (slterm_nbr_uvars t1)) - (fun uu___4 -> - (fun lnbr -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2654)) (Prims.of_int (17)) - (Prims.of_int (2654)) (Prims.of_int (36))) - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2655)) (Prims.of_int (6)) - (Prims.of_int (2685)) (Prims.of_int (18))) - (Obj.magic (slterm_nbr_uvars t2)) - (fun uu___4 -> - (fun rnbr -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2656)) - (Prims.of_int (8)) - (Prims.of_int (2658)) - (Prims.of_int (18))) - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2655)) - (Prims.of_int (6)) - (Prims.of_int (2685)) - (Prims.of_int (18))) - (if (lnbr + rnbr) <= Prims.int_one - then - Obj.magic - (Obj.repr (unfold_guard ())) - else - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___5 -> false)))) - (fun uu___4 -> - (fun uu___4 -> - if uu___4 - then - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2661)) - (Prims.of_int (8)) - (Prims.of_int (2683)) - (Prims.of_int (62))) - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2684)) - (Prims.of_int (8)) - (Prims.of_int (2684)) - (Prims.of_int (12))) - (Obj.magic - (FStar_Tactics_Derived.focus - (fun uu___5 -> - FStar_Tactics_Effect.tac_bind - ( - Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2661)) - (Prims.of_int (24)) - (Prims.of_int (2661)) - (Prims.of_int (30))) - ( - Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2662)) - (Prims.of_int (22)) - (Prims.of_int (2683)) - (Prims.of_int (61))) - ( - Obj.magic - (FStar_Tactics_Builtins.norm - [])) - ( - fun - uu___6 -> + FStar_Pervasives.primops; + FStar_Pervasives.iota; + FStar_Pervasives.zeta])) (fun - uu___6 -> + uu___7 -> + (fun + uu___7 -> Obj.magic - (FStar_Tactics_Effect.tac_bind + (canon' + false + (FStar_Reflection_Builtins.pack_ln + (FStar_Reflection_Data.Tv_FVar + (FStar_Reflection_Builtins.pack_fv + ["Steel"; + "Effect"; + "Common"; + "true_p"]))) + (FStar_Reflection_Builtins.pack_ln + (FStar_Reflection_Data.Tv_FVar + (FStar_Reflection_Builtins.pack_fv + ["Steel"; + "Effect"; + "Common"; + "true_p"]))))) + uu___7))) + uu___6))) + uu___5))))) + (fun uu___3 -> + FStar_Tactics_Effect.lift_div_tac + (fun uu___4 -> true)))) + else + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___4 -> false)))) + uu___2))) uu___2))) uu___2) + | uu___ -> + FStar_Tactics_Derived.fail "Ill-formed equiv, should not happen" +let (solve_can_be_split_post : + FStar_Reflection_Data.argv Prims.list -> + (Prims.bool, unit) FStar_Tactics_Effect.tac_repr) + = + fun args -> + match args with + | uu___::uu___1::(t1, uu___2)::(t2, uu___3)::[] -> + FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2653)) + (Prims.of_int (17)) (Prims.of_int (2653)) (Prims.of_int (36))) + (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2654)) + (Prims.of_int (6)) (Prims.of_int (2685)) (Prims.of_int (18))) + (Obj.magic (slterm_nbr_uvars t1)) + (fun uu___4 -> + (fun lnbr -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2654)) (Prims.of_int (17)) + (Prims.of_int (2654)) (Prims.of_int (36))) + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2655)) (Prims.of_int (6)) + (Prims.of_int (2685)) (Prims.of_int (18))) + (Obj.magic (slterm_nbr_uvars t2)) + (fun uu___4 -> + (fun rnbr -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2656)) (Prims.of_int (8)) + (Prims.of_int (2658)) (Prims.of_int (18))) + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2655)) (Prims.of_int (6)) + (Prims.of_int (2685)) (Prims.of_int (18))) + (if (lnbr + rnbr) <= Prims.int_one + then Obj.magic (Obj.repr (unfold_guard ())) + else + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___5 -> false)))) + (fun uu___4 -> + (fun uu___4 -> + if uu___4 + then + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2661)) + (Prims.of_int (8)) + (Prims.of_int (2683)) + (Prims.of_int (62))) + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2684)) + (Prims.of_int (8)) + (Prims.of_int (2684)) + (Prims.of_int (12))) + (Obj.magic + (FStar_Tactics_Derived.focus + (fun uu___5 -> + FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2661)) + (Prims.of_int (24)) + (Prims.of_int (2661)) + (Prims.of_int (30))) + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2662)) + (Prims.of_int (22)) + (Prims.of_int (2683)) + (Prims.of_int (61))) + (Obj.magic + (FStar_Tactics_Builtins.norm + [])) + (fun uu___6 -> + (fun uu___6 -> + Obj.magic + (FStar_Tactics_Effect.tac_bind (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2662)) @@ -11638,22 +11538,17 @@ let (solve_can_be_split_post : uu___8))) uu___7))) uu___7))) - uu___6)))) - (fun uu___5 -> - FStar_Tactics_Effect.lift_div_tac - (fun uu___6 -> - true)))) - else - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___6 -> false)))) - uu___4))) uu___4))) uu___4))) - | uu___ -> - Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail "ill-formed can_be_split_post"))) - uu___ + uu___6)))) + (fun uu___5 -> + FStar_Tactics_Effect.lift_div_tac + (fun uu___6 -> true)))) + else + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___6 -> false)))) + uu___4))) uu___4))) uu___4) + | uu___ -> FStar_Tactics_Derived.fail "ill-formed can_be_split_post" let (is_return_eq : FStar_Reflection_Types.term -> FStar_Reflection_Types.term -> @@ -12029,118 +11924,116 @@ let (goal_to_equiv : match f with | FStar_Reflection_Formula.App (hd0, t1) -> Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2746)) - (Prims.of_int (6)) - (Prims.of_int (2747)) - (Prims.of_int (70))) - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2748)) - (Prims.of_int (6)) - (Prims.of_int (2768)) - (Prims.of_int (51))) - (if - Prims.op_Negation - (FStar_Reflection_Derived.is_fvar hd0 - "Prims.squash") - then - FStar_Tactics_Derived.fail - (Prims.strcat loc - " unexpected non-squash goal in goal_to_equiv") - else - FStar_Tactics_Effect.lift_div_tac - (fun uu___1 -> ())) + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2746)) (Prims.of_int (6)) + (Prims.of_int (2747)) (Prims.of_int (70))) + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2748)) (Prims.of_int (6)) + (Prims.of_int (2768)) (Prims.of_int (51))) + (if + Prims.op_Negation + (FStar_Reflection_Derived.is_fvar hd0 + "Prims.squash") + then + Obj.magic + (Obj.repr + (FStar_Tactics_Derived.fail + (Prims.strcat loc + " unexpected non-squash goal in goal_to_equiv"))) + else + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___1 -> ())))) + (fun uu___ -> (fun uu___ -> - (fun uu___ -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2748)) - (Prims.of_int (21)) - (Prims.of_int (2748)) - (Prims.of_int (34))) - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2748)) - (Prims.of_int (6)) - (Prims.of_int (2768)) - (Prims.of_int (51))) - (FStar_Tactics_Effect.lift_div_tac - (fun uu___1 -> - FStar_Reflection_Derived.collect_app - t1)) + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2748)) + (Prims.of_int (21)) + (Prims.of_int (2748)) + (Prims.of_int (34))) + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2748)) + (Prims.of_int (6)) + (Prims.of_int (2768)) + (Prims.of_int (51))) + (FStar_Tactics_Effect.lift_div_tac (fun uu___1 -> - (fun uu___1 -> - match uu___1 with - | (hd, args) -> - if - FStar_Reflection_Derived.is_fvar - hd - "Steel.Effect.Common.can_be_split" - then - Obj.magic - (Obj.repr - (FStar_Tactics_Derived.apply_lemma - (FStar_Reflection_Builtins.pack_ln - (FStar_Reflection_Data.Tv_FVar - (FStar_Reflection_Builtins.pack_fv + FStar_Reflection_Derived.collect_app + t1)) + (fun uu___1 -> + (fun uu___1 -> + match uu___1 with + | (hd, args) -> + if + FStar_Reflection_Derived.is_fvar + hd + "Steel.Effect.Common.can_be_split" + then + Obj.magic + (FStar_Tactics_Derived.apply_lemma + (FStar_Reflection_Builtins.pack_ln + (FStar_Reflection_Data.Tv_FVar + (FStar_Reflection_Builtins.pack_fv ["Steel"; "Effect"; "Common"; - "equiv_can_be_split"]))))) - else - Obj.magic - (Obj.repr - (if - FStar_Reflection_Derived.is_fvar - hd - "Steel.Effect.Common.can_be_split_forall" - then - Obj.repr + "equiv_can_be_split"])))) + else + if + FStar_Reflection_Derived.is_fvar + hd + "Steel.Effect.Common.can_be_split_forall" + then + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2752)) + (Prims.of_int (8)) + (Prims.of_int (2752)) + (Prims.of_int (32))) + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2753)) + (Prims.of_int (8)) + (Prims.of_int (2753)) + (Prims.of_int (41))) + (Obj.magic + (FStar_Tactics_Effect.tac_bind ( - FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2752)) - (Prims.of_int (8)) - (Prims.of_int (2752)) - (Prims.of_int (32))) - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2753)) - (Prims.of_int (8)) - (Prims.of_int (2753)) - (Prims.of_int (41))) - (Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range + Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2752)) (Prims.of_int (15)) (Prims.of_int (2752)) (Prims.of_int (32))) - (Prims.mk_range + ( + Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2752)) (Prims.of_int (8)) (Prims.of_int (2752)) (Prims.of_int (32))) - (Obj.magic + ( + Obj.magic (FStar_Tactics_Logic.forall_intro ())) - (fun + ( + fun uu___3 -> FStar_Tactics_Effect.lift_div_tac (fun uu___4 -> ())))) - (fun - uu___3 -> - (fun - uu___3 -> + (fun uu___3 -> + (fun uu___3 + -> Obj.magic (FStar_Tactics_Derived.apply_lemma (FStar_Reflection_Builtins.pack_ln @@ -12150,31 +12043,29 @@ let (goal_to_equiv : "Effect"; "Common"; "equiv_can_be_split"]))))) - uu___3)) - else - Obj.repr - ( - if - FStar_Reflection_Derived.is_fvar - hd - "Steel.Effect.Common.equiv_forall" - then - Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2755)) - (Prims.of_int (8)) - (Prims.of_int (2755)) - (Prims.of_int (40))) - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2756)) - (Prims.of_int (8)) - (Prims.of_int (2756)) - (Prims.of_int (32))) - (Obj.magic - (FStar_Tactics_Derived.apply_lemma + uu___3)) + else + if + FStar_Reflection_Derived.is_fvar + hd + "Steel.Effect.Common.equiv_forall" + then + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2755)) + (Prims.of_int (8)) + (Prims.of_int (2755)) + (Prims.of_int (40))) + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2756)) + (Prims.of_int (8)) + (Prims.of_int (2756)) + (Prims.of_int (32))) + (Obj.magic + (FStar_Tactics_Derived.apply_lemma (FStar_Reflection_Builtins.pack_ln (FStar_Reflection_Data.Tv_FVar (FStar_Reflection_Builtins.pack_fv @@ -12182,9 +12073,9 @@ let (goal_to_equiv : "Effect"; "Common"; "equiv_forall_elim"]))))) - (fun - uu___4 -> - (fun + (fun uu___4 + -> + (fun uu___4 -> Obj.magic (FStar_Tactics_Effect.tac_bind @@ -12210,28 +12101,27 @@ let (goal_to_equiv : uu___6 -> ())))) uu___4)) - else - Obj.repr - (if - FStar_Reflection_Derived.is_fvar - hd - "Steel.Effect.Common.can_be_split_post" - then - Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range + else + if + FStar_Reflection_Derived.is_fvar + hd + "Steel.Effect.Common.can_be_split_post" + then + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2758)) (Prims.of_int (8)) (Prims.of_int (2758)) (Prims.of_int (45))) - (Prims.mk_range + (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2759)) (Prims.of_int (8)) (Prims.of_int (2761)) (Prims.of_int (32))) - (Obj.magic + (Obj.magic (FStar_Tactics_Derived.apply_lemma (FStar_Reflection_Builtins.pack_ln (FStar_Reflection_Data.Tv_FVar @@ -12240,8 +12130,8 @@ let (goal_to_equiv : "Effect"; "Common"; "can_be_split_post_elim"]))))) - (fun - uu___5 -> + (fun uu___5 + -> (fun uu___5 -> Obj.magic @@ -12332,39 +12222,41 @@ let (goal_to_equiv : uu___7))) uu___6))) uu___5)) - else - Obj.repr - (if - FStar_Reflection_Derived.is_fvar - hd - "Steel.Effect.Common.can_be_split_dep" - then - FStar_Tactics_Derived.fail - (Prims.strcat + else + if + FStar_Reflection_Derived.is_fvar + hd + "Steel.Effect.Common.can_be_split_dep" + then + Obj.magic + (FStar_Tactics_Derived.fail + ( + Prims.strcat "can_be_split_dep not supported in " - loc) - else - if - FStar_Reflection_Derived.is_fvar - hd - "Steel.Effect.Common.can_be_split_forall_dep" - then - FStar_Tactics_Derived.fail + loc)) + else + if + FStar_Reflection_Derived.is_fvar + hd + "Steel.Effect.Common.can_be_split_forall_dep" + then + Obj.magic + (FStar_Tactics_Derived.fail (Prims.strcat "can_be_split_forall_dep not supported in " - loc) - else - FStar_Tactics_Derived.fail + loc)) + else + Obj.magic + (FStar_Tactics_Derived.fail (Prims.strcat loc - " goal in unexpected position"))))))) - uu___1))) uu___))) + " goal in unexpected position"))) + uu___1))) uu___)) | uu___ -> Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - (Prims.strcat loc " unexpected goal")))) - uu___))) uu___) + (FStar_Tactics_Derived.fail + (Prims.strcat loc " unexpected goal"))) uu___))) + uu___) let rec term_dict_assoc : 'a . FStar_Reflection_Types.term -> @@ -13419,13 +13311,11 @@ let rec (resolve_tac : if uu___2 then Obj.magic - (Obj.repr - (resolve_tac dict)) + (resolve_tac dict) else Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "Could not make progress, no solvable goal found"))) + (FStar_Tactics_Derived.fail + "Could not make progress, no solvable goal found")) uu___2))) uu___2))) uu___1)))) uu___) let rec (pick_next_logical : @@ -13714,16 +13604,14 @@ let (is_true : match uu___1 with | FStar_Reflection_Formula.True_ -> Obj.magic - (Obj.repr - (FStar_Tactics_Derived.exact - (FStar_Reflection_Builtins.pack_ln - (FStar_Reflection_Data.Tv_Const - FStar_Reflection_Data.C_Unit)))) + (FStar_Tactics_Derived.exact + (FStar_Reflection_Builtins.pack_ln + (FStar_Reflection_Data.Tv_Const + FStar_Reflection_Data.C_Unit))) | uu___2 -> Obj.magic - (Obj.repr - (FStar_Tactics_Effect.raise - FStar_Tactics_Derived.Goal_not_trivial))) uu___1) + (FStar_Tactics_Effect.raise + FStar_Tactics_Derived.Goal_not_trivial)) uu___1) let rec (solve_maybe_emps : Prims.nat -> (unit, unit) FStar_Tactics_Effect.tac_repr) = fun uu___ -> @@ -14641,13 +14529,11 @@ let (ite_soundness_tac : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) match uu___4 with | [] -> Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "should not happen")) + (FStar_Tactics_Derived.fail + "should not happen") | uu___5::tl -> Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind + (FStar_Tactics_Effect.tac_bind (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (3153)) @@ -14865,7 +14751,7 @@ let (ite_soundness_tac : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) uu___9))) uu___8))) uu___7))) - uu___6)))) + uu___6))) uu___4))) uu___4))) uu___3))) uu___2))) uu___1) let (vc_norm : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = diff --git a/ocaml/fstar-lib/generated/Steel_ST_GenElim_Base.ml b/ocaml/fstar-lib/generated/Steel_ST_GenElim_Base.ml index 2ca503ea1dc..ad134920e48 100644 --- a/ocaml/fstar-lib/generated/Steel_ST_GenElim_Base.ml +++ b/ocaml/fstar-lib/generated/Steel_ST_GenElim_Base.ml @@ -240,28 +240,24 @@ let rec (solve_gen_unit_elim : (t2, FStar_Reflection_Data.Q_Explicit)::[] -> - Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.ST.GenElim.Base.fsti" - (Prims.of_int (335)) - (Prims.of_int (20)) - (Prims.of_int (335)) - (Prims.of_int (42))) - (Prims.mk_range - "Steel.ST.GenElim.Base.fsti" - (Prims.of_int (336)) - (Prims.of_int (10)) - (Prims.of_int (337)) - (Prims.of_int (68))) - (Obj.magic - ( - solve_gen_unit_elim + FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.ST.GenElim.Base.fsti" + (Prims.of_int (335)) + (Prims.of_int (20)) + (Prims.of_int (335)) + (Prims.of_int (42))) + (Prims.mk_range + "Steel.ST.GenElim.Base.fsti" + (Prims.of_int (336)) + (Prims.of_int (10)) + (Prims.of_int (337)) + (Prims.of_int (68))) + (Obj.magic + (solve_gen_unit_elim t1)) - (fun uu___5 - -> - (fun t1' - -> + (fun uu___5 -> + (fun t1' -> Obj.magic (FStar_Tactics_Effect.tac_bind (Prims.mk_range @@ -298,11 +294,10 @@ let rec (solve_gen_unit_elim : FStar_Reflection_Data.Q_Explicit); (t2', FStar_Reflection_Data.Q_Explicit)])))) - uu___5)) + uu___5) | uu___5 -> - Obj.repr - (FStar_Tactics_Derived.fail - "ill-formed star"))) + FStar_Tactics_Derived.fail + "ill-formed star")) else Obj.magic (Obj.repr @@ -434,19 +429,25 @@ let rec (solve_gen_elim : (body, FStar_Reflection_Data.Q_Explicit)::[] -> - FStar_Tactics_Effect.lift_div_tac - (fun uu___3 -> - ([(ty, - FStar_Reflection_Data.Q_Implicit)], - body)) + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___3 -> + ([(ty, + FStar_Reflection_Data.Q_Implicit)], + body)))) | (body, FStar_Reflection_Data.Q_Explicit)::[] -> - FStar_Tactics_Effect.lift_div_tac - (fun uu___3 -> ([], body)) + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___3 -> ([], body)))) | uu___3 -> - FStar_Tactics_Derived.fail - "ill-formed exists_") + Obj.magic + (Obj.repr + (FStar_Tactics_Derived.fail + "ill-formed exists_"))) (fun uu___3 -> (fun uu___3 -> match uu___3 with @@ -804,37 +805,35 @@ let rec (solve_gen_elim : (tr, FStar_Reflection_Data.Q_Explicit)::[] -> - Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.ST.GenElim.Base.fsti" - (Prims.of_int (383)) - (Prims.of_int (15)) - (Prims.of_int (383)) - (Prims.of_int (42))) - (Prims.mk_range - "Steel.ST.GenElim.Base.fsti" - (Prims.of_int (383)) - (Prims.of_int (12)) - (Prims.of_int (396)) - (Prims.of_int (72))) - (Obj.magic - (term_has_head - tl - (FStar_Reflection_Builtins.pack_ln + FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.ST.GenElim.Base.fsti" + (Prims.of_int (383)) + (Prims.of_int (15)) + (Prims.of_int (383)) + (Prims.of_int (42))) + (Prims.mk_range + "Steel.ST.GenElim.Base.fsti" + (Prims.of_int (383)) + (Prims.of_int (12)) + (Prims.of_int (396)) + (Prims.of_int (72))) + (Obj.magic + (term_has_head tl + (FStar_Reflection_Builtins.pack_ln + (FStar_Reflection_Data.Tv_FVar ( - FStar_Reflection_Data.Tv_FVar - (FStar_Reflection_Builtins.pack_fv + FStar_Reflection_Builtins.pack_fv ["Steel"; "ST"; "Util"; "exists_"]))))) + (fun uu___5 -> (fun uu___5 -> - (fun uu___5 -> - if uu___5 - then - Obj.magic - (FStar_Tactics_Effect.tac_bind + if uu___5 + then + Obj.magic + (FStar_Tactics_Effect.tac_bind (Prims.mk_range "Steel.ST.GenElim.Base.fsti" (Prims.of_int (385)) @@ -959,9 +958,9 @@ let rec (solve_gen_elim : FStar_Reflection_Data.Q_Explicit)])))) uu___6))) uu___6)) - else - Obj.magic - (FStar_Tactics_Effect.tac_bind + else + Obj.magic + (FStar_Tactics_Effect.tac_bind (Prims.mk_range "Steel.ST.GenElim.Base.fsti" (Prims.of_int (394)) @@ -1018,11 +1017,10 @@ let rec (solve_gen_elim : (tr', FStar_Reflection_Data.Q_Explicit)])))) uu___7))) - uu___5)) + uu___5) | uu___5 -> - Obj.repr - (FStar_Tactics_Derived.fail - "ill-formed star"))) + FStar_Tactics_Derived.fail + "ill-formed star")) else Obj.magic (Obj.repr @@ -1784,70 +1782,78 @@ let (solve_gen_elim_prop : FStar_Tactics_Effect.lift_div_tac (fun uu___3 -> Prims.op_Negation uu___2)))) (fun uu___2 -> - if uu___2 - then - FStar_Tactics_Derived.fail - "not a squash goal" - else - FStar_Tactics_Effect.lift_div_tac - (fun uu___4 -> ())))) + (fun uu___2 -> + if uu___2 + then + Obj.magic + (Obj.repr + (FStar_Tactics_Derived.fail + "not a squash goal")) + else + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___4 -> ())))) uu___2))) (fun uu___2 -> (fun uu___2 -> match tl with | (body1, FStar_Reflection_Data.Q_Explicit)::[] -> Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.ST.GenElim.Base.fsti" - (Prims.of_int (571)) - (Prims.of_int (21)) - (Prims.of_int (571)) - (Prims.of_int (40))) - (Prims.mk_range - "Steel.ST.GenElim.Base.fsti" - (Prims.of_int (571)) - (Prims.of_int (4)) - (Prims.of_int (603)) - (Prims.of_int (7))) - (FStar_Tactics_Effect.lift_div_tac - (fun uu___3 -> - FStar_Reflection_Derived.collect_app - body1)) + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.ST.GenElim.Base.fsti" + (Prims.of_int (571)) + (Prims.of_int (21)) + (Prims.of_int (571)) + (Prims.of_int (40))) + (Prims.mk_range + "Steel.ST.GenElim.Base.fsti" + (Prims.of_int (571)) + (Prims.of_int (4)) + (Prims.of_int (603)) + (Prims.of_int (7))) + (FStar_Tactics_Effect.lift_div_tac (fun uu___3 -> - (fun uu___3 -> - match uu___3 with - | (hd1, tl1) -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.ST.GenElim.Base.fsti" - (Prims.of_int (572)) - (Prims.of_int (4)) - (Prims.of_int (573)) - (Prims.of_int (42))) - (Prims.mk_range - "Steel.ST.GenElim.Base.fsti" - (Prims.of_int (574)) - (Prims.of_int (10)) - (Prims.of_int (602)) - (Prims.of_int (44))) - (if - Prims.op_Negation - (is_fvar hd1 - "Steel.ST.GenElim.Base.gen_elim_prop") - then - FStar_Tactics_Derived.fail - "not a gen_elim_prop goal" - else - FStar_Tactics_Effect.lift_div_tac - (fun uu___5 -> ())) + FStar_Reflection_Derived.collect_app + body1)) + (fun uu___3 -> + (fun uu___3 -> + match uu___3 with + | (hd1, tl1) -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.ST.GenElim.Base.fsti" + (Prims.of_int (572)) + (Prims.of_int (4)) + (Prims.of_int (573)) + (Prims.of_int (42))) + (Prims.mk_range + "Steel.ST.GenElim.Base.fsti" + (Prims.of_int (574)) + (Prims.of_int (10)) + (Prims.of_int (602)) + (Prims.of_int (44))) + (if + Prims.op_Negation + (is_fvar hd1 + "Steel.ST.GenElim.Base.gen_elim_prop") + then + Obj.magic + (Obj.repr + (FStar_Tactics_Derived.fail + "not a gen_elim_prop goal")) + else + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___5 -> + ())))) + (fun uu___4 -> (fun uu___4 -> - (fun uu___4 -> - match FStar_List_Tot_Base.filter - ( - fun - uu___5 -> + match FStar_List_Tot_Base.filter + (fun uu___5 + -> match uu___5 with | @@ -1855,37 +1861,35 @@ let (solve_gen_elim_prop : x) -> FStar_Reflection_Data.uu___is_Q_Explicit x) tl1 - with - | (enable_nondep_opt_tm, - uu___5):: - (p, uu___6):: - (a, uu___7):: - (q, uu___8):: - (post, - uu___9)::[] - -> - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range + with + | (enable_nondep_opt_tm, + uu___5):: + (p, uu___6):: + (a, uu___7):: + (q, uu___8):: + (post, uu___9)::[] + -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.ST.GenElim.Base.fsti" (Prims.of_int (576)) (Prims.of_int (30)) (Prims.of_int (576)) (Prims.of_int (74))) - (Prims.mk_range + (Prims.mk_range "Steel.ST.GenElim.Base.fsti" (Prims.of_int (577)) (Prims.of_int (6)) (Prims.of_int (601)) (Prims.of_int (44))) - (Obj.magic + (Obj.magic (FStar_Tactics_Builtins.term_eq_old enable_nondep_opt_tm (FStar_Reflection_Builtins.pack_ln (FStar_Reflection_Data.Tv_Const FStar_Reflection_Data.C_True)))) - (fun + (fun uu___10 -> (fun @@ -2412,19 +2416,16 @@ let (solve_gen_elim_prop : uu___10))) uu___10))) uu___10))) - uu___10))) - | uu___5 -> - Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "ill formed gen_elim_prop"))) - uu___4))) uu___3))) + uu___10)) + | uu___5 -> + Obj.magic + (FStar_Tactics_Derived.fail + "ill formed gen_elim_prop")) + uu___4))) uu___3)) | uu___3 -> Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "ill-formed squash"))) uu___2))) - uu___1) + (FStar_Tactics_Derived.fail + "ill-formed squash")) uu___2))) uu___1) let (solve_gen_elim_prop_placeholder : unit -> (Prims.bool, unit) FStar_Tactics_Effect.tac_repr) = fun uu___ -> @@ -2476,70 +2477,78 @@ let (solve_gen_elim_prop_placeholder : FStar_Tactics_Effect.lift_div_tac (fun uu___3 -> Prims.op_Negation uu___2)))) (fun uu___2 -> - if uu___2 - then - FStar_Tactics_Derived.fail - "not a squash goal" - else - FStar_Tactics_Effect.lift_div_tac - (fun uu___4 -> ())))) + (fun uu___2 -> + if uu___2 + then + Obj.magic + (Obj.repr + (FStar_Tactics_Derived.fail + "not a squash goal")) + else + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___4 -> ())))) uu___2))) (fun uu___2 -> (fun uu___2 -> match tl with | (body1, FStar_Reflection_Data.Q_Explicit)::[] -> Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.ST.GenElim.Base.fsti" - (Prims.of_int (615)) - (Prims.of_int (21)) - (Prims.of_int (615)) - (Prims.of_int (40))) - (Prims.mk_range - "Steel.ST.GenElim.Base.fsti" - (Prims.of_int (615)) - (Prims.of_int (4)) - (Prims.of_int (643)) - (Prims.of_int (7))) - (FStar_Tactics_Effect.lift_div_tac - (fun uu___3 -> - FStar_Reflection_Derived.collect_app - body1)) + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.ST.GenElim.Base.fsti" + (Prims.of_int (615)) + (Prims.of_int (21)) + (Prims.of_int (615)) + (Prims.of_int (40))) + (Prims.mk_range + "Steel.ST.GenElim.Base.fsti" + (Prims.of_int (615)) + (Prims.of_int (4)) + (Prims.of_int (643)) + (Prims.of_int (7))) + (FStar_Tactics_Effect.lift_div_tac (fun uu___3 -> - (fun uu___3 -> - match uu___3 with - | (hd1, tl1) -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.ST.GenElim.Base.fsti" - (Prims.of_int (616)) - (Prims.of_int (4)) - (Prims.of_int (617)) - (Prims.of_int (54))) - (Prims.mk_range - "Steel.ST.GenElim.Base.fsti" - (Prims.of_int (618)) - (Prims.of_int (10)) - (Prims.of_int (642)) - (Prims.of_int (56))) - (if - Prims.op_Negation - (is_fvar hd1 - "Steel.ST.GenElim.Base.gen_elim_prop_placeholder") - then - FStar_Tactics_Derived.fail - "not a gen_elim_prop_placeholder goal" - else - FStar_Tactics_Effect.lift_div_tac - (fun uu___5 -> ())) + FStar_Reflection_Derived.collect_app + body1)) + (fun uu___3 -> + (fun uu___3 -> + match uu___3 with + | (hd1, tl1) -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.ST.GenElim.Base.fsti" + (Prims.of_int (616)) + (Prims.of_int (4)) + (Prims.of_int (617)) + (Prims.of_int (54))) + (Prims.mk_range + "Steel.ST.GenElim.Base.fsti" + (Prims.of_int (618)) + (Prims.of_int (10)) + (Prims.of_int (642)) + (Prims.of_int (56))) + (if + Prims.op_Negation + (is_fvar hd1 + "Steel.ST.GenElim.Base.gen_elim_prop_placeholder") + then + Obj.magic + (Obj.repr + (FStar_Tactics_Derived.fail + "not a gen_elim_prop_placeholder goal")) + else + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___5 -> + ())))) + (fun uu___4 -> (fun uu___4 -> - (fun uu___4 -> - match FStar_List_Tot_Base.filter - ( - fun - uu___5 -> + match FStar_List_Tot_Base.filter + (fun uu___5 + -> match uu___5 with | @@ -2547,31 +2556,29 @@ let (solve_gen_elim_prop_placeholder : x) -> FStar_Reflection_Data.uu___is_Q_Explicit x) tl1 - with - | (enable_nondep_opt_tm, - uu___5):: - (p, uu___6):: - (a, uu___7):: - (q, uu___8):: - (post, - uu___9)::[] - -> - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range + with + | (enable_nondep_opt_tm, + uu___5):: + (p, uu___6):: + (a, uu___7):: + (q, uu___8):: + (post, uu___9)::[] + -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.ST.GenElim.Base.fsti" (Prims.of_int (620)) (Prims.of_int (6)) (Prims.of_int (621)) (Prims.of_int (47))) - (Prims.mk_range + (Prims.mk_range "Steel.ST.GenElim.Base.fsti" (Prims.of_int (622)) (Prims.of_int (6)) (Prims.of_int (641)) (Prims.of_int (10))) - (Obj.magic + (Obj.magic (FStar_Tactics_Effect.tac_bind (Prims.mk_range "Steel.ST.GenElim.Base.fsti" @@ -2615,17 +2622,25 @@ let (solve_gen_elim_prop_placeholder : (fun uu___10 -> + (fun + uu___10 + -> if uu___10 then - FStar_Tactics_Derived.fail - "pre-resource not solved yet" + Obj.magic + (Obj.repr + (FStar_Tactics_Derived.fail + "pre-resource not solved yet")) else - FStar_Tactics_Effect.lift_div_tac + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac (fun uu___12 -> ())))) - (fun + uu___10))) + (fun uu___10 -> (fun @@ -2791,13 +2806,17 @@ let (solve_gen_elim_prop_placeholder : && post_is_uvar) then - FStar_Tactics_Derived.fail - "gen_elim_prop_placeholder is already solved" + Obj.magic + (Obj.repr + (FStar_Tactics_Derived.fail + "gen_elim_prop_placeholder is already solved")) else - FStar_Tactics_Effect.lift_div_tac + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac (fun uu___12 - -> ())) + -> ())))) (fun uu___11 -> @@ -3196,19 +3215,16 @@ let (solve_gen_elim_prop_placeholder : uu___11))) uu___11))) uu___11))) - uu___10))) - | uu___5 -> - Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "ill-formed gen_elim_prop_placeholder"))) - uu___4))) uu___3))) + uu___10)) + | uu___5 -> + Obj.magic + (FStar_Tactics_Derived.fail + "ill-formed gen_elim_prop_placeholder")) + uu___4))) uu___3)) | uu___3 -> Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "ill-formed squash"))) uu___2))) - uu___1) + (FStar_Tactics_Derived.fail + "ill-formed squash")) uu___2))) uu___1) let (init_resolve_tac : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = fun uu___ -> Steel_Effect_Common.init_resolve_tac' From 60bd9aa6699448ac6ae71bfad9cd0a5ebce2af02 Mon Sep 17 00:00:00 2001 From: Nikhil Swamy Date: Mon, 20 Mar 2023 22:43:54 -0700 Subject: [PATCH 23/48] support verify and lax to a point --- .../generated/FStar_Interactive_Ide.ml | 63 +++- .../generated/FStar_Interactive_Ide_Types.ml | 29 +- .../FStar_Interactive_Incremental.ml | 335 ++++++++++-------- .../generated/FStar_Interactive_PushHelper.ml | 5 +- src/fstar/FStar.Interactive.Ide.Types.fst | 10 +- src/fstar/FStar.Interactive.Ide.fst | 32 +- src/fstar/FStar.Interactive.Incremental.fst | 50 ++- src/fstar/FStar.Interactive.Incremental.fsti | 5 +- src/fstar/FStar.Interactive.PushHelper.fst | 4 +- 9 files changed, 332 insertions(+), 201 deletions(-) diff --git a/ocaml/fstar-lib/generated/FStar_Interactive_Ide.ml b/ocaml/fstar-lib/generated/FStar_Interactive_Ide.ml index 7984ba9497a..7456aecb260 100644 --- a/ocaml/fstar-lib/generated/FStar_Interactive_Ide.ml +++ b/ocaml/fstar-lib/generated/FStar_Interactive_Ide.ml @@ -341,11 +341,31 @@ let (unpack_interactive_query : FStar_Compiler_Effect.op_Bar_Greater uu___4 FStar_Interactive_JsonHelper.js_int in (uu___1, uu___2, uu___3) in + let read_to_position uu___1 = + let to_pos = + let uu___2 = arg "to-position" in + FStar_Compiler_Effect.op_Bar_Greater uu___2 + FStar_Interactive_JsonHelper.js_assoc in + let uu___2 = + let uu___3 = assoc "to-position.line" "line" to_pos in + FStar_Compiler_Effect.op_Bar_Greater uu___3 + FStar_Interactive_JsonHelper.js_int in + let uu___3 = + let uu___4 = assoc "to-position.column" "column" to_pos in + FStar_Compiler_Effect.op_Bar_Greater uu___4 + FStar_Interactive_JsonHelper.js_int in + ("", uu___2, uu___3) in let parse_full_buffer_kind uu___1 = match uu___1 with | "full" -> FStar_Interactive_Ide_Types.Full | "cache" -> FStar_Interactive_Ide_Types.Cache | "reload-deps" -> FStar_Interactive_Ide_Types.ReloadDeps + | "verify-to-position" -> + let uu___2 = read_to_position () in + FStar_Interactive_Ide_Types.VerifyToPosition uu___2 + | "lax-to-position" -> + let uu___2 = read_to_position () in + FStar_Interactive_Ide_Types.LaxToPosition uu___2 | uu___2 -> FStar_Compiler_Effect.raise (FStar_Interactive_JsonHelper.InvalidQuery @@ -1397,7 +1417,8 @@ let (write_full_buffer_fragment_progress : [uu___1] in write_progress (FStar_Pervasives_Native.Some "full-buffer-fragment-started") uu___ - | FStar_Interactive_Incremental.FragmentSuccess (d, cf) -> + | FStar_Interactive_Incremental.FragmentSuccess + (d, cf, FStar_Interactive_Ide_Types.FullCheck) -> let uu___ = let uu___1 = let uu___2 = @@ -1412,6 +1433,22 @@ let (write_full_buffer_fragment_progress : uu___1 :: uu___2 in write_progress (FStar_Pervasives_Native.Some "full-buffer-fragment-ok") uu___ + | FStar_Interactive_Incremental.FragmentSuccess + (d, cf, FStar_Interactive_Ide_Types.LaxCheck) -> + let uu___ = + let uu___1 = + let uu___2 = + FStar_Compiler_Range.json_of_def_range + d.FStar_Parser_AST.drange in + ("ranges", uu___2) in + let uu___2 = + let uu___3 = + let uu___4 = json_of_code_fragment cf in + ("code-fragment", uu___4) in + [uu___3] in + uu___1 :: uu___2 in + write_progress + (FStar_Pervasives_Native.Some "full-buffer-fragment-lax-ok") uu___ | FStar_Interactive_Incremental.FragmentFailed d -> let uu___ = let uu___1 = @@ -1595,7 +1632,7 @@ let (run_push_without_deps : let st1 = set_nosynth_flag st peek_only in let uu___2 = run_repl_transaction st1 push_kind peek_only - (FStar_Interactive_Ide_Types.PushFragment frag) in + (FStar_Interactive_Ide_Types.PushFragment (frag, push_kind)) in match uu___2 with | (success, st2) -> let st3 = set_nosynth_flag st2 false in @@ -1612,15 +1649,15 @@ let (run_push_without_deps : | FStar_Errors.ENotImplemented -> true | uu___3 -> false) errs in ((match code_or_decl with - | FStar_Pervasives.Inr ds -> + | FStar_Pervasives.Inr (d, s) -> if Prims.op_Negation has_error then write_full_buffer_fragment_progress - (FStar_Interactive_Incremental.FragmentSuccess ds) + (FStar_Interactive_Incremental.FragmentSuccess + (d, s, push_kind)) else write_full_buffer_fragment_progress - (FStar_Interactive_Incremental.FragmentFailed - (FStar_Pervasives_Native.fst ds)) + (FStar_Interactive_Incremental.FragmentFailed d) | uu___4 -> ()); (let json_errors = let uu___4 = @@ -2511,6 +2548,15 @@ let (maybe_cancel_queries : = fun st -> fun l -> + let log_cancellation l1 = + let uu___ = FStar_Options.debug_any () in + if uu___ + then + FStar_Compiler_List.iter + (fun q -> + let uu___1 = FStar_Interactive_Ide_Types.query_to_string q in + FStar_Compiler_Util.print1 "Cancelling query: %s\n" uu___1) l1 + else () in match st.FStar_Interactive_Ide_Types.repl_buffered_input_queries with | { FStar_Interactive_Ide_Types.qq = FStar_Interactive_Ide_Types.Cancel @@ -2537,7 +2583,7 @@ let (maybe_cancel_queries : FStar_Interactive_Ide_Types.repl_buffered_input_queries = rest } in (match p with - | FStar_Pervasives_Native.None -> ([], st1) + | FStar_Pervasives_Native.None -> (log_cancellation l; ([], st1)) | FStar_Pervasives_Native.Some p1 -> let query_ahead_of p2 q = let uu___1 = p2 in @@ -2552,7 +2598,8 @@ let (maybe_cancel_queries : FStar_Compiler_Util.prefix_until (query_ahead_of p1) l in match uu___1 with | FStar_Pervasives_Native.None -> l - | FStar_Pervasives_Native.Some (l2, uu___2, uu___3) -> l2 in + | FStar_Pervasives_Native.Some (l2, q, qs) -> + (log_cancellation (q :: qs); l2) in (l1, st1)) | uu___ -> (l, st) let rec (fold_query : diff --git a/ocaml/fstar-lib/generated/FStar_Interactive_Ide_Types.ml b/ocaml/fstar-lib/generated/FStar_Interactive_Ide_Types.ml index 59ef2d6e3c0..0a2ba4e0a3b 100644 --- a/ocaml/fstar-lib/generated/FStar_Interactive_Ide_Types.ml +++ b/ocaml/fstar-lib/generated/FStar_Interactive_Ide_Types.ml @@ -116,8 +116,8 @@ type repl_task = | LDInterleaved of (timed_fname * timed_fname) | LDSingle of timed_fname | LDInterfaceOfCurrentFile of timed_fname - | PushFragment of (FStar_Parser_ParseIt.input_frag, FStar_Parser_AST.decl) - FStar_Pervasives.either + | PushFragment of ((FStar_Parser_ParseIt.input_frag, FStar_Parser_AST.decl) + FStar_Pervasives.either * push_kind) | Noop let (uu___is_LDInterleaved : repl_task -> Prims.bool) = fun projectee -> @@ -142,8 +142,8 @@ let (uu___is_PushFragment : repl_task -> Prims.bool) = match projectee with | PushFragment _0 -> true | uu___ -> false let (__proj__PushFragment__item___0 : repl_task -> - (FStar_Parser_ParseIt.input_frag, FStar_Parser_AST.decl) - FStar_Pervasives.either) + ((FStar_Parser_ParseIt.input_frag, FStar_Parser_AST.decl) + FStar_Pervasives.either * push_kind)) = fun projectee -> match projectee with | PushFragment _0 -> _0 let (uu___is_Noop : repl_task -> Prims.bool) = fun projectee -> match projectee with | Noop -> true | uu___ -> false @@ -151,12 +151,25 @@ type full_buffer_request_kind = | Full | Cache | ReloadDeps + | VerifyToPosition of position + | LaxToPosition of position let (uu___is_Full : full_buffer_request_kind -> Prims.bool) = fun projectee -> match projectee with | Full -> true | uu___ -> false let (uu___is_Cache : full_buffer_request_kind -> Prims.bool) = fun projectee -> match projectee with | Cache -> true | uu___ -> false let (uu___is_ReloadDeps : full_buffer_request_kind -> Prims.bool) = fun projectee -> match projectee with | ReloadDeps -> true | uu___ -> false +let (uu___is_VerifyToPosition : full_buffer_request_kind -> Prims.bool) = + fun projectee -> + match projectee with | VerifyToPosition _0 -> true | uu___ -> false +let (__proj__VerifyToPosition__item___0 : + full_buffer_request_kind -> position) = + fun projectee -> match projectee with | VerifyToPosition _0 -> _0 +let (uu___is_LaxToPosition : full_buffer_request_kind -> Prims.bool) = + fun projectee -> + match projectee with | LaxToPosition _0 -> true | uu___ -> false +let (__proj__LaxToPosition__item___0 : full_buffer_request_kind -> position) + = fun projectee -> match projectee with | LaxToPosition _0 -> _0 type query' = | Exit | DescribeProtocol @@ -384,12 +397,12 @@ let (string_of_repl_task : repl_task -> Prims.string) = | LDInterfaceOfCurrentFile intf -> let uu___1 = string_of_timed_fname intf in FStar_Compiler_Util.format1 "LDInterfaceOfCurrentFile %s" uu___1 - | PushFragment (FStar_Pervasives.Inl frag) -> + | PushFragment (FStar_Pervasives.Inl frag, uu___1) -> FStar_Compiler_Util.format1 "PushFragment { code = %s }" frag.FStar_Parser_ParseIt.frag_text - | PushFragment (FStar_Pervasives.Inr d) -> - let uu___1 = FStar_Parser_AST.decl_to_string d in - FStar_Compiler_Util.format1 "PushFragment { decl = %s }" uu___1 + | PushFragment (FStar_Pervasives.Inr d, uu___1) -> + let uu___2 = FStar_Parser_AST.decl_to_string d in + FStar_Compiler_Util.format1 "PushFragment { decl = %s }" uu___2 | Noop -> "Noop {}" let (string_of_repl_stack_entry : repl_stack_entry_t -> Prims.string) = fun uu___ -> diff --git a/ocaml/fstar-lib/generated/FStar_Interactive_Incremental.ml b/ocaml/fstar-lib/generated/FStar_Interactive_Incremental.ml index 046efebbc9f..21b0424bc6c 100644 --- a/ocaml/fstar-lib/generated/FStar_Interactive_Incremental.ml +++ b/ocaml/fstar-lib/generated/FStar_Interactive_Incremental.ml @@ -2,7 +2,8 @@ open Prims type fragment_progress = | FragmentStarted of FStar_Parser_AST.decl | FragmentSuccess of (FStar_Parser_AST.decl * - FStar_Parser_ParseIt.code_fragment) + FStar_Parser_ParseIt.code_fragment * FStar_Interactive_Ide_Types.push_kind) + | FragmentFailed of FStar_Parser_AST.decl | FragmentError of FStar_Errors.issue Prims.list let (uu___is_FragmentStarted : fragment_progress -> Prims.bool) = @@ -16,7 +17,8 @@ let (uu___is_FragmentSuccess : fragment_progress -> Prims.bool) = match projectee with | FragmentSuccess _0 -> true | uu___ -> false let (__proj__FragmentSuccess__item___0 : fragment_progress -> - (FStar_Parser_AST.decl * FStar_Parser_ParseIt.code_fragment)) + (FStar_Parser_AST.decl * FStar_Parser_ParseIt.code_fragment * + FStar_Interactive_Ide_Types.push_kind)) = fun projectee -> match projectee with | FragmentSuccess _0 -> _0 let (uu___is_FragmentFailed : fragment_progress -> Prims.bool) = fun projectee -> @@ -82,52 +84,58 @@ let (as_query : } in return uu___1) let (push_decl : - (fragment_progress -> unit) -> - (FStar_Parser_AST.decl * FStar_Parser_ParseIt.code_fragment) -> - FStar_Interactive_Ide_Types.query Prims.list qst) + FStar_Interactive_Ide_Types.push_kind -> + (fragment_progress -> unit) -> + (FStar_Parser_AST.decl * FStar_Parser_ParseIt.code_fragment) -> + FStar_Interactive_Ide_Types.query Prims.list qst) = - fun write_full_buffer_fragment_progress -> - fun ds -> - let uu___ = ds in - match uu___ with - | (d, s) -> - let pq = - let uu___1 = + fun push_kind -> + fun write_full_buffer_fragment_progress -> + fun ds -> + let uu___ = ds in + match uu___ with + | (d, s) -> + let pq = + let uu___1 = + let uu___2 = + FStar_Compiler_Range.start_of_range + d.FStar_Parser_AST.drange in + FStar_Compiler_Range.line_of_pos uu___2 in let uu___2 = - FStar_Compiler_Range.start_of_range d.FStar_Parser_AST.drange in - FStar_Compiler_Range.line_of_pos uu___2 in - let uu___2 = - let uu___3 = - FStar_Compiler_Range.start_of_range d.FStar_Parser_AST.drange in - FStar_Compiler_Range.col_of_pos uu___3 in - { - FStar_Interactive_Ide_Types.push_kind = - FStar_Interactive_Ide_Types.FullCheck; - FStar_Interactive_Ide_Types.push_line = uu___1; - FStar_Interactive_Ide_Types.push_column = uu___2; - FStar_Interactive_Ide_Types.push_peek_only = false; - FStar_Interactive_Ide_Types.push_code_or_decl = - (FStar_Pervasives.Inr ds) - } in - let progress st = - write_full_buffer_fragment_progress (FragmentStarted d); - ((FStar_Interactive_Ide_Types.QueryOK, []), - (FStar_Pervasives.Inl st)) in - let uu___1 = - as_query (FStar_Interactive_Ide_Types.Callback progress) in - op_let_Bang uu___1 - (fun cb -> - let uu___2 = as_query (FStar_Interactive_Ide_Types.Push pq) in - op_let_Bang uu___2 (fun push -> return [cb; push])) + let uu___3 = + FStar_Compiler_Range.start_of_range + d.FStar_Parser_AST.drange in + FStar_Compiler_Range.col_of_pos uu___3 in + { + FStar_Interactive_Ide_Types.push_kind = push_kind; + FStar_Interactive_Ide_Types.push_line = uu___1; + FStar_Interactive_Ide_Types.push_column = uu___2; + FStar_Interactive_Ide_Types.push_peek_only = false; + FStar_Interactive_Ide_Types.push_code_or_decl = + (FStar_Pervasives.Inr ds) + } in + let progress st = + write_full_buffer_fragment_progress (FragmentStarted d); + ((FStar_Interactive_Ide_Types.QueryOK, []), + (FStar_Pervasives.Inl st)) in + let uu___1 = + as_query (FStar_Interactive_Ide_Types.Callback progress) in + op_let_Bang uu___1 + (fun cb -> + let uu___2 = as_query (FStar_Interactive_Ide_Types.Push pq) in + op_let_Bang uu___2 (fun push -> return [cb; push])) let (push_decls : - (fragment_progress -> unit) -> - (FStar_Parser_AST.decl * FStar_Parser_ParseIt.code_fragment) Prims.list - -> FStar_Interactive_Ide_Types.query Prims.list qst) + FStar_Interactive_Ide_Types.push_kind -> + (fragment_progress -> unit) -> + (FStar_Parser_AST.decl * FStar_Parser_ParseIt.code_fragment) Prims.list + -> FStar_Interactive_Ide_Types.query Prims.list qst) = - fun write_full_buffer_fragment_progress -> - fun ds -> - let uu___ = map (push_decl write_full_buffer_fragment_progress) ds in - op_let_Bang uu___ (fun qs -> return (FStar_Compiler_List.flatten qs)) + fun push_kind -> + fun write_full_buffer_fragment_progress -> + fun ds -> + let uu___ = + map (push_decl push_kind write_full_buffer_fragment_progress) ds in + op_let_Bang uu___ (fun qs -> return (FStar_Compiler_List.flatten qs)) let (pop_entries : FStar_Interactive_Ide_Types.repl_stack_entry_t Prims.list -> FStar_Interactive_Ide_Types.query Prims.list qst) @@ -161,73 +169,80 @@ let (inspect_repl_stack : FStar_Interactive_Ide_Types.repl_stack_t -> (FStar_Parser_AST.decl * FStar_Parser_ParseIt.code_fragment) Prims.list -> - (fragment_progress -> unit) -> - FStar_Interactive_Ide_Types.query Prims.list qst) + FStar_Interactive_Ide_Types.push_kind -> + (fragment_progress -> unit) -> + FStar_Interactive_Ide_Types.query Prims.list qst) = fun s -> fun ds -> - fun write_full_buffer_fragment_progress -> - let entries = FStar_Compiler_List.rev s in - let push_decls1 = push_decls write_full_buffer_fragment_progress in - let uu___ = - FStar_Compiler_Util.prefix_until - (fun uu___1 -> - match uu___1 with - | (uu___2, - (FStar_Interactive_Ide_Types.PushFragment uu___3, uu___4)) - -> true - | uu___2 -> false) entries in - match uu___ with - | FStar_Pervasives_Native.None -> - let uu___1 = push_decls1 ds in - op_let_Bang uu___1 (fun ds1 -> return ds1) - | FStar_Pervasives_Native.Some (prefix, first_push, rest) -> - let entries1 = first_push :: rest in - let repl_task1 uu___1 = - match uu___1 with | (uu___2, (p, uu___3)) -> p in - let rec matching_prefix entries2 ds1 = - match (entries2, ds1) with - | ([], []) -> return [] - | (e::entries3, d::ds2) -> - (match repl_task1 e with - | FStar_Interactive_Ide_Types.Noop -> - matching_prefix entries3 (d :: ds2) - | FStar_Interactive_Ide_Types.PushFragment - (FStar_Pervasives.Inl frag) -> - let uu___1 = pop_entries (e :: entries3) in - op_let_Bang uu___1 - (fun pops -> - let uu___2 = push_decls1 (d :: ds2) in - op_let_Bang uu___2 - (fun pushes -> - return - (FStar_Compiler_List.op_At pops pushes))) - | FStar_Interactive_Ide_Types.PushFragment - (FStar_Pervasives.Inr d') -> - let uu___1 = - FStar_Parser_AST_Comparison.eq_decl - (FStar_Pervasives_Native.fst d) d' in - if uu___1 - then - (write_full_buffer_fragment_progress - (FragmentSuccess d); - matching_prefix entries3 ds2) - else - (let uu___3 = pop_entries (e :: entries3) in - op_let_Bang uu___3 - (fun pops -> - let uu___4 = push_decls1 (d :: ds2) in - op_let_Bang uu___4 - (fun pushes -> - return - (FStar_Compiler_List.op_At pops pushes))))) - | ([], ds2) -> - let uu___1 = push_decls1 ds2 in - op_let_Bang uu___1 (fun pushes -> return pushes) - | (es, []) -> - let uu___1 = pop_entries es in - op_let_Bang uu___1 (fun pops -> return pops) in - matching_prefix entries1 ds + fun push_kind -> + fun write_full_buffer_fragment_progress -> + let entries = FStar_Compiler_List.rev s in + let push_decls1 = + push_decls push_kind write_full_buffer_fragment_progress in + let uu___ = + FStar_Compiler_Util.prefix_until + (fun uu___1 -> + match uu___1 with + | (uu___2, + (FStar_Interactive_Ide_Types.PushFragment uu___3, uu___4)) + -> true + | uu___2 -> false) entries in + match uu___ with + | FStar_Pervasives_Native.None -> + let uu___1 = push_decls1 ds in + op_let_Bang uu___1 (fun ds1 -> return ds1) + | FStar_Pervasives_Native.Some (prefix, first_push, rest) -> + let entries1 = first_push :: rest in + let repl_task1 uu___1 = + match uu___1 with | (uu___2, (p, uu___3)) -> p in + let rec matching_prefix entries2 ds1 = + match (entries2, ds1) with + | ([], []) -> return [] + | (e::entries3, d::ds2) -> + (match repl_task1 e with + | FStar_Interactive_Ide_Types.Noop -> + matching_prefix entries3 (d :: ds2) + | FStar_Interactive_Ide_Types.PushFragment + (FStar_Pervasives.Inl frag, uu___1) -> + let uu___2 = pop_entries (e :: entries3) in + op_let_Bang uu___2 + (fun pops -> + let uu___3 = push_decls1 (d :: ds2) in + op_let_Bang uu___3 + (fun pushes -> + return + (FStar_Compiler_List.op_At pops pushes))) + | FStar_Interactive_Ide_Types.PushFragment + (FStar_Pervasives.Inr d', pk) -> + let uu___1 = + FStar_Parser_AST_Comparison.eq_decl + (FStar_Pervasives_Native.fst d) d' in + if uu___1 + then + let uu___2 = d in + (match uu___2 with + | (d1, s1) -> + (write_full_buffer_fragment_progress + (FragmentSuccess (d1, s1, pk)); + matching_prefix entries3 ds2)) + else + (let uu___3 = pop_entries (e :: entries3) in + op_let_Bang uu___3 + (fun pops -> + let uu___4 = push_decls1 (d :: ds2) in + op_let_Bang uu___4 + (fun pushes -> + return + (FStar_Compiler_List.op_At pops + pushes))))) + | ([], ds2) -> + let uu___1 = push_decls1 ds2 in + op_let_Bang uu___1 (fun pushes -> return pushes) + | (es, []) -> + let uu___1 = pop_entries es in + op_let_Bang uu___1 (fun pops -> return pops) in + matching_prefix entries1 ds let reload_deps : 'uuuuu 'uuuuu1 . ('uuuuu * (FStar_Interactive_Ide_Types.repl_task * 'uuuuu1)) Prims.list @@ -318,72 +333,82 @@ let (run_full_buffer : | FStar_Pervasives_Native.Some err1 -> let issue = syntax_issue err1 in write_full_buffer_fragment_progress (FragmentError [issue]) in + let filter_decls decls = + match request_type with + | FStar_Interactive_Ide_Types.VerifyToPosition + (uu___, line, _col) -> + ((let uu___2 = FStar_Compiler_Util.string_of_int line in + FStar_Compiler_Util.print1 "Got to-position: %s" uu___2); + FStar_Compiler_List.filter + (fun uu___2 -> + match uu___2 with + | (d, uu___3) -> + let start = + FStar_Compiler_Range.start_of_range + d.FStar_Parser_AST.drange in + let start_line = + FStar_Compiler_Range.line_of_pos start in + start_line <= line) decls) + | FStar_Interactive_Ide_Types.LaxToPosition (uu___, line, _col) + -> + ((let uu___2 = FStar_Compiler_Util.string_of_int line in + FStar_Compiler_Util.print1 "Got to-position: %s" uu___2); + FStar_Compiler_List.filter + (fun uu___2 -> + match uu___2 with + | (d, uu___3) -> + let start = + FStar_Compiler_Range.start_of_range + d.FStar_Parser_AST.drange in + let start_line = + FStar_Compiler_Range.line_of_pos start in + start_line <= line) decls) + | uu___ -> decls in let qs = match parse_result with | FStar_Parser_ParseIt.IncrementalFragment (decls, uu___, err_opt) -> + let decls1 = filter_decls decls in (match request_type with - | FStar_Interactive_Ide_Types.Full -> - let queries = - let uu___1 = - let uu___2 = - FStar_Compiler_Effect.op_Bang - FStar_Interactive_PushHelper.repl_stack in - inspect_repl_stack uu___2 decls - write_full_buffer_fragment_progress in - run_qst uu___1 qid1 in - (if request_type = FStar_Interactive_Ide_Types.Full - then log_syntax_issues err_opt - else (); - (let uu___3 = FStar_Options.debug_any () in - if uu___3 - then - let uu___4 = - let uu___5 = - FStar_Compiler_List.map - FStar_Interactive_Ide_Types.query_to_string - queries in - FStar_String.concat "\n" uu___5 in - FStar_Compiler_Util.print1 - "Generating queries\n%s\n" uu___4 - else ()); - if request_type = FStar_Interactive_Ide_Types.Full - then queries - else []) - | FStar_Interactive_Ide_Types.Cache -> + | FStar_Interactive_Ide_Types.ReloadDeps -> + let uu___1 = + let uu___2 = + FStar_Compiler_Effect.op_Bang + FStar_Interactive_PushHelper.repl_stack in + reload_deps uu___2 in + run_qst uu___1 qid1 + | uu___1 -> + let push_kind = + match request_type with + | FStar_Interactive_Ide_Types.LaxToPosition uu___2 + -> FStar_Interactive_Ide_Types.LaxCheck + | uu___2 -> FStar_Interactive_Ide_Types.FullCheck in let queries = - let uu___1 = - let uu___2 = + let uu___2 = + let uu___3 = FStar_Compiler_Effect.op_Bang FStar_Interactive_PushHelper.repl_stack in - inspect_repl_stack uu___2 decls + inspect_repl_stack uu___3 decls1 push_kind write_full_buffer_fragment_progress in - run_qst uu___1 qid1 in + run_qst uu___2 qid1 in (if request_type = FStar_Interactive_Ide_Types.Full then log_syntax_issues err_opt else (); - (let uu___3 = FStar_Options.debug_any () in - if uu___3 + (let uu___4 = FStar_Options.debug_any () in + if uu___4 then - let uu___4 = - let uu___5 = + let uu___5 = + let uu___6 = FStar_Compiler_List.map FStar_Interactive_Ide_Types.query_to_string queries in - FStar_String.concat "\n" uu___5 in + FStar_String.concat "\n" uu___6 in FStar_Compiler_Util.print1 - "Generating queries\n%s\n" uu___4 + "Generating queries\n%s\n" uu___5 else ()); - if request_type = FStar_Interactive_Ide_Types.Full + if request_type <> FStar_Interactive_Ide_Types.Cache then queries - else []) - | FStar_Interactive_Ide_Types.ReloadDeps -> - let uu___1 = - let uu___2 = - FStar_Compiler_Effect.op_Bang - FStar_Interactive_PushHelper.repl_stack in - reload_deps uu___2 in - run_qst uu___1 qid1) + else [])) | FStar_Parser_ParseIt.ParseError err -> (if request_type = FStar_Interactive_Ide_Types.Full then log_syntax_issues (FStar_Pervasives_Native.Some err) diff --git a/ocaml/fstar-lib/generated/FStar_Interactive_PushHelper.ml b/ocaml/fstar-lib/generated/FStar_Interactive_PushHelper.ml index 8cc0e5e5b68..045eaf632b0 100644 --- a/ocaml/fstar-lib/generated/FStar_Interactive_PushHelper.ml +++ b/ocaml/fstar-lib/generated/FStar_Interactive_PushHelper.ml @@ -368,7 +368,7 @@ let (run_repl_task : FStar_Universal.load_interface_decls env intf.FStar_Interactive_Ide_Types.tf_fname in (curmod, uu___) - | FStar_Interactive_Ide_Types.PushFragment frag -> + | FStar_Interactive_Ide_Types.PushFragment (frag, uu___) -> FStar_Universal.tc_one_fragment curmod env frag | FStar_Interactive_Ide_Types.Noop -> (curmod, env) let (query_of_ids : @@ -903,5 +903,6 @@ let (full_lax : (st1.FStar_Interactive_Ide_Types.repl_buffered_input_queries) } FStar_Interactive_Ide_Types.LaxCheck (FStar_Interactive_Ide_Types.PushFragment - (FStar_Pervasives.Inl frag)) + ((FStar_Pervasives.Inl frag), + FStar_Interactive_Ide_Types.LaxCheck)) | FStar_Pervasives.Inr st1 -> (FStar_Pervasives_Native.None, st1)) \ No newline at end of file diff --git a/src/fstar/FStar.Interactive.Ide.Types.fst b/src/fstar/FStar.Interactive.Ide.Types.fst index 417702138af..50a6969c69d 100755 --- a/src/fstar/FStar.Interactive.Ide.Types.fst +++ b/src/fstar/FStar.Interactive.Ide.Types.fst @@ -95,14 +95,16 @@ type repl_task = | LDInterleaved of timed_fname * timed_fname (* (interface * implementation) *) | LDSingle of timed_fname (* interface or implementation *) | LDInterfaceOfCurrentFile of timed_fname (* interface *) - | PushFragment of either PI.input_frag FStar.Parser.AST.decl (* code fragment *) + | PushFragment of either PI.input_frag FStar.Parser.AST.decl * push_kind (* code fragment *) | Noop (* Used by compute *) type full_buffer_request_kind = | Full : full_buffer_request_kind | Cache : full_buffer_request_kind | ReloadDeps : full_buffer_request_kind - + | VerifyToPosition of position + | LaxToPosition of position + type query' = | Exit | DescribeProtocol @@ -163,9 +165,9 @@ let string_of_repl_task = function Util.format1 "LDSingle %s" (string_of_timed_fname intf_or_impl) | LDInterfaceOfCurrentFile intf -> Util.format1 "LDInterfaceOfCurrentFile %s" (string_of_timed_fname intf) - | PushFragment (Inl frag) -> + | PushFragment (Inl frag, _) -> Util.format1 "PushFragment { code = %s }" frag.frag_text - | PushFragment (Inr d) -> + | PushFragment (Inr d, _) -> Util.format1 "PushFragment { decl = %s }" (FStar.Parser.AST.decl_to_string d) | Noop -> "Noop {}" diff --git a/src/fstar/FStar.Interactive.Ide.fst b/src/fstar/FStar.Interactive.Ide.fst index b3fb90a1255..94c30fb0dfc 100644 --- a/src/fstar/FStar.Interactive.Ide.fst +++ b/src/fstar/FStar.Interactive.Ide.fst @@ -218,10 +218,18 @@ let unpack_interactive_query json = assoc err "line" loc |> js_int, assoc err "column" loc |> js_int in + let read_to_position () = + let to_pos = arg "to-position" |> js_assoc in + "", + assoc "to-position.line" "line" to_pos |> js_int, + assoc "to-position.column" "column" to_pos |> js_int + in let parse_full_buffer_kind = function | "full" -> Full | "cache" -> Cache | "reload-deps" -> ReloadDeps + | "verify-to-position" -> VerifyToPosition (read_to_position ()) + | "lax-to-position" -> LaxToPosition (read_to_position ()) | _ -> raise (InvalidQuery "Invalid full-buffer kind") in { qid = qid; @@ -596,10 +604,14 @@ let write_full_buffer_fragment_progress (di:Incremental.fragment_progress) = | FragmentStarted d -> write_progress (Some "full-buffer-fragment-started") ["ranges", json_of_def_range d.FStar.Parser.AST.drange] - | FragmentSuccess (d, cf) -> + | FragmentSuccess (d, cf, FullCheck) -> write_progress (Some "full-buffer-fragment-ok") ["ranges", json_of_def_range d.FStar.Parser.AST.drange; "code-fragment", json_of_code_fragment cf] + | FragmentSuccess (d, cf, LaxCheck) -> + write_progress (Some "full-buffer-fragment-lax-ok") + ["ranges", json_of_def_range d.FStar.Parser.AST.drange; + "code-fragment", json_of_code_fragment cf] | FragmentFailed d -> write_progress (Some "full-buffer-fragment-failed") ["ranges", json_of_def_range d.FStar.Parser.AST.drange] @@ -637,7 +649,7 @@ let run_push_without_deps st query Inr decl in let st = set_nosynth_flag st peek_only in - let success, st = run_repl_transaction st push_kind peek_only (PushFragment frag) in + let success, st = run_repl_transaction st push_kind peek_only (PushFragment (frag, push_kind)) in let st = set_nosynth_flag st false in let status = if success || peek_only then QueryOK else QueryNOK in @@ -652,10 +664,10 @@ let run_push_without_deps st query in let _ = match code_or_decl with - | Inr ds -> + | Inr (d, s) -> if not has_error - then write_full_buffer_fragment_progress (Incremental.FragmentSuccess ds) - else write_full_buffer_fragment_progress (Incremental.FragmentFailed (fst ds)) + then write_full_buffer_fragment_progress (Incremental.FragmentSuccess (d, s, push_kind)) + else write_full_buffer_fragment_progress (Incremental.FragmentFailed d) | _ -> () in let json_errors = JsonList (errs |> List.map json_of_issue) in @@ -969,11 +981,16 @@ let as_json_list (q: (query_status & json) & either repl_state int) let run_query_result = (query_status * list json) * either repl_state int let maybe_cancel_queries st l = + let log_cancellation l = + if Options.debug_any() + then List.iter (fun q -> BU.print1 "Cancelling query: %s\n" (query_to_string q)) l + in match st.repl_buffered_input_queries with | { qq = Cancel p } :: rest -> ( let st = { st with repl_buffered_input_queries = rest } in match p with | None -> //If no range, then cancel all remaining queries + log_cancellation l; [], st | Some p -> //Cancel all queries that are within the range let query_ahead_of p q = @@ -985,7 +1002,9 @@ let maybe_cancel_queries st l = let l = match BU.prefix_until (query_ahead_of p) l with | None -> l - | Some (l, _, _) -> l + | Some (l, q, qs) -> + log_cancellation (q::qs); + l in l, st ) @@ -1153,7 +1172,6 @@ let interactive_mode' init_st = let interactive_mode (filename:string): unit = install_ide_mode_hooks write_json; - // Ignore unexpected interrupts (some methods override this handler) Util.set_sigint_handler Util.sigint_ignore; diff --git a/src/fstar/FStar.Interactive.Incremental.fst b/src/fstar/FStar.Interactive.Incremental.fst index 267c0d37c5b..f194c493d25 100755 --- a/src/fstar/FStar.Interactive.Incremental.fst +++ b/src/fstar/FStar.Interactive.Incremental.fst @@ -87,13 +87,14 @@ let as_query (q:query') qid=qid_prefix ^ "." ^ string_of_int i } -let push_decl (write_full_buffer_fragment_progress: fragment_progress -> unit) +let push_decl (push_kind:push_kind) + (write_full_buffer_fragment_progress: fragment_progress -> unit) (ds:decl * code_fragment) : qst (list query) = let open FStar.Compiler.Range in let d, s = ds in let pq = { - push_kind = FullCheck; + push_kind; push_line = line_of_pos (start_of_range d.drange); push_column = col_of_pos (start_of_range d.drange); push_peek_only = false; @@ -107,10 +108,11 @@ let push_decl (write_full_buffer_fragment_progress: fragment_progress -> unit) let! push = as_query (Push pq) in return [cb; push] -let push_decls (write_full_buffer_fragment_progress : fragment_progress -> unit) +let push_decls (push_kind:push_kind) + (write_full_buffer_fragment_progress : fragment_progress -> unit) (ds:list (decl & code_fragment)) : qst (list query) - = let! qs = map (push_decl write_full_buffer_fragment_progress) ds in + = let! qs = map (push_decl push_kind write_full_buffer_fragment_progress) ds in return (List.flatten qs) let pop_entries (e:list repl_stack_entry_t) @@ -130,10 +132,11 @@ let repl_task (_, (p, _)) = p let inspect_repl_stack (s:repl_stack_t) (ds:list (decl & code_fragment)) + (push_kind : push_kind) (write_full_buffer_fragment_progress: fragment_progress -> unit) : qst (list query) // & list json) = let entries = List.rev s in - let push_decls = push_decls write_full_buffer_fragment_progress in + let push_decls = push_decls push_kind write_full_buffer_fragment_progress in match BU.prefix_until (function (_, (PushFragment _, _)) -> true | _ -> false) entries @@ -155,14 +158,15 @@ let inspect_repl_stack (s:repl_stack_t) match repl_task e with | Noop -> matching_prefix entries (d::ds) - | PushFragment (Inl frag) -> + | PushFragment (Inl frag, _) -> let! pops = pop_entries (e::entries) in let! pushes = push_decls (d::ds) in return (pops @ pushes) - | PushFragment (Inr d') -> + | PushFragment (Inr d', pk) -> if eq_decl (fst d) d' then ( - write_full_buffer_fragment_progress (FragmentSuccess d); + let d, s = d in + write_full_buffer_fragment_progress (FragmentSuccess (d, s, pk)); matching_prefix entries ds ) else let! pops = pop_entries (e::entries) in @@ -230,13 +234,35 @@ let run_full_buffer (st:repl_state) let issue = syntax_issue err in write_full_buffer_fragment_progress (FragmentError [issue]) in + let filter_decls decls = + match request_type with + | VerifyToPosition (_, line, _col) + | LaxToPosition (_, line, _col) -> + BU.print1 "Got to-position: %s" (string_of_int line); + List.filter + (fun (d, _) -> + let start = Range.start_of_range d.drange in + let start_line = Range.line_of_pos start in + start_line <= line) + decls + | _ -> decls + in let qs = match parse_result with | IncrementalFragment (decls, _, err_opt) -> ( + let decls = filter_decls decls in match request_type with - | Full | Cache -> + | ReloadDeps -> + run_qst (reload_deps (!repl_stack)) qid + + | _ -> + let push_kind = + match request_type with + | LaxToPosition _ -> LaxCheck + | _ -> FullCheck + in let queries = - run_qst (inspect_repl_stack (!repl_stack) decls write_full_buffer_fragment_progress) qid + run_qst (inspect_repl_stack (!repl_stack) decls push_kind write_full_buffer_fragment_progress) qid in if request_type = Full then log_syntax_issues err_opt; if Options.debug_any() @@ -244,10 +270,8 @@ let run_full_buffer (st:repl_state) BU.print1 "Generating queries\n%s\n" (String.concat "\n" (List.map query_to_string queries)) ); - if request_type = Full then queries else [] + if request_type <> Cache then queries else [] - | ReloadDeps -> - run_qst (reload_deps (!repl_stack)) qid ) | ParseError err -> diff --git a/src/fstar/FStar.Interactive.Incremental.fsti b/src/fstar/FStar.Interactive.Incremental.fsti index 19638da7e99..8918f269377 100755 --- a/src/fstar/FStar.Interactive.Incremental.fsti +++ b/src/fstar/FStar.Interactive.Incremental.fsti @@ -27,7 +27,7 @@ open FStar.Interactive.Ide.Types type fragment_progress = | FragmentStarted of decl - | FragmentSuccess of (decl & FStar.Parser.ParseIt.code_fragment) + | FragmentSuccess of (decl & FStar.Parser.ParseIt.code_fragment & push_kind) | FragmentFailed of decl | FragmentError of list issue @@ -50,4 +50,5 @@ val run_full_buffer (st:repl_state) val format_code (st:repl_state) (code:string) - : either string (list issue) \ No newline at end of file + : either string (list issue) + \ No newline at end of file diff --git a/src/fstar/FStar.Interactive.PushHelper.fst b/src/fstar/FStar.Interactive.PushHelper.fst index ff872a02db2..6538435d0e5 100644 --- a/src/fstar/FStar.Interactive.PushHelper.fst +++ b/src/fstar/FStar.Interactive.PushHelper.fst @@ -166,7 +166,7 @@ let run_repl_task (curmod: optmod_t) (env: env_t) (task: repl_task) : optmod_t * curmod, tc_one env None intf_or_impl.tf_fname | LDInterfaceOfCurrentFile intf -> curmod, Universal.load_interface_decls env intf.tf_fname - | PushFragment frag -> + | PushFragment (frag, _) -> tc_one_fragment curmod env frag | Noop -> curmod, env @@ -371,5 +371,5 @@ let full_lax text st = match ld_deps st with | Inl (st, deps) -> let names = add_module_completions st.repl_fname deps st.repl_names in - repl_tx ({ st with repl_names = names }) LaxCheck (PushFragment (Inl frag)) + repl_tx ({ st with repl_names = names }) LaxCheck (PushFragment (Inl frag, LaxCheck)) | Inr st -> None, st From ec203019c20f1c83c7038277a62dd92eba0b7e84 Mon Sep 17 00:00:00 2001 From: Nikhil Swamy Date: Tue, 21 Mar 2023 07:14:12 -0700 Subject: [PATCH 24/48] try fix off-by-one in ParseIt --- ocaml/fstar-lib/FStar_Parser_ParseIt.ml | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/ocaml/fstar-lib/FStar_Parser_ParseIt.ml b/ocaml/fstar-lib/FStar_Parser_ParseIt.ml index 9bc10ea9f5f..554dac2d99d 100644 --- a/ocaml/fstar-lib/FStar_Parser_ParseIt.ml +++ b/ocaml/fstar-lib/FStar_Parser_ParseIt.ml @@ -147,7 +147,7 @@ let parse fn = let contents_at = let lines = U.splitlines contents in fun (start_pos:Lexing.position) (end_pos:Lexing.position) -> - let suffix = FStar_Compiler_Util.nth_tail (Z.of_int (start_pos.pos_lnum - 1)) lines in + let suffix = FStar_Compiler_Util.nth_tail (Z.of_int (if start_pos.pos_lnum > 0 then start_pos.pos_lnum - 1 else 0)) lines in let text, _ = FStar_Compiler_Util.first_N (Z.of_int (end_pos.pos_lnum - start_pos.pos_lnum)) suffix in let range = range_of_positions start_pos end_pos in { range; From 78ee2d0ed5d63d04f3ee9b3672a7c69ec7b4e8a9 Mon Sep 17 00:00:00 2001 From: Nikhil Swamy Date: Tue, 21 Mar 2023 22:15:22 +0000 Subject: [PATCH 25/48] snap --- .../FStar_InteractiveHelpers_Base.ml | 4 +- .../FStar_InteractiveHelpers_Effectful.ml | 21 +- .../FStar_InteractiveHelpers_ExploreTerm.ml | 265 +- .../FStar_InteractiveHelpers_PostProcess.ml | 109 +- .../generated/FStar_Interactive_Ide_Types.ml | 2 +- ocaml/fstar-lib/generated/FStar_Tactics_BV.ml | 8 +- .../generated/FStar_Tactics_Canon.ml | 8 +- .../FStar_Tactics_CanonCommMonoid.ml | 68 +- .../FStar_Tactics_CanonCommMonoidSimple.ml | 188 +- ...tar_Tactics_CanonCommMonoidSimple_Equiv.ml | 99 +- .../FStar_Tactics_CanonCommSemiring.ml | 56 +- .../generated/FStar_Tactics_CanonMonoid.ml | 148 +- .../generated/FStar_Tactics_Derived.ml | 1062 +++---- .../generated/FStar_Tactics_Effect.ml | 4 +- .../generated/FStar_Tactics_Logic.ml | 225 +- .../FStar_Tactics_PatternMatching.ml | 678 +++-- .../generated/FStar_Tactics_Simplifier.ml | 324 +- .../generated/FStar_Tactics_SyntaxHelpers.ml | 6 +- .../generated/FStar_Tactics_Typeclasses.ml | 73 +- .../generated/Steel_Effect_Common.ml | 2686 +++++++++-------- .../generated/Steel_ST_GenElim_Base.ml | 498 ++- 21 files changed, 3389 insertions(+), 3143 deletions(-) diff --git a/ocaml/fstar-lib/generated/FStar_InteractiveHelpers_Base.ml b/ocaml/fstar-lib/generated/FStar_InteractiveHelpers_Base.ml index bea6412242a..e8b1b42a385 100644 --- a/ocaml/fstar-lib/generated/FStar_InteractiveHelpers_Base.ml +++ b/ocaml/fstar-lib/generated/FStar_InteractiveHelpers_Base.ml @@ -946,7 +946,9 @@ let (__proj__MetaAnalysis__item__uu___ : Prims.exn -> Prims.string) = fun projectee -> match projectee with | MetaAnalysis uu___ -> uu___ let mfail : 'uuuuu . Prims.string -> ('uuuuu, unit) FStar_Tactics_Effect.tac_repr = - fun str -> FStar_Tactics_Effect.raise (MetaAnalysis str) + fun uu___ -> + (fun str -> Obj.magic (FStar_Tactics_Effect.raise (MetaAnalysis str))) + uu___ let (print_dbg : Prims.bool -> Prims.string -> (unit, unit) FStar_Tactics_Effect.tac_repr) = fun uu___1 -> diff --git a/ocaml/fstar-lib/generated/FStar_InteractiveHelpers_Effectful.ml b/ocaml/fstar-lib/generated/FStar_InteractiveHelpers_Effectful.ml index d113c2ed518..b6e37006c5d 100644 --- a/ocaml/fstar-lib/generated/FStar_InteractiveHelpers_Effectful.ml +++ b/ocaml/fstar-lib/generated/FStar_InteractiveHelpers_Effectful.ml @@ -1431,13 +1431,20 @@ let (compute_eterm_info : parameters)))) uu___2))) uu___2)) (fun uu___1 -> - match uu___1 with - | FStar_Tactics_Common.TacticFailure msg -> - FStar_InteractiveHelpers_Base.mfail - (Prims.strcat - "compute_eterm_info: failure: '" - (Prims.strcat msg "'")) - | e1 -> FStar_Tactics_Effect.raise e1))) uu___) + (fun uu___1 -> + match uu___1 with + | FStar_Tactics_Common.TacticFailure msg -> + Obj.magic + (Obj.repr + (FStar_InteractiveHelpers_Base.mfail + (Prims.strcat + "compute_eterm_info: failure: '" + (Prims.strcat msg "'")))) + | e1 -> + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.raise e1))) + uu___1))) uu___) let (has_refinement : FStar_InteractiveHelpers_ExploreTerm.type_info -> Prims.bool) = fun ty -> diff --git a/ocaml/fstar-lib/generated/FStar_InteractiveHelpers_ExploreTerm.ml b/ocaml/fstar-lib/generated/FStar_InteractiveHelpers_ExploreTerm.ml index 7428ba9e9f6..fa2b8a6733f 100644 --- a/ocaml/fstar-lib/generated/FStar_InteractiveHelpers_ExploreTerm.ml +++ b/ocaml/fstar-lib/generated/FStar_InteractiveHelpers_ExploreTerm.ml @@ -1556,13 +1556,21 @@ let rec (inst_comp : match () with | () -> inst_comp_once e c t) (fun uu___ -> - match uu___ with - | FStar_InteractiveHelpers_Base.MetaAnalysis - msg -> - FStar_InteractiveHelpers_Base.mfail - (Prims.strcat "inst_comp: error: " - msg) - | err -> FStar_Tactics_Effect.raise err))) + (fun uu___ -> + match uu___ with + | FStar_InteractiveHelpers_Base.MetaAnalysis + msg -> + Obj.magic + (Obj.repr + (FStar_InteractiveHelpers_Base.mfail + (Prims.strcat + "inst_comp: error: " + msg))) + | err -> + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.raise + err))) uu___))) (fun uu___ -> (fun c' -> Obj.magic (inst_comp e c' tl')) uu___)))) uu___2 uu___1 uu___ @@ -1661,54 +1669,68 @@ let (_abs_update_typ : "_abs_update_typ: inconsistent state")) uu___1))) uu___1)) (fun uu___ -> - match uu___ with - | FStar_InteractiveHelpers_Base.MetaAnalysis msg -> - FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.InteractiveHelpers.ExploreTerm.fst" - (Prims.of_int (303)) (Prims.of_int (10)) - (Prims.of_int (303)) (Prims.of_int (93))) - (Prims.mk_range - "FStar.InteractiveHelpers.ExploreTerm.fst" - (Prims.of_int (303)) (Prims.of_int (4)) - (Prims.of_int (303)) (Prims.of_int (93))) - (Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.InteractiveHelpers.ExploreTerm.fst" - (Prims.of_int (303)) (Prims.of_int (61)) - (Prims.of_int (303)) (Prims.of_int (92))) - (Prims.mk_range "prims.fst" (Prims.of_int (606)) - (Prims.of_int (19)) (Prims.of_int (606)) - (Prims.of_int (31))) - (Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.InteractiveHelpers.ExploreTerm.fst" - (Prims.of_int (303)) (Prims.of_int (61)) - (Prims.of_int (303)) (Prims.of_int (78))) - (Prims.mk_range "prims.fst" - (Prims.of_int (606)) (Prims.of_int (19)) - (Prims.of_int (606)) (Prims.of_int (31))) - (Obj.magic - (FStar_Tactics_Builtins.term_to_string ty)) + (fun uu___ -> + match uu___ with + | FStar_InteractiveHelpers_Base.MetaAnalysis msg -> + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.InteractiveHelpers.ExploreTerm.fst" + (Prims.of_int (303)) (Prims.of_int (10)) + (Prims.of_int (303)) (Prims.of_int (93))) + (Prims.mk_range + "FStar.InteractiveHelpers.ExploreTerm.fst" + (Prims.of_int (303)) (Prims.of_int (4)) + (Prims.of_int (303)) (Prims.of_int (93))) + (Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.InteractiveHelpers.ExploreTerm.fst" + (Prims.of_int (303)) + (Prims.of_int (61)) + (Prims.of_int (303)) + (Prims.of_int (92))) + (Prims.mk_range "prims.fst" + (Prims.of_int (606)) + (Prims.of_int (19)) + (Prims.of_int (606)) + (Prims.of_int (31))) + (Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.InteractiveHelpers.ExploreTerm.fst" + (Prims.of_int (303)) + (Prims.of_int (61)) + (Prims.of_int (303)) + (Prims.of_int (78))) + (Prims.mk_range "prims.fst" + (Prims.of_int (606)) + (Prims.of_int (19)) + (Prims.of_int (606)) + (Prims.of_int (31))) + (Obj.magic + (FStar_Tactics_Builtins.term_to_string + ty)) + (fun uu___1 -> + FStar_Tactics_Effect.lift_div_tac + (fun uu___2 -> + Prims.strcat uu___1 + (Prims.strcat ":\n" msg))))) + (fun uu___1 -> + FStar_Tactics_Effect.lift_div_tac + (fun uu___2 -> + Prims.strcat + "_abs_update_typ: could not find an arrow in: " + uu___1)))) + (fun uu___1 -> (fun uu___1 -> - FStar_Tactics_Effect.lift_div_tac - (fun uu___2 -> - Prims.strcat uu___1 - (Prims.strcat ":\n" msg))))) - (fun uu___1 -> - FStar_Tactics_Effect.lift_div_tac - (fun uu___2 -> - Prims.strcat - "_abs_update_typ: could not find an arrow in: " - uu___1)))) - (fun uu___1 -> - (fun uu___1 -> - Obj.magic - (FStar_InteractiveHelpers_Base.mfail uu___1)) - uu___1) - | err -> FStar_Tactics_Effect.raise err) + Obj.magic + (FStar_InteractiveHelpers_Base.mfail + uu___1)) uu___1))) + | err -> + Obj.magic (Obj.repr (FStar_Tactics_Effect.raise err))) + uu___) let (abs_update_typ_or_comp : FStar_Reflection_Types.binder -> typ_or_comp -> @@ -1780,14 +1802,12 @@ let (abs_update_opt_typ_or_comp : | FStar_InteractiveHelpers_Base.MetaAnalysis msg -> Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___1 -> - FStar_Pervasives_Native.None))) + (FStar_Tactics_Effect.lift_div_tac + (fun uu___1 -> + FStar_Pervasives_Native.None)) | err -> Obj.magic - (Obj.repr - (FStar_Tactics_Effect.raise err))) + (FStar_Tactics_Effect.raise err)) uu___)))) uu___2 uu___1 uu___ let rec (_flush_typ_or_comp_comp : Prims.bool -> @@ -2315,62 +2335,77 @@ let (flush_typ_or_comp : uu___1) | TC_Comp (c, pl, n) -> flush_comp pl n c)) (fun uu___ -> - match uu___ with - | FStar_InteractiveHelpers_Base.MetaAnalysis msg -> - FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.InteractiveHelpers.ExploreTerm.fst" - (Prims.of_int (379)) (Prims.of_int (15)) - (Prims.of_int (379)) (Prims.of_int (90))) - (Prims.mk_range - "FStar.InteractiveHelpers.ExploreTerm.fst" - (Prims.of_int (379)) (Prims.of_int (9)) - (Prims.of_int (379)) (Prims.of_int (90))) - (Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.InteractiveHelpers.ExploreTerm.fst" - (Prims.of_int (379)) - (Prims.of_int (50)) - (Prims.of_int (379)) - (Prims.of_int (89))) - (Prims.mk_range "prims.fst" - (Prims.of_int (606)) - (Prims.of_int (19)) - (Prims.of_int (606)) - (Prims.of_int (31))) - (Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.InteractiveHelpers.ExploreTerm.fst" - (Prims.of_int (379)) - (Prims.of_int (50)) - (Prims.of_int (379)) - (Prims.of_int (75))) - (Prims.mk_range "prims.fst" - (Prims.of_int (606)) - (Prims.of_int (19)) - (Prims.of_int (606)) - (Prims.of_int (31))) - (Obj.magic - (typ_or_comp_to_string tyc)) + (fun uu___ -> + match uu___ with + | FStar_InteractiveHelpers_Base.MetaAnalysis msg + -> + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.InteractiveHelpers.ExploreTerm.fst" + (Prims.of_int (379)) + (Prims.of_int (15)) + (Prims.of_int (379)) + (Prims.of_int (90))) + (Prims.mk_range + "FStar.InteractiveHelpers.ExploreTerm.fst" + (Prims.of_int (379)) + (Prims.of_int (9)) + (Prims.of_int (379)) + (Prims.of_int (90))) + (Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.InteractiveHelpers.ExploreTerm.fst" + (Prims.of_int (379)) + (Prims.of_int (50)) + (Prims.of_int (379)) + (Prims.of_int (89))) + (Prims.mk_range "prims.fst" + (Prims.of_int (606)) + (Prims.of_int (19)) + (Prims.of_int (606)) + (Prims.of_int (31))) + (Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.InteractiveHelpers.ExploreTerm.fst" + (Prims.of_int (379)) + (Prims.of_int (50)) + (Prims.of_int (379)) + (Prims.of_int (75))) + (Prims.mk_range + "prims.fst" + (Prims.of_int (606)) + (Prims.of_int (19)) + (Prims.of_int (606)) + (Prims.of_int (31))) + (Obj.magic + (typ_or_comp_to_string + tyc)) + (fun uu___1 -> + FStar_Tactics_Effect.lift_div_tac + (fun uu___2 -> + Prims.strcat + uu___1 + (Prims.strcat + ":\n" msg))))) + (fun uu___1 -> + FStar_Tactics_Effect.lift_div_tac + (fun uu___2 -> + Prims.strcat + "flush_typ_or_comp failed on: " + uu___1)))) + (fun uu___1 -> (fun uu___1 -> - FStar_Tactics_Effect.lift_div_tac - (fun uu___2 -> - Prims.strcat uu___1 - (Prims.strcat ":\n" msg))))) - (fun uu___1 -> - FStar_Tactics_Effect.lift_div_tac - (fun uu___2 -> - Prims.strcat - "flush_typ_or_comp failed on: " - uu___1)))) - (fun uu___1 -> - (fun uu___1 -> - Obj.magic - (FStar_InteractiveHelpers_Base.mfail - uu___1)) uu___1) - | err -> FStar_Tactics_Effect.raise err))) uu___) + Obj.magic + (FStar_InteractiveHelpers_Base.mfail + uu___1)) uu___1))) + | err -> + Obj.magic + (Obj.repr (FStar_Tactics_Effect.raise err))) + uu___))) uu___) let (safe_arg_typ_or_comp : Prims.bool -> FStar_Reflection_Types.env -> diff --git a/ocaml/fstar-lib/generated/FStar_InteractiveHelpers_PostProcess.ml b/ocaml/fstar-lib/generated/FStar_InteractiveHelpers_PostProcess.ml index a3e884e6b41..a2430d7f851 100644 --- a/ocaml/fstar-lib/generated/FStar_InteractiveHelpers_PostProcess.ml +++ b/ocaml/fstar-lib/generated/FStar_InteractiveHelpers_PostProcess.ml @@ -2146,22 +2146,29 @@ let (pp_analyze_effectful_term : (fun uu___2 -> Obj.magic (end_proof ())) uu___2))) uu___2)) (fun uu___1 -> - match uu___1 with - | FStar_InteractiveHelpers_Base.MetaAnalysis msg -> - FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.InteractiveHelpers.PostProcess.fst" - (Prims.of_int (256)) (Prims.of_int (29)) - (Prims.of_int (256)) (Prims.of_int (49))) - (Prims.mk_range - "FStar.InteractiveHelpers.PostProcess.fst" - (Prims.of_int (256)) (Prims.of_int (51)) - (Prims.of_int (256)) (Prims.of_int (63))) - (Obj.magic - (FStar_InteractiveHelpers_Output.printout_failure msg)) - (fun uu___2 -> - (fun uu___2 -> Obj.magic (end_proof ())) uu___2) - | err -> FStar_Tactics_Effect.raise err) + (fun uu___1 -> + match uu___1 with + | FStar_InteractiveHelpers_Base.MetaAnalysis msg -> + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.InteractiveHelpers.PostProcess.fst" + (Prims.of_int (256)) (Prims.of_int (29)) + (Prims.of_int (256)) (Prims.of_int (49))) + (Prims.mk_range + "FStar.InteractiveHelpers.PostProcess.fst" + (Prims.of_int (256)) (Prims.of_int (51)) + (Prims.of_int (256)) (Prims.of_int (63))) + (Obj.magic + (FStar_InteractiveHelpers_Output.printout_failure + msg)) + (fun uu___2 -> + (fun uu___2 -> Obj.magic (end_proof ())) + uu___2))) + | err -> + Obj.magic (Obj.repr (FStar_Tactics_Effect.raise err))) + uu___1) let _ = FStar_Tactics_Native.register_tactic "FStar.InteractiveHelpers.PostProcess.pp_analyze_effectful_term" @@ -2605,20 +2612,27 @@ let (pp_split_assert_conjs : (fun uu___2 -> Obj.magic (end_proof ())) uu___2))) uu___2)) (fun uu___1 -> - match uu___1 with - | FStar_InteractiveHelpers_Base.MetaAnalysis msg -> - FStar_Tactics_Effect.tac_bind - (Prims.mk_range "FStar.InteractiveHelpers.PostProcess.fst" - (Prims.of_int (356)) (Prims.of_int (29)) - (Prims.of_int (356)) (Prims.of_int (49))) - (Prims.mk_range "FStar.InteractiveHelpers.PostProcess.fst" - (Prims.of_int (356)) (Prims.of_int (51)) - (Prims.of_int (356)) (Prims.of_int (63))) - (Obj.magic - (FStar_InteractiveHelpers_Output.printout_failure msg)) - (fun uu___2 -> - (fun uu___2 -> Obj.magic (end_proof ())) uu___2) - | err -> FStar_Tactics_Effect.raise err) + (fun uu___1 -> + match uu___1 with + | FStar_InteractiveHelpers_Base.MetaAnalysis msg -> + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.InteractiveHelpers.PostProcess.fst" + (Prims.of_int (356)) (Prims.of_int (29)) + (Prims.of_int (356)) (Prims.of_int (49))) + (Prims.mk_range + "FStar.InteractiveHelpers.PostProcess.fst" + (Prims.of_int (356)) (Prims.of_int (51)) + (Prims.of_int (356)) (Prims.of_int (63))) + (Obj.magic + (FStar_InteractiveHelpers_Output.printout_failure + msg)) + (fun uu___2 -> + (fun uu___2 -> Obj.magic (end_proof ())) uu___2))) + | err -> Obj.magic (Obj.repr (FStar_Tactics_Effect.raise err))) + uu___1) let _ = FStar_Tactics_Native.register_tactic "FStar.InteractiveHelpers.PostProcess.pp_split_assert_conjs" @@ -6913,20 +6927,27 @@ let (pp_unfold_in_assert_or_assume : (fun uu___2 -> Obj.magic (end_proof ())) uu___2))) uu___2)) (fun uu___1 -> - match uu___1 with - | FStar_InteractiveHelpers_Base.MetaAnalysis msg -> - FStar_Tactics_Effect.tac_bind - (Prims.mk_range "FStar.InteractiveHelpers.PostProcess.fst" - (Prims.of_int (748)) (Prims.of_int (29)) - (Prims.of_int (748)) (Prims.of_int (49))) - (Prims.mk_range "FStar.InteractiveHelpers.PostProcess.fst" - (Prims.of_int (748)) (Prims.of_int (51)) - (Prims.of_int (748)) (Prims.of_int (63))) - (Obj.magic - (FStar_InteractiveHelpers_Output.printout_failure msg)) - (fun uu___2 -> - (fun uu___2 -> Obj.magic (end_proof ())) uu___2) - | err -> FStar_Tactics_Effect.raise err) + (fun uu___1 -> + match uu___1 with + | FStar_InteractiveHelpers_Base.MetaAnalysis msg -> + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.InteractiveHelpers.PostProcess.fst" + (Prims.of_int (748)) (Prims.of_int (29)) + (Prims.of_int (748)) (Prims.of_int (49))) + (Prims.mk_range + "FStar.InteractiveHelpers.PostProcess.fst" + (Prims.of_int (748)) (Prims.of_int (51)) + (Prims.of_int (748)) (Prims.of_int (63))) + (Obj.magic + (FStar_InteractiveHelpers_Output.printout_failure + msg)) + (fun uu___2 -> + (fun uu___2 -> Obj.magic (end_proof ())) uu___2))) + | err -> Obj.magic (Obj.repr (FStar_Tactics_Effect.raise err))) + uu___1) let _ = FStar_Tactics_Native.register_tactic "FStar.InteractiveHelpers.PostProcess.pp_unfold_in_assert_or_assume" diff --git a/ocaml/fstar-lib/generated/FStar_Interactive_Ide_Types.ml b/ocaml/fstar-lib/generated/FStar_Interactive_Ide_Types.ml index 0a2ba4e0a3b..295c26eb7e2 100644 --- a/ocaml/fstar-lib/generated/FStar_Interactive_Ide_Types.ml +++ b/ocaml/fstar-lib/generated/FStar_Interactive_Ide_Types.ml @@ -437,7 +437,7 @@ let (repl_state_to_string : repl_state -> Prims.string) = uu___3 :: uu___4 in uu___1 :: uu___2 in FStar_Compiler_Util.format - "{\n\trepl_line=%s;\n\trepl_column=%s;\n\trepl_fname=%s;\n\trepl_cur_mod=%s;\n\t\\ \r\n repl_deps_stack={%s}\n}" + "{\n\trepl_line=%s;\n\trepl_column=%s;\n\trepl_fname=%s;\n\trepl_cur_mod=%s;\n\t\\ \n repl_deps_stack={%s}\n}" uu___ let (push_query_to_string : push_query -> Prims.string) = fun pq -> diff --git a/ocaml/fstar-lib/generated/FStar_Tactics_BV.ml b/ocaml/fstar-lib/generated/FStar_Tactics_BV.ml index cc7174ea77a..496adf26e40 100644 --- a/ocaml/fstar-lib/generated/FStar_Tactics_BV.ml +++ b/ocaml/fstar-lib/generated/FStar_Tactics_BV.ml @@ -892,11 +892,9 @@ let (arith_to_bv_tac : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = "arith_to_bv_tac: unexpected: " uu___4)))) (fun uu___4 -> - (fun uu___4 -> - Obj.magic - (FStar_Tactics_Derived.fail - uu___4)) uu___4))) - uu___3))) uu___3))) uu___2)) + FStar_Tactics_Derived.fail + uu___4))) uu___3))) + uu___3))) uu___2)) let (bv_tac : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = fun uu___ -> FStar_Tactics_Derived.focus diff --git a/ocaml/fstar-lib/generated/FStar_Tactics_Canon.ml b/ocaml/fstar-lib/generated/FStar_Tactics_Canon.ml index 93e57a67c65..ef14886be48 100644 --- a/ocaml/fstar-lib/generated/FStar_Tactics_Canon.ml +++ b/ocaml/fstar-lib/generated/FStar_Tactics_Canon.ml @@ -1028,10 +1028,8 @@ let (canon_point_entry : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) "impossible: " uu___4)))) (fun uu___4 -> - (fun uu___4 -> - Obj.magic - (FStar_Tactics_Derived.fail - uu___4)) uu___4))) - uu___2))) uu___2))) uu___1) + FStar_Tactics_Derived.fail + uu___4))) uu___2))) uu___2))) + uu___1) let (canon : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = fun uu___ -> FStar_Tactics_Derived.pointwise canon_point_entry \ No newline at end of file diff --git a/ocaml/fstar-lib/generated/FStar_Tactics_CanonCommMonoid.ml b/ocaml/fstar-lib/generated/FStar_Tactics_CanonCommMonoid.ml index d71fd4339c0..a6482cacc8d 100644 --- a/ocaml/fstar-lib/generated/FStar_Tactics_CanonCommMonoid.ml +++ b/ocaml/fstar-lib/generated/FStar_Tactics_CanonCommMonoid.ml @@ -1669,28 +1669,30 @@ let canon_monoid_aux : t), t1, t2) -> Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.CanonCommMonoid.fst" - (Prims.of_int (340)) - (Prims.of_int (9)) - (Prims.of_int (340)) - (Prims.of_int (25))) - (Prims.mk_range - "FStar.Tactics.CanonCommMonoid.fst" - (Prims.of_int (340)) - (Prims.of_int (6)) - (Prims.of_int (411)) - (Prims.of_int (69))) - (Obj.magic - (FStar_Tactics_Builtins.term_eq_old - t ta)) - (fun uu___2 -> + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.CanonCommMonoid.fst" + (Prims.of_int (340)) + (Prims.of_int (9)) + (Prims.of_int (340)) + (Prims.of_int (25))) + (Prims.mk_range + "FStar.Tactics.CanonCommMonoid.fst" + (Prims.of_int (340)) + (Prims.of_int (6)) + (Prims.of_int (411)) + (Prims.of_int (69))) + (Obj.magic + (FStar_Tactics_Builtins.term_eq_old + t ta)) (fun uu___2 -> - if uu___2 - then - Obj.magic - (FStar_Tactics_Effect.tac_bind + (fun uu___2 -> + if uu___2 + then + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.CanonCommMonoid.fst" (Prims.of_int (341)) @@ -1731,6 +1733,7 @@ let canon_monoid_aux : (r1::r2::[], vm) -> Obj.magic + (Obj.repr (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.CanonCommMonoid.fst" @@ -1997,22 +2000,25 @@ let canon_monoid_aux : uu___4))) uu___4))) uu___4))) - uu___4)) + uu___4))) | uu___4 -> Obj.magic + (Obj.repr (FStar_Tactics_Derived.fail - "Unexpected")) - uu___3)) - else - Obj.magic - (FStar_Tactics_Derived.fail - "Goal should be an equality at the right monoid type")) - uu___2)) + "Unexpected"))) + uu___3))) + else + Obj.magic + (Obj.repr + (FStar_Tactics_Derived.fail + "Goal should be an equality at the right monoid type"))) + uu___2))) | uu___2 -> Obj.magic - (FStar_Tactics_Derived.fail - "Goal should be an equality")) + (Obj.repr + (FStar_Tactics_Derived.fail + "Goal should be an equality"))) uu___1))) uu___) let canon_monoid_with : 'b . diff --git a/ocaml/fstar-lib/generated/FStar_Tactics_CanonCommMonoidSimple.ml b/ocaml/fstar-lib/generated/FStar_Tactics_CanonCommMonoidSimple.ml index d4730d05f43..16012c10b64 100644 --- a/ocaml/fstar-lib/generated/FStar_Tactics_CanonCommMonoidSimple.ml +++ b/ocaml/fstar-lib/generated/FStar_Tactics_CanonCommMonoidSimple.ml @@ -624,89 +624,97 @@ let canon_monoid : (FStar_Pervasives_Native.Some t), t1, t2) -> Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.CanonCommMonoidSimple.fst" - (Prims.of_int (255)) (Prims.of_int (9)) - (Prims.of_int (255)) (Prims.of_int (28))) - (Prims.mk_range - "FStar.Tactics.CanonCommMonoidSimple.fst" - (Prims.of_int (255)) (Prims.of_int (6)) - (Prims.of_int (273)) (Prims.of_int (69))) - (Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.CanonCommMonoidSimple.fst" - (Prims.of_int (255)) - (Prims.of_int (19)) - (Prims.of_int (255)) - (Prims.of_int (28))) - (Prims.mk_range - "FStar.Tactics.CanonCommMonoidSimple.fst" - (Prims.of_int (255)) - (Prims.of_int (9)) - (Prims.of_int (255)) - (Prims.of_int (28))) - (FStar_Tactics_Effect.lift_div_tac - (fun uu___2 -> + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.CanonCommMonoidSimple.fst" + (Prims.of_int (255)) (Prims.of_int (9)) + (Prims.of_int (255)) + (Prims.of_int (28))) + (Prims.mk_range + "FStar.Tactics.CanonCommMonoidSimple.fst" + (Prims.of_int (255)) (Prims.of_int (6)) + (Prims.of_int (273)) + (Prims.of_int (69))) + (Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.CanonCommMonoidSimple.fst" + (Prims.of_int (255)) + (Prims.of_int (19)) + (Prims.of_int (255)) + (Prims.of_int (28))) + (Prims.mk_range + "FStar.Tactics.CanonCommMonoidSimple.fst" + (Prims.of_int (255)) + (Prims.of_int (9)) + (Prims.of_int (255)) + (Prims.of_int (28))) + (FStar_Tactics_Effect.lift_div_tac (fun uu___2 -> - Obj.magic - (failwith - "Cannot evaluate open quotation at runtime")) - uu___2)) - (fun uu___2 -> + (fun uu___2 -> + Obj.magic + (failwith + "Cannot evaluate open quotation at runtime")) + uu___2)) (fun uu___2 -> - Obj.magic (term_eq t uu___2)) - uu___2))) - (fun uu___2 -> + (fun uu___2 -> + Obj.magic (term_eq t uu___2)) + uu___2))) (fun uu___2 -> - if uu___2 - then - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.CanonCommMonoidSimple.fst" - (Prims.of_int (256)) - (Prims.of_int (27)) - (Prims.of_int (256)) - (Prims.of_int (67))) - (Prims.mk_range - "FStar.Tactics.CanonCommMonoidSimple.fst" - (Prims.of_int (256)) - (Prims.of_int (8)) - (Prims.of_int (271)) - (Prims.of_int (22))) - (Obj.magic - (reification m [] - (const - (FStar_Algebra_CommMonoid.__proj__CM__item__unit - m)) t1)) - (fun uu___3 -> - (fun uu___3 -> - match uu___3 with - | (r1, ts, am) -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.CanonCommMonoidSimple.fst" - (Prims.of_int (257)) - (Prims.of_int (26)) - (Prims.of_int (257)) - (Prims.of_int (48))) - (Prims.mk_range - "FStar.Tactics.CanonCommMonoidSimple.fst" - (Prims.of_int (257)) - (Prims.of_int (8)) - (Prims.of_int (271)) - (Prims.of_int (22))) - (Obj.magic - (reification m - ts am t2)) - (fun uu___4 -> - (fun uu___4 -> - match uu___4 - with - | (r2, + (fun uu___2 -> + if uu___2 + then + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.CanonCommMonoidSimple.fst" + (Prims.of_int (256)) + (Prims.of_int (27)) + (Prims.of_int (256)) + (Prims.of_int (67))) + (Prims.mk_range + "FStar.Tactics.CanonCommMonoidSimple.fst" + (Prims.of_int (256)) + (Prims.of_int (8)) + (Prims.of_int (271)) + (Prims.of_int (22))) + (Obj.magic + (reification m [] + (const + (FStar_Algebra_CommMonoid.__proj__CM__item__unit + m)) t1)) + (fun uu___3 -> + (fun uu___3 -> + match uu___3 with + | (r1, ts, am) -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.CanonCommMonoidSimple.fst" + (Prims.of_int (257)) + (Prims.of_int (26)) + (Prims.of_int (257)) + (Prims.of_int (48))) + (Prims.mk_range + "FStar.Tactics.CanonCommMonoidSimple.fst" + (Prims.of_int (257)) + (Prims.of_int (8)) + (Prims.of_int (271)) + (Prims.of_int (22))) + (Obj.magic + (reification + m ts am + t2)) + (fun uu___4 + -> + (fun + uu___4 -> + match uu___4 + with + | + (r2, uu___5, am1) -> Obj.magic @@ -900,15 +908,17 @@ let canon_monoid : uu___8))) uu___7))) uu___6))) - uu___4))) - uu___3)) - else - Obj.magic - (FStar_Tactics_Derived.fail - "Goal should be an equality at the right monoid type")) - uu___2)) + uu___4))) + uu___3))) + else + Obj.magic + (Obj.repr + (FStar_Tactics_Derived.fail + "Goal should be an equality at the right monoid type"))) + uu___2))) | uu___2 -> Obj.magic - (FStar_Tactics_Derived.fail - "Goal should be an equality")) uu___1))) + (Obj.repr + (FStar_Tactics_Derived.fail + "Goal should be an equality"))) uu___1))) uu___) \ No newline at end of file diff --git a/ocaml/fstar-lib/generated/FStar_Tactics_CanonCommMonoidSimple_Equiv.ml b/ocaml/fstar-lib/generated/FStar_Tactics_CanonCommMonoidSimple_Equiv.ml index 6bd65b0452a..5d9bec7196b 100644 --- a/ocaml/fstar-lib/generated/FStar_Tactics_CanonCommMonoidSimple_Equiv.ml +++ b/ocaml/fstar-lib/generated/FStar_Tactics_CanonCommMonoidSimple_Equiv.ml @@ -1075,67 +1075,74 @@ let (canon_monoid : (match rel_xy with | (rel_xy1, uu___2)::[] -> Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.CanonCommMonoidSimple.Equiv.fst" - (Prims.of_int (381)) - (Prims.of_int (21)) - (Prims.of_int (381)) - (Prims.of_int (43))) - (Prims.mk_range - "FStar.Tactics.CanonCommMonoidSimple.Equiv.fst" - (Prims.of_int (380)) - (Prims.of_int (21)) - (Prims.of_int (391)) - (Prims.of_int (6))) - (FStar_Tactics_Effect.lift_div_tac + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.CanonCommMonoidSimple.Equiv.fst" + (Prims.of_int (381)) + (Prims.of_int (21)) + (Prims.of_int (381)) + (Prims.of_int (43))) + (Prims.mk_range + "FStar.Tactics.CanonCommMonoidSimple.Equiv.fst" + (Prims.of_int (380)) + (Prims.of_int (21)) + (Prims.of_int (391)) + (Prims.of_int (6))) + (FStar_Tactics_Effect.lift_div_tac + (fun uu___3 -> + FStar_Reflection_Derived_Lemmas.collect_app_ref + rel_xy1)) (fun uu___3 -> - FStar_Reflection_Derived_Lemmas.collect_app_ref - rel_xy1)) - (fun uu___3 -> - (fun uu___3 -> - match uu___3 with - | (rel, xy) -> - if - (FStar_List_Tot_Base.length - xy) - >= - (Prims.of_int (2)) - then - (match - ((FStar_List_Tot_Base.index + (fun uu___3 -> + match uu___3 with + | (rel, xy) -> + if + (FStar_List_Tot_Base.length + xy) + >= + (Prims.of_int (2)) + then + Obj.magic + (Obj.repr + (match + ((FStar_List_Tot_Base.index xy ((FStar_List_Tot_Base.length xy) - (Prims.of_int (2)))), - (FStar_List_Tot_Base.index + (FStar_List_Tot_Base.index xy ((FStar_List_Tot_Base.length xy) - Prims.int_one))) - with - | ((lhs, - FStar_Reflection_Data.Q_Explicit), - (rhs, - FStar_Reflection_Data.Q_Explicit)) - -> - Obj.magic + with + | + ((lhs, + FStar_Reflection_Data.Q_Explicit), + (rhs, + FStar_Reflection_Data.Q_Explicit)) + -> + Obj.repr (canon_lhs_rhs eq m lhs rhs) - | uu___4 -> - Obj.magic + | + uu___4 -> + Obj.repr + (FStar_Tactics_Derived.fail + "Goal should have been an application of a binary relation to 2 explicit arguments"))) + else + Obj.magic + (Obj.repr (FStar_Tactics_Derived.fail - "Goal should have been an application of a binary relation to 2 explicit arguments")) - else - Obj.magic - (FStar_Tactics_Derived.fail - "Goal should have been an application of a binary relation to n implicit and 2 explicit arguments")) - uu___3)) + "Goal should have been an application of a binary relation to n implicit and 2 explicit arguments"))) + uu___3))) | uu___2 -> Obj.magic - (FStar_Tactics_Derived.fail - "Goal should be squash applied to a binary relation"))) + (Obj.repr + (FStar_Tactics_Derived.fail + "Goal should be squash applied to a binary relation")))) uu___1))) uu___1))) uu___) let _ = FStar_Tactics_Native.register_tactic diff --git a/ocaml/fstar-lib/generated/FStar_Tactics_CanonCommSemiring.ml b/ocaml/fstar-lib/generated/FStar_Tactics_CanonCommSemiring.ml index a7843079ede..547fce6978e 100644 --- a/ocaml/fstar-lib/generated/FStar_Tactics_CanonCommSemiring.ml +++ b/ocaml/fstar-lib/generated/FStar_Tactics_CanonCommSemiring.ml @@ -2136,29 +2136,31 @@ let canon_semiring_aux : t), t1, t2) -> Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.CanonCommSemiring.fst" - (Prims.of_int (1634)) - (Prims.of_int (9)) - (Prims.of_int (1634)) - (Prims.of_int (21))) - (Prims.mk_range - "FStar.Tactics.CanonCommSemiring.fst" - (Prims.of_int (1634)) - (Prims.of_int (6)) - (Prims.of_int (1671)) - (Prims.of_int (73))) - (Obj.magic - (term_eq t - ta)) - (fun uu___3 - -> + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.CanonCommSemiring.fst" + (Prims.of_int (1634)) + (Prims.of_int (9)) + (Prims.of_int (1634)) + (Prims.of_int (21))) + (Prims.mk_range + "FStar.Tactics.CanonCommSemiring.fst" + (Prims.of_int (1634)) + (Prims.of_int (6)) + (Prims.of_int (1671)) + (Prims.of_int (73))) + (Obj.magic + (term_eq + t ta)) (fun + uu___3 -> + (fun uu___3 -> if uu___3 then Obj.magic + (Obj.repr (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.CanonCommSemiring.fst" @@ -2191,6 +2193,7 @@ let canon_semiring_aux : (e1::e2::[], vm) -> Obj.magic + (Obj.repr (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.CanonCommSemiring.fst" @@ -2429,22 +2432,25 @@ let canon_semiring_aux : uu___5))) uu___5))) uu___5))) - uu___5)) + uu___5))) | uu___5 -> Obj.magic + (Obj.repr (FStar_Tactics_Derived.fail - "Unexpected")) - uu___4)) + "Unexpected"))) + uu___4))) else Obj.magic + (Obj.repr (FStar_Tactics_Derived.fail - "Found equality, but terms do not have the expected type")) - uu___3)) + "Found equality, but terms do not have the expected type"))) + uu___3))) | uu___3 -> Obj.magic - (FStar_Tactics_Derived.fail - "Goal should be an equality")) + (Obj.repr + (FStar_Tactics_Derived.fail + "Goal should be an equality"))) uu___2))) uu___2))) uu___1)) let canon_semiring : 'a . 'a cr -> (unit, unit) FStar_Tactics_Effect.tac_repr diff --git a/ocaml/fstar-lib/generated/FStar_Tactics_CanonMonoid.ml b/ocaml/fstar-lib/generated/FStar_Tactics_CanonMonoid.ml index 648100aa857..be918e0117b 100644 --- a/ocaml/fstar-lib/generated/FStar_Tactics_CanonMonoid.ml +++ b/ocaml/fstar-lib/generated/FStar_Tactics_CanonMonoid.ml @@ -448,85 +448,89 @@ let canon_monoid : me2) -> Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.CanonMonoid.fst" - (Prims.of_int (110)) - (Prims.of_int (9)) - (Prims.of_int (110)) - (Prims.of_int (32))) - (Prims.mk_range - "FStar.Tactics.CanonMonoid.fst" - (Prims.of_int (110)) - (Prims.of_int (6)) - (Prims.of_int (119)) - (Prims.of_int (69))) - (Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.CanonMonoid.fst" - (Prims.of_int (110)) - (Prims.of_int (23)) - (Prims.of_int (110)) - (Prims.of_int (32))) - (Prims.mk_range - "FStar.Tactics.CanonMonoid.fst" - (Prims.of_int (110)) - (Prims.of_int (9)) - (Prims.of_int (110)) - (Prims.of_int (32))) - (FStar_Tactics_Effect.lift_div_tac + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.CanonMonoid.fst" + (Prims.of_int (110)) + (Prims.of_int (9)) + (Prims.of_int (110)) + (Prims.of_int (32))) + (Prims.mk_range + "FStar.Tactics.CanonMonoid.fst" + (Prims.of_int (110)) + (Prims.of_int (6)) + (Prims.of_int (119)) + (Prims.of_int (69))) + (Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.CanonMonoid.fst" + (Prims.of_int (110)) + (Prims.of_int (23)) + (Prims.of_int (110)) + (Prims.of_int (32))) + (Prims.mk_range + "FStar.Tactics.CanonMonoid.fst" + (Prims.of_int (110)) + (Prims.of_int (9)) + (Prims.of_int (110)) + (Prims.of_int (32))) + (FStar_Tactics_Effect.lift_div_tac + (fun uu___2 -> + (fun uu___2 -> + Obj.magic + (failwith + "Cannot evaluate open quotation at runtime")) + uu___2)) (fun uu___2 -> (fun uu___2 -> Obj.magic - (failwith - "Cannot evaluate open quotation at runtime")) - uu___2)) - (fun uu___2 -> - (fun uu___2 -> - Obj.magic - (FStar_Tactics_Builtins.term_eq_old - t uu___2)) uu___2))) - (fun uu___2 -> + (FStar_Tactics_Builtins.term_eq_old + t uu___2)) + uu___2))) (fun uu___2 -> - if uu___2 - then - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.CanonMonoid.fst" - (Prims.of_int (111)) - (Prims.of_int (17)) - (Prims.of_int (111)) - (Prims.of_int (34))) - (Prims.mk_range - "FStar.Tactics.CanonMonoid.fst" - (Prims.of_int (112)) - (Prims.of_int (8)) - (Prims.of_int (118)) - (Prims.of_int (51))) - (Obj.magic - (reification m me1)) - (fun uu___3 -> - (fun r1 -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range + (fun uu___2 -> + if uu___2 + then + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.CanonMonoid.fst" + (Prims.of_int (111)) + (Prims.of_int (17)) + (Prims.of_int (111)) + (Prims.of_int (34))) + (Prims.mk_range + "FStar.Tactics.CanonMonoid.fst" + (Prims.of_int (112)) + (Prims.of_int (8)) + (Prims.of_int (118)) + (Prims.of_int (51))) + (Obj.magic + (reification + m me1)) + (fun uu___3 -> + (fun r1 -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "FStar.Tactics.CanonMonoid.fst" (Prims.of_int (112)) (Prims.of_int (17)) (Prims.of_int (112)) (Prims.of_int (34))) - (Prims.mk_range + (Prims.mk_range "FStar.Tactics.CanonMonoid.fst" (Prims.of_int (113)) (Prims.of_int (8)) (Prims.of_int (118)) (Prims.of_int (51))) - (Obj.magic + (Obj.magic (reification m me2)) - (fun + (fun uu___3 -> (fun r2 -> @@ -617,14 +621,16 @@ let canon_monoid : uu___4))) uu___3))) uu___3))) - uu___3)) - else - Obj.magic - (FStar_Tactics_Derived.fail - "Goal should be an equality at the right monoid type")) - uu___2)) + uu___3))) + else + Obj.magic + (Obj.repr + (FStar_Tactics_Derived.fail + "Goal should be an equality at the right monoid type"))) + uu___2))) | uu___2 -> Obj.magic - (FStar_Tactics_Derived.fail - "Goal should be an equality")) + (Obj.repr + (FStar_Tactics_Derived.fail + "Goal should be an equality"))) uu___1))) uu___1))) uu___) \ No newline at end of file diff --git a/ocaml/fstar-lib/generated/FStar_Tactics_Derived.ml b/ocaml/fstar-lib/generated/FStar_Tactics_Derived.ml index 4c8008fdce0..c8d79f468ed 100644 --- a/ocaml/fstar-lib/generated/FStar_Tactics_Derived.ml +++ b/ocaml/fstar-lib/generated/FStar_Tactics_Derived.ml @@ -36,7 +36,7 @@ let (goals : (Prims.of_int (42)) (Prims.of_int (40)) (Prims.of_int (50))) (Prims.mk_range "FStar.Tactics.Derived.fst" (Prims.of_int (40)) (Prims.of_int (33)) (Prims.of_int (40)) (Prims.of_int (50))) - (Obj.magic (FStar_Tactics_Effect.get ())) + (FStar_Tactics_Effect.get ()) (fun uu___1 -> FStar_Tactics_Effect.lift_div_tac (fun uu___2 -> FStar_Tactics_Types.goals_of uu___1)) @@ -50,11 +50,11 @@ let (smt_goals : (Prims.of_int (50)) (Prims.of_int (41)) (Prims.of_int (58))) (Prims.mk_range "FStar.Tactics.Derived.fst" (Prims.of_int (41)) (Prims.of_int (37)) (Prims.of_int (41)) (Prims.of_int (58))) - (Obj.magic (FStar_Tactics_Effect.get ())) + (FStar_Tactics_Effect.get ()) (fun uu___1 -> FStar_Tactics_Effect.lift_div_tac (fun uu___2 -> FStar_Tactics_Types.smt_goals_of uu___1)) -let fail : 'a . Prims.string -> ('a, unit) FStar_Tactics_Effect.tac_repr = +let fail : 'a . Prims.string -> ('a, Obj.t) FStar_Tactics_Effect.tac_repr = fun m -> FStar_Tactics_Effect.raise (FStar_Tactics_Common.TacticFailure m) let fail_silently : 'a . Prims.string -> ('a, unit) FStar_Tactics_Effect.tac_repr = @@ -66,10 +66,7 @@ let fail_silently : (Prims.of_int (4)) (Prims.of_int (50)) (Prims.of_int (30))) (Obj.magic (FStar_Tactics_Builtins.set_urgency Prims.int_zero)) (fun uu___ -> - (fun uu___ -> - Obj.magic - (FStar_Tactics_Effect.raise - (FStar_Tactics_Common.TacticFailure m))) uu___) + FStar_Tactics_Effect.raise (FStar_Tactics_Common.TacticFailure m)) let (_cur_goal : unit -> (FStar_Tactics_Types.goal, unit) FStar_Tactics_Effect.tac_repr) = fun uu___ -> @@ -80,14 +77,9 @@ let (_cur_goal : (Prims.of_int (4)) (Prims.of_int (56)) (Prims.of_int (15))) (Obj.magic (goals ())) (fun uu___1 -> - (fun uu___1 -> - match uu___1 with - | [] -> Obj.magic (Obj.repr (fail "no more goals")) - | g::uu___2 -> - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac (fun uu___3 -> g)))) - uu___1) + match uu___1 with + | [] -> fail "no more goals" + | g::uu___2 -> FStar_Tactics_Effect.lift_div_tac (fun uu___3 -> g)) let (cur_env : unit -> (FStar_Reflection_Types.env, unit) FStar_Tactics_Effect.tac_repr) = fun uu___ -> @@ -139,7 +131,7 @@ let (cur_goal_safe : (Prims.of_int (18)) (Prims.of_int (72)) (Prims.of_int (26))) (Prims.mk_range "FStar.Tactics.Derived.fst" (Prims.of_int (72)) (Prims.of_int (9)) (Prims.of_int (72)) (Prims.of_int (26))) - (Obj.magic (FStar_Tactics_Effect.get ())) + (FStar_Tactics_Effect.get ()) (fun uu___1 -> FStar_Tactics_Effect.lift_div_tac (fun uu___2 -> FStar_Tactics_Types.goals_of uu___1)))) @@ -275,14 +267,16 @@ let (trivial : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = match uu___2 with | FStar_Reflection_Formula.True_ -> Obj.magic - (exact - (FStar_Reflection_Builtins.pack_ln - (FStar_Reflection_Data.Tv_Const - FStar_Reflection_Data.C_Unit))) + (Obj.repr + (exact + (FStar_Reflection_Builtins.pack_ln + (FStar_Reflection_Data.Tv_Const + FStar_Reflection_Data.C_Unit)))) | uu___3 -> Obj.magic - (FStar_Tactics_Effect.raise - Goal_not_trivial)) uu___2))) + (Obj.repr + (FStar_Tactics_Effect.raise + Goal_not_trivial))) uu___2))) uu___2))) uu___1) let (dismiss : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = fun uu___ -> @@ -295,8 +289,9 @@ let (dismiss : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = (fun uu___1 -> (fun uu___1 -> match uu___1 with - | [] -> Obj.magic (fail "dismiss: no more goals") - | uu___2::gs -> Obj.magic (FStar_Tactics_Builtins.set_goals gs)) + | [] -> Obj.magic (Obj.repr (fail "dismiss: no more goals")) + | uu___2::gs -> + Obj.magic (Obj.repr (FStar_Tactics_Builtins.set_goals gs))) uu___1) let (flip : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = fun uu___ -> @@ -320,13 +315,17 @@ let (flip : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = (fun uu___1 -> (fun uu___1 -> match uu___1 with - | [] -> Obj.magic (fail "flip: less than two goals") + | [] -> + Obj.magic + (Obj.repr (fail "flip: less than two goals")) | uu___2::[] -> - Obj.magic (fail "flip: less than two goals") + Obj.magic + (Obj.repr (fail "flip: less than two goals")) | g1::g2::gs1 -> Obj.magic - (FStar_Tactics_Builtins.set_goals (g2 :: g1 :: - gs1))) uu___1))) uu___1) + (Obj.repr + (FStar_Tactics_Builtins.set_goals (g2 :: g1 + :: gs1)))) uu___1))) uu___1) let (qed : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = fun uu___ -> FStar_Tactics_Effect.tac_bind @@ -336,13 +335,9 @@ let (qed : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = (Prims.of_int (4)) (Prims.of_int (130)) (Prims.of_int (32))) (Obj.magic (goals ())) (fun uu___1 -> - (fun uu___1 -> - match uu___1 with - | [] -> - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac (fun uu___2 -> ()))) - | uu___2 -> Obj.magic (Obj.repr (fail "qed: not done!"))) uu___1) + match uu___1 with + | [] -> FStar_Tactics_Effect.lift_div_tac (fun uu___2 -> ()) + | uu___2 -> fail "qed: not done!") let (debug : Prims.string -> (unit, unit) FStar_Tactics_Effect.tac_repr) = fun m -> FStar_Tactics_Effect.tac_bind @@ -391,22 +386,24 @@ let (smt : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = (fun uu___1 -> (fun uu___1 -> match uu___1 with - | ([], uu___2) -> Obj.magic (fail "smt: no active goals") + | ([], uu___2) -> + Obj.magic (Obj.repr (fail "smt: no active goals")) | (g::gs, gs') -> Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "FStar.Tactics.Derived.fst" - (Prims.of_int (147)) (Prims.of_int (8)) - (Prims.of_int (147)) (Prims.of_int (20))) - (Prims.mk_range "FStar.Tactics.Derived.fst" - (Prims.of_int (148)) (Prims.of_int (8)) - (Prims.of_int (148)) (Prims.of_int (32))) - (Obj.magic (FStar_Tactics_Builtins.set_goals gs)) - (fun uu___2 -> + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "FStar.Tactics.Derived.fst" + (Prims.of_int (147)) (Prims.of_int (8)) + (Prims.of_int (147)) (Prims.of_int (20))) + (Prims.mk_range "FStar.Tactics.Derived.fst" + (Prims.of_int (148)) (Prims.of_int (8)) + (Prims.of_int (148)) (Prims.of_int (32))) + (Obj.magic (FStar_Tactics_Builtins.set_goals gs)) (fun uu___2 -> - Obj.magic - (FStar_Tactics_Builtins.set_smt_goals (g :: gs'))) - uu___2))) uu___1) + (fun uu___2 -> + Obj.magic + (FStar_Tactics_Builtins.set_smt_goals (g :: + gs'))) uu___2)))) uu___1) let (idtac : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = fun uu___ -> (fun uu___ -> @@ -425,9 +422,10 @@ let (later : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = match uu___1 with | g::gs -> Obj.magic - (FStar_Tactics_Builtins.set_goals - (FStar_List_Tot_Base.op_At gs [g])) - | uu___2 -> Obj.magic (fail "later: no goals")) uu___1) + (Obj.repr + (FStar_Tactics_Builtins.set_goals + (FStar_List_Tot_Base.op_At gs [g]))) + | uu___2 -> Obj.magic (Obj.repr (fail "later: no goals"))) uu___1) let (apply : FStar_Reflection_Types.term -> (unit, unit) FStar_Tactics_Effect.tac_repr) = fun t -> FStar_Tactics_Builtins.t_apply true false false t @@ -536,29 +534,21 @@ let (topdown_rewrite : (Prims.of_int (254)) (Prims.of_int (10))) (match i with | uu___2 when uu___2 = Prims.int_zero -> - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___3 -> - FStar_Tactics_Types.Continue))) + FStar_Tactics_Effect.lift_div_tac + (fun uu___3 -> + FStar_Tactics_Types.Continue) | uu___2 when uu___2 = Prims.int_one -> - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___3 -> - FStar_Tactics_Types.Skip))) + FStar_Tactics_Effect.lift_div_tac + (fun uu___3 -> + FStar_Tactics_Types.Skip) | uu___2 when uu___2 = (Prims.of_int (2)) -> - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___3 -> - FStar_Tactics_Types.Abort))) + FStar_Tactics_Effect.lift_div_tac + (fun uu___3 -> + FStar_Tactics_Types.Abort) | uu___2 -> - Obj.magic - (Obj.repr - (fail - "topdown_rewrite: bad value from ctrl"))) + fail + "topdown_rewrite: bad value from ctrl") (fun f -> FStar_Tactics_Effect.lift_div_tac (fun uu___2 -> (b, f))))) uu___1))) @@ -681,11 +671,8 @@ let divide : (Prims.mk_range "FStar.Tactics.Derived.fst" (Prims.of_int (289)) (Prims.of_int (4)) (Prims.of_int (301)) (Prims.of_int (10))) (if n < Prims.int_zero - then Obj.magic (Obj.repr (fail "divide: negative n")) - else - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac (fun uu___1 -> ())))) + then fail "divide: negative n" + else FStar_Tactics_Effect.lift_div_tac (fun uu___1 -> ())) (fun uu___ -> (fun uu___ -> Obj.magic @@ -1106,84 +1093,91 @@ let focus : (fun uu___ -> (fun uu___ -> match uu___ with - | [] -> Obj.magic (fail "focus: no goals") + | [] -> Obj.magic (Obj.repr (fail "focus: no goals")) | g::gs -> Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "FStar.Tactics.Derived.fst" - (Prims.of_int (314)) (Prims.of_int (18)) - (Prims.of_int (314)) (Prims.of_int (30))) - (Prims.mk_range "FStar.Tactics.Derived.fst" - (Prims.of_int (315)) (Prims.of_int (8)) - (Prims.of_int (318)) (Prims.of_int (9))) - (Obj.magic (smt_goals ())) - (fun uu___1 -> - (fun sgs -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "FStar.Tactics.Derived.fst" - (Prims.of_int (315)) (Prims.of_int (8)) - (Prims.of_int (315)) (Prims.of_int (21))) - (Prims.mk_range "FStar.Tactics.Derived.fst" - (Prims.of_int (315)) (Prims.of_int (23)) - (Prims.of_int (318)) (Prims.of_int (9))) - (Obj.magic - (FStar_Tactics_Builtins.set_goals [g])) - (fun uu___1 -> + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "FStar.Tactics.Derived.fst" + (Prims.of_int (314)) (Prims.of_int (18)) + (Prims.of_int (314)) (Prims.of_int (30))) + (Prims.mk_range "FStar.Tactics.Derived.fst" + (Prims.of_int (315)) (Prims.of_int (8)) + (Prims.of_int (318)) (Prims.of_int (9))) + (Obj.magic (smt_goals ())) + (fun uu___1 -> + (fun sgs -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.Derived.fst" + (Prims.of_int (315)) (Prims.of_int (8)) + (Prims.of_int (315)) + (Prims.of_int (21))) + (Prims.mk_range + "FStar.Tactics.Derived.fst" + (Prims.of_int (315)) + (Prims.of_int (23)) + (Prims.of_int (318)) (Prims.of_int (9))) + (Obj.magic + (FStar_Tactics_Builtins.set_goals [g])) (fun uu___1 -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.Derived.fst" - (Prims.of_int (315)) - (Prims.of_int (23)) - (Prims.of_int (315)) - (Prims.of_int (39))) - (Prims.mk_range - "FStar.Tactics.Derived.fst" - (Prims.of_int (316)) - (Prims.of_int (8)) - (Prims.of_int (318)) - (Prims.of_int (9))) - (Obj.magic - (FStar_Tactics_Builtins.set_smt_goals - [])) - (fun uu___2 -> + (fun uu___1 -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.Derived.fst" + (Prims.of_int (315)) + (Prims.of_int (23)) + (Prims.of_int (315)) + (Prims.of_int (39))) + (Prims.mk_range + "FStar.Tactics.Derived.fst" + (Prims.of_int (316)) + (Prims.of_int (8)) + (Prims.of_int (318)) + (Prims.of_int (9))) + (Obj.magic + (FStar_Tactics_Builtins.set_smt_goals + [])) (fun uu___2 -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.Derived.fst" - (Prims.of_int (316)) - (Prims.of_int (16)) - (Prims.of_int (316)) - (Prims.of_int (20))) - (Prims.mk_range - "FStar.Tactics.Derived.fst" - (Prims.of_int (317)) - (Prims.of_int (8)) - (Prims.of_int (318)) - (Prims.of_int (9))) - (Obj.magic (t ())) - (fun uu___3 -> - (fun x -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range + (fun uu___2 -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.Derived.fst" + (Prims.of_int (316)) + (Prims.of_int (16)) + (Prims.of_int (316)) + (Prims.of_int (20))) + (Prims.mk_range + "FStar.Tactics.Derived.fst" + (Prims.of_int (317)) + (Prims.of_int (8)) + (Prims.of_int (318)) + (Prims.of_int (9))) + (Obj.magic (t ())) + (fun uu___3 -> + (fun x -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + ( + Prims.mk_range "FStar.Tactics.Derived.fst" (Prims.of_int (317)) (Prims.of_int (8)) (Prims.of_int (317)) (Prims.of_int (33))) - (Prims.mk_range + ( + Prims.mk_range "FStar.Tactics.Derived.fst" (Prims.of_int (317)) (Prims.of_int (35)) (Prims.of_int (318)) (Prims.of_int (9))) - (Obj.magic ( - FStar_Tactics_Effect.tac_bind + Obj.magic + (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.Derived.fst" (Prims.of_int (317)) @@ -1227,8 +1221,9 @@ let focus : (FStar_Tactics_Builtins.set_goals uu___3)) uu___3))) - (fun uu___3 - -> + ( + fun + uu___3 -> (fun uu___3 -> Obj.magic @@ -1299,8 +1294,9 @@ let focus : uu___5 -> x)))) uu___3))) - uu___3))) uu___2))) - uu___1))) uu___1))) uu___) + uu___3))) + uu___2))) uu___1))) uu___1)))) + uu___) let (dump1 : Prims.string -> (unit, unit) FStar_Tactics_Effect.tac_repr) = fun m -> focus (fun uu___ -> FStar_Tactics_Builtins.dump m) let rec mapAll : @@ -1786,10 +1782,8 @@ let (guard : Prims.bool -> (unit, unit) FStar_Tactics_Effect.tac_repr) = fun uu___ -> (fun b -> if Prims.op_Negation b - then Obj.magic (Obj.repr (fail "guard failed")) - else - Obj.magic - (Obj.repr (FStar_Tactics_Effect.lift_div_tac (fun uu___1 -> ())))) + then Obj.magic (fail "guard failed") + else Obj.magic (FStar_Tactics_Effect.lift_div_tac (fun uu___1 -> ()))) uu___ let try_with : 'a . @@ -1863,7 +1857,8 @@ let first : = fun ts -> FStar_List_Tot_Base.fold_right op_Less_Bar_Greater ts - (fun uu___ -> fail "no tactics to try") () + (fun uu___ -> (fun uu___ -> Obj.magic (fail "no tactics to try")) uu___) + () let rec repeat : 'a . (unit -> ('a, unit) FStar_Tactics_Effect.tac_repr) -> @@ -2141,8 +2136,9 @@ let (skip_guard : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = (Obj.magic (is_guard ())) (fun uu___1 -> (fun uu___1 -> - if uu___1 then Obj.magic (smt ()) else Obj.magic (fail "")) - uu___1) + if uu___1 + then Obj.magic (Obj.repr (smt ())) + else Obj.magic (Obj.repr (fail ""))) uu___1) let (guards_to_smt : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = fun uu___ -> FStar_Tactics_Effect.tac_bind @@ -2387,50 +2383,57 @@ let rec (__assumption_aux : FStar_Reflection_Types.binders -> (unit, unit) FStar_Tactics_Effect.tac_repr) = - fun bs -> - match bs with - | [] -> fail "no assumption matches goal" - | b::bs1 -> - FStar_Tactics_Effect.tac_bind - (Prims.mk_range "FStar.Tactics.Derived.fst" (Prims.of_int (539)) - (Prims.of_int (16)) (Prims.of_int (539)) (Prims.of_int (32))) - (Prims.mk_range "FStar.Tactics.Derived.fst" (Prims.of_int (540)) - (Prims.of_int (8)) (Prims.of_int (543)) (Prims.of_int (27))) - (Obj.magic (binder_to_term b)) - (fun uu___ -> - (fun t -> - Obj.magic - (try_with (fun uu___ -> match () with | () -> exact t) - (fun uu___ -> - try_with - (fun uu___1 -> - match () with - | () -> - FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.Derived.fst" - (Prims.of_int (541)) - (Prims.of_int (13)) - (Prims.of_int (541)) - (Prims.of_int (48))) - (Prims.mk_range - "FStar.Tactics.Derived.fst" - (Prims.of_int (542)) - (Prims.of_int (13)) - (Prims.of_int (542)) - (Prims.of_int (20))) - (Obj.magic - (apply - (FStar_Reflection_Builtins.pack_ln - (FStar_Reflection_Data.Tv_FVar - (FStar_Reflection_Builtins.pack_fv - ["FStar"; - "Squash"; - "return_squash"]))))) - (fun uu___2 -> - (fun uu___2 -> Obj.magic (exact t)) - uu___2)) - (fun uu___1 -> __assumption_aux bs1)))) uu___) + fun uu___ -> + (fun bs -> + match bs with + | [] -> Obj.magic (Obj.repr (fail "no assumption matches goal")) + | b::bs1 -> + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "FStar.Tactics.Derived.fst" + (Prims.of_int (539)) (Prims.of_int (16)) + (Prims.of_int (539)) (Prims.of_int (32))) + (Prims.mk_range "FStar.Tactics.Derived.fst" + (Prims.of_int (540)) (Prims.of_int (8)) + (Prims.of_int (543)) (Prims.of_int (27))) + (Obj.magic (binder_to_term b)) + (fun uu___ -> + (fun t -> + Obj.magic + (try_with + (fun uu___ -> match () with | () -> exact t) + (fun uu___ -> + try_with + (fun uu___1 -> + match () with + | () -> + FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.Derived.fst" + (Prims.of_int (541)) + (Prims.of_int (13)) + (Prims.of_int (541)) + (Prims.of_int (48))) + (Prims.mk_range + "FStar.Tactics.Derived.fst" + (Prims.of_int (542)) + (Prims.of_int (13)) + (Prims.of_int (542)) + (Prims.of_int (20))) + (Obj.magic + (apply + (FStar_Reflection_Builtins.pack_ln + (FStar_Reflection_Data.Tv_FVar + (FStar_Reflection_Builtins.pack_fv + ["FStar"; + "Squash"; + "return_squash"]))))) + (fun uu___2 -> + (fun uu___2 -> + Obj.magic (exact t)) uu___2)) + (fun uu___1 -> __assumption_aux bs1)))) + uu___)))) uu___ let (assumption : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = fun uu___ -> FStar_Tactics_Effect.tac_bind @@ -2522,7 +2525,8 @@ let (rewrite' : (fun uu___2 -> Obj.magic (FStar_Tactics_Builtins.rewrite b)) uu___2))) uu___1))) - (fun uu___ -> fail "rewrite' failed") () + (fun uu___ -> (fun uu___ -> Obj.magic (fail "rewrite' failed")) uu___) + () let rec (try_rewrite_equality : FStar_Reflection_Types.term -> FStar_Reflection_Types.binders -> @@ -2638,23 +2642,26 @@ let (unfold_def : match uu___ with | FStar_Reflection_Data.Tv_FVar fv -> Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "FStar.Tactics.Derived.fst" - (Prims.of_int (601)) (Prims.of_int (16)) - (Prims.of_int (601)) (Prims.of_int (42))) - (Prims.mk_range "FStar.Tactics.Derived.fst" - (Prims.of_int (602)) (Prims.of_int (8)) - (Prims.of_int (602)) (Prims.of_int (30))) - (FStar_Tactics_Effect.lift_div_tac + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "FStar.Tactics.Derived.fst" + (Prims.of_int (601)) (Prims.of_int (16)) + (Prims.of_int (601)) (Prims.of_int (42))) + (Prims.mk_range "FStar.Tactics.Derived.fst" + (Prims.of_int (602)) (Prims.of_int (8)) + (Prims.of_int (602)) (Prims.of_int (30))) + (FStar_Tactics_Effect.lift_div_tac + (fun uu___1 -> + FStar_Reflection_Builtins.implode_qn + (FStar_Reflection_Builtins.inspect_fv fv))) (fun uu___1 -> - FStar_Reflection_Builtins.implode_qn - (FStar_Reflection_Builtins.inspect_fv fv))) - (fun uu___1 -> - (fun n -> - Obj.magic - (FStar_Tactics_Builtins.norm - [FStar_Pervasives.delta_fully [n]])) uu___1)) - | uu___1 -> Obj.magic (fail "unfold_def: term is not a fv")) + (fun n -> + Obj.magic + (FStar_Tactics_Builtins.norm + [FStar_Pervasives.delta_fully [n]])) + uu___1))) + | uu___1 -> + Obj.magic (Obj.repr (fail "unfold_def: term is not a fv"))) uu___) let (l_to_r : FStar_Reflection_Types.term Prims.list -> @@ -2885,79 +2892,83 @@ let (grewrite_eq : | FStar_Reflection_Formula.Comp (FStar_Reflection_Formula.Eq uu___3, l, r) -> Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.Derived.fst" - (Prims.of_int (655)) - (Prims.of_int (6)) - (Prims.of_int (655)) - (Prims.of_int (18))) - (Prims.mk_range - "FStar.Tactics.Derived.fst" - (Prims.of_int (656)) - (Prims.of_int (6)) - (Prims.of_int (657)) - (Prims.of_int (56))) - (Obj.magic (grewrite l r)) - (fun uu___4 -> + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.Derived.fst" + (Prims.of_int (655)) + (Prims.of_int (6)) + (Prims.of_int (655)) + (Prims.of_int (18))) + (Prims.mk_range + "FStar.Tactics.Derived.fst" + (Prims.of_int (656)) + (Prims.of_int (6)) + (Prims.of_int (657)) + (Prims.of_int (56))) + (Obj.magic (grewrite l r)) (fun uu___4 -> - Obj.magic - (iseq - [idtac; - (fun uu___5 -> - FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.Derived.fst" - (Prims.of_int (656)) - (Prims.of_int (30)) - (Prims.of_int (656)) - (Prims.of_int (55))) - (Prims.mk_range - "FStar.Tactics.Derived.fst" - (Prims.of_int (657)) - (Prims.of_int (30)) - (Prims.of_int (657)) - (Prims.of_int (54))) - (Obj.magic - (apply_lemma - (FStar_Reflection_Builtins.pack_ln - (FStar_Reflection_Data.Tv_FVar - (FStar_Reflection_Builtins.pack_fv - ["FStar"; - "Tactics"; - "Derived"; - "__un_sq_eq"]))))) - (fun uu___6 -> + (fun uu___4 -> + Obj.magic + (iseq + [idtac; + (fun uu___5 -> + FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.Derived.fst" + (Prims.of_int (656)) + (Prims.of_int (30)) + (Prims.of_int (656)) + (Prims.of_int (55))) + (Prims.mk_range + "FStar.Tactics.Derived.fst" + (Prims.of_int (657)) + (Prims.of_int (30)) + (Prims.of_int (657)) + (Prims.of_int (54))) + (Obj.magic + (apply_lemma + (FStar_Reflection_Builtins.pack_ln + (FStar_Reflection_Data.Tv_FVar + (FStar_Reflection_Builtins.pack_fv + ["FStar"; + "Tactics"; + "Derived"; + "__un_sq_eq"]))))) (fun uu___6 -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.Derived.fst" - (Prims.of_int (657)) - (Prims.of_int (36)) - (Prims.of_int (657)) - (Prims.of_int (54))) - (Prims.mk_range - "FStar.Tactics.Derived.fst" - (Prims.of_int (657)) - (Prims.of_int (30)) - (Prims.of_int (657)) - (Prims.of_int (54))) - (Obj.magic - (binder_to_term + (fun uu___6 -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.Derived.fst" + (Prims.of_int (657)) + (Prims.of_int (36)) + (Prims.of_int (657)) + (Prims.of_int (54))) + (Prims.mk_range + "FStar.Tactics.Derived.fst" + (Prims.of_int (657)) + (Prims.of_int (30)) + (Prims.of_int (657)) + (Prims.of_int (54))) + (Obj.magic + (binder_to_term b)) - (fun uu___7 -> (fun uu___7 -> + (fun + uu___7 -> Obj.magic (exact uu___7)) uu___7))) - uu___6))])) uu___4)) + uu___6))])) + uu___4))) | uu___3 -> Obj.magic - (fail - "grewrite_eq: binder type is not an equality")) + (Obj.repr + (fail + "grewrite_eq: binder type is not an equality"))) uu___2))) uu___) let rec (apply_squash_or_lem : Prims.nat -> @@ -2989,101 +3000,121 @@ let rec (apply_squash_or_lem : (fun uu___1 -> try_with (fun uu___2 -> match () with | () -> apply_lemma t) (fun uu___2 -> - if d <= Prims.int_zero - then fail "mapply: out of fuel" - else - FStar_Tactics_Effect.tac_bind - (Prims.mk_range "FStar.Tactics.Derived.fst" - (Prims.of_int (688)) (Prims.of_int (13)) - (Prims.of_int (688)) (Prims.of_int (30))) - (Prims.mk_range "FStar.Tactics.Derived.fst" - (Prims.of_int (689)) (Prims.of_int (4)) - (Prims.of_int (737)) (Prims.of_int (41))) - (Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "FStar.Tactics.Derived.fst" - (Prims.of_int (688)) (Prims.of_int (16)) - (Prims.of_int (688)) (Prims.of_int (28))) - (Prims.mk_range "FStar.Tactics.Derived.fst" - (Prims.of_int (688)) (Prims.of_int (13)) - (Prims.of_int (688)) (Prims.of_int (30))) - (Obj.magic (cur_env ())) - (fun uu___4 -> + (fun uu___2 -> + if d <= Prims.int_zero + then + Obj.magic (Obj.repr (fail "mapply: out of fuel")) + else + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "FStar.Tactics.Derived.fst" + (Prims.of_int (688)) (Prims.of_int (13)) + (Prims.of_int (688)) (Prims.of_int (30))) + (Prims.mk_range "FStar.Tactics.Derived.fst" + (Prims.of_int (689)) (Prims.of_int (4)) + (Prims.of_int (737)) (Prims.of_int (41))) + (Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.Derived.fst" + (Prims.of_int (688)) + (Prims.of_int (16)) + (Prims.of_int (688)) + (Prims.of_int (28))) + (Prims.mk_range + "FStar.Tactics.Derived.fst" + (Prims.of_int (688)) + (Prims.of_int (13)) + (Prims.of_int (688)) + (Prims.of_int (30))) + (Obj.magic (cur_env ())) + (fun uu___4 -> + (fun uu___4 -> + Obj.magic + (FStar_Tactics_Builtins.tc + uu___4 t)) uu___4))) (fun uu___4 -> - Obj.magic - (FStar_Tactics_Builtins.tc uu___4 t)) - uu___4))) - (fun uu___4 -> - (fun ty -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.Derived.fst" - (Prims.of_int (689)) - (Prims.of_int (17)) - (Prims.of_int (689)) - (Prims.of_int (31))) - (Prims.mk_range - "FStar.Tactics.Derived.fst" - (Prims.of_int (689)) - (Prims.of_int (4)) - (Prims.of_int (737)) - (Prims.of_int (41))) - (Obj.magic - (FStar_Tactics_SyntaxHelpers.collect_arr - ty)) - (fun uu___4 -> - (fun uu___4 -> - match uu___4 with - | (tys, c) -> - (match FStar_Reflection_Builtins.inspect_comp - c - with - | FStar_Reflection_Data.C_Lemma - (pre, post, uu___5) -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.Derived.fst" - (Prims.of_int (693)) - (Prims.of_int (18)) - (Prims.of_int (693)) - (Prims.of_int (32))) - (Prims.mk_range - "FStar.Tactics.Derived.fst" - (Prims.of_int (694)) - (Prims.of_int (7)) - (Prims.of_int (702)) - (Prims.of_int (41))) - (FStar_Tactics_Effect.lift_div_tac - (fun uu___6 -> - FStar_Reflection_Builtins.pack_ln - (FStar_Reflection_Data.Tv_App - (post, + (fun ty -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.Derived.fst" + (Prims.of_int (689)) + (Prims.of_int (17)) + (Prims.of_int (689)) + (Prims.of_int (31))) + (Prims.mk_range + "FStar.Tactics.Derived.fst" + (Prims.of_int (689)) + (Prims.of_int (4)) + (Prims.of_int (737)) + (Prims.of_int (41))) + (Obj.magic + (FStar_Tactics_SyntaxHelpers.collect_arr + ty)) + (fun uu___4 -> + (fun uu___4 -> + match uu___4 with + | (tys, c) -> + (match FStar_Reflection_Builtins.inspect_comp + c + with + | FStar_Reflection_Data.C_Lemma + (pre, post, + uu___5) + -> + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.tac_bind + ( + Prims.mk_range + "FStar.Tactics.Derived.fst" + (Prims.of_int (693)) + (Prims.of_int (18)) + (Prims.of_int (693)) + (Prims.of_int (32))) + ( + Prims.mk_range + "FStar.Tactics.Derived.fst" + (Prims.of_int (694)) + (Prims.of_int (7)) + (Prims.of_int (702)) + (Prims.of_int (41))) + ( + FStar_Tactics_Effect.lift_div_tac + (fun + uu___6 -> + FStar_Reflection_Builtins.pack_ln + (FStar_Reflection_Data.Tv_App + (post, ((FStar_Reflection_Builtins.pack_ln (FStar_Reflection_Data.Tv_Const FStar_Reflection_Data.C_Unit)), FStar_Reflection_Data.Q_Explicit))))) - (fun uu___6 -> - (fun post1 -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range + ( + fun + uu___6 -> + (fun + post1 -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "FStar.Tactics.Derived.fst" (Prims.of_int (694)) (Prims.of_int (18)) (Prims.of_int (694)) (Prims.of_int (35))) - (Prims.mk_range + (Prims.mk_range "FStar.Tactics.Derived.fst" (Prims.of_int (696)) (Prims.of_int (7)) (Prims.of_int (702)) (Prims.of_int (41))) - (Obj.magic + (Obj.magic (norm_term [] post1)) - (fun + (fun uu___6 -> (fun post2 -> @@ -3114,6 +3145,7 @@ let rec (apply_squash_or_lem : FStar_Reflection_Formula.Implies (p, q) -> Obj.magic + (Obj.repr (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.Derived.fst" @@ -3145,42 +3177,47 @@ let rec (apply_squash_or_lem : (d - Prims.int_one) t)) - uu___7)) + uu___7))) | uu___7 -> Obj.magic + (Obj.repr (fail - "mapply: can't apply (1)")) + "mapply: can't apply (1)"))) + uu___6))) uu___6))) uu___6))) - uu___6)) - | FStar_Reflection_Data.C_Total - rt -> - (match FStar_Reflection_Derived.unsquash_term - rt - with - | FStar_Pervasives_Native.Some - rt1 -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.Derived.fst" - (Prims.of_int (710)) - (Prims.of_int (18)) - (Prims.of_int (710)) - (Prims.of_int (33))) - (Prims.mk_range - "FStar.Tactics.Derived.fst" - (Prims.of_int (712)) - (Prims.of_int (9)) - (Prims.of_int (718)) - (Prims.of_int (43))) - (Obj.magic - (norm_term [] - rt1)) - (fun uu___5 -> - (fun rt2 -> - Obj.magic + | FStar_Reflection_Data.C_Total + rt -> + Obj.magic + (Obj.repr + (match + FStar_Reflection_Derived.unsquash_term + rt + with + | FStar_Pervasives_Native.Some + rt1 -> + FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.Derived.fst" + (Prims.of_int (710)) + (Prims.of_int (18)) + (Prims.of_int (710)) + (Prims.of_int (33))) + (Prims.mk_range + "FStar.Tactics.Derived.fst" + (Prims.of_int (712)) + (Prims.of_int (9)) + (Prims.of_int (718)) + (Prims.of_int (43))) + (Obj.magic + (norm_term + [] rt1)) + (fun + uu___5 -> + (fun rt2 + -> + Obj.magic (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.Derived.fst" @@ -3207,6 +3244,7 @@ let rec (apply_squash_or_lem : FStar_Reflection_Formula.Implies (p, q) -> Obj.magic + (Obj.repr (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.Derived.fst" @@ -3238,36 +3276,38 @@ let rec (apply_squash_or_lem : (d - Prims.int_one) t)) - uu___6)) + uu___6))) | uu___6 -> Obj.magic + (Obj.repr (fail - "mapply: can't apply (1)")) + "mapply: can't apply (1)"))) uu___5))) - uu___5)) - | FStar_Pervasives_Native.None - -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.Derived.fst" - (Prims.of_int (725)) - (Prims.of_int (18)) - (Prims.of_int (725)) - (Prims.of_int (33))) - (Prims.mk_range - "FStar.Tactics.Derived.fst" - (Prims.of_int (727)) - (Prims.of_int (9)) - (Prims.of_int (734)) - (Prims.of_int (20))) - (Obj.magic - (norm_term [] - rt)) - (fun uu___5 -> - (fun rt1 -> - Obj.magic + uu___5) + | FStar_Pervasives_Native.None + -> + FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.Derived.fst" + (Prims.of_int (725)) + (Prims.of_int (18)) + (Prims.of_int (725)) + (Prims.of_int (33))) + (Prims.mk_range + "FStar.Tactics.Derived.fst" + (Prims.of_int (727)) + (Prims.of_int (9)) + (Prims.of_int (734)) + (Prims.of_int (20))) + (Obj.magic + (norm_term + [] rt)) + (fun + uu___5 -> + (fun rt1 + -> + Obj.magic (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.Derived.fst" @@ -3358,12 +3398,13 @@ let rec (apply_squash_or_lem : (apply t)) uu___7))) uu___5))) - uu___5))) - | uu___5 -> - Obj.magic - (fail - "mapply: can't apply (2)"))) - uu___4))) uu___4)))) + uu___5))) + | uu___5 -> + Obj.magic + (Obj.repr + (fail + "mapply: can't apply (2)")))) + uu___4))) uu___4)))) uu___2))) let (mapply : FStar_Reflection_Types.term -> (unit, unit) FStar_Tactics_Effect.tac_repr) = fun t -> apply_squash_or_lem (Prims.of_int (10)) t @@ -3477,7 +3518,10 @@ let finish_by : (Prims.of_int (773)) (Prims.of_int (9))) (Obj.magic (or_else qed - (fun uu___ -> fail "finish_by: not finished"))) + (fun uu___ -> + (fun uu___ -> + Obj.magic (fail "finish_by: not finished")) + uu___))) (fun uu___ -> FStar_Tactics_Effect.lift_div_tac (fun uu___1 -> x)))) uu___) @@ -3632,11 +3676,12 @@ let (tlabel : Prims.string -> (unit, unit) FStar_Tactics_Effect.tac_repr) = (fun uu___ -> (fun uu___ -> match uu___ with - | [] -> Obj.magic (fail "tlabel: no goals") + | [] -> Obj.magic (Obj.repr (fail "tlabel: no goals")) | h::t -> Obj.magic - (FStar_Tactics_Builtins.set_goals - ((FStar_Tactics_Types.set_label l h) :: t))) uu___) + (Obj.repr + (FStar_Tactics_Builtins.set_goals + ((FStar_Tactics_Types.set_label l h) :: t)))) uu___) let (tlabel' : Prims.string -> (unit, unit) FStar_Tactics_Effect.tac_repr) = fun l -> FStar_Tactics_Effect.tac_bind @@ -3648,26 +3693,27 @@ let (tlabel' : Prims.string -> (unit, unit) FStar_Tactics_Effect.tac_repr) = (fun uu___ -> (fun uu___ -> match uu___ with - | [] -> Obj.magic (fail "tlabel': no goals") + | [] -> Obj.magic (Obj.repr (fail "tlabel': no goals")) | h::t -> Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "FStar.Tactics.Derived.fst" - (Prims.of_int (821)) (Prims.of_int (16)) - (Prims.of_int (821)) (Prims.of_int (45))) - (Prims.mk_range "FStar.Tactics.Derived.fst" - (Prims.of_int (822)) (Prims.of_int (8)) - (Prims.of_int (822)) (Prims.of_int (26))) - (FStar_Tactics_Effect.lift_div_tac + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "FStar.Tactics.Derived.fst" + (Prims.of_int (821)) (Prims.of_int (16)) + (Prims.of_int (821)) (Prims.of_int (45))) + (Prims.mk_range "FStar.Tactics.Derived.fst" + (Prims.of_int (822)) (Prims.of_int (8)) + (Prims.of_int (822)) (Prims.of_int (26))) + (FStar_Tactics_Effect.lift_div_tac + (fun uu___1 -> + FStar_Tactics_Types.set_label + (Prims.strcat l + (FStar_Tactics_Types.get_label h)) h)) (fun uu___1 -> - FStar_Tactics_Types.set_label - (Prims.strcat l - (FStar_Tactics_Types.get_label h)) h)) - (fun uu___1 -> - (fun h1 -> - Obj.magic - (FStar_Tactics_Builtins.set_goals (h1 :: t))) - uu___1))) uu___) + (fun h1 -> + Obj.magic + (FStar_Tactics_Builtins.set_goals (h1 :: t))) + uu___1)))) uu___) let (focus_all : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = fun uu___ -> FStar_Tactics_Effect.tac_bind @@ -3749,9 +3795,11 @@ let (bump_nth : Prims.pos -> (unit, unit) FStar_Tactics_Effect.tac_repr) = (fun uu___ -> match uu___ with | FStar_Pervasives_Native.None -> - Obj.magic (fail "bump_nth: not that many goals") + Obj.magic (Obj.repr (fail "bump_nth: not that many goals")) | FStar_Pervasives_Native.Some (h, t) -> - Obj.magic (FStar_Tactics_Builtins.set_goals (h :: t))) uu___) + Obj.magic + (Obj.repr (FStar_Tactics_Builtins.set_goals (h :: t)))) + uu___) let rec (destruct_list : FStar_Reflection_Types.term -> (FStar_Reflection_Types.term Prims.list, unit) @@ -3781,20 +3829,22 @@ let rec (destruct_list : (FStar_Reflection_Builtins.inspect_fv fv) = FStar_Reflection_Const.cons_qn then - FStar_Tactics_Effect.tac_bind - (Prims.mk_range "FStar.Tactics.Derived.fst" - (Prims.of_int (851)) (Prims.of_int (17)) - (Prims.of_int (851)) (Prims.of_int (33))) - (Prims.mk_range "FStar.Tactics.Derived.fst" - (Prims.of_int (851)) (Prims.of_int (14)) - (Prims.of_int (851)) (Prims.of_int (16))) - (Obj.magic (destruct_list a2)) - (fun uu___1 -> - FStar_Tactics_Effect.lift_div_tac - (fun uu___2 -> a1 :: uu___1)) + Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "FStar.Tactics.Derived.fst" + (Prims.of_int (851)) (Prims.of_int (17)) + (Prims.of_int (851)) (Prims.of_int (33))) + (Prims.mk_range "FStar.Tactics.Derived.fst" + (Prims.of_int (851)) (Prims.of_int (14)) + (Prims.of_int (851)) (Prims.of_int (16))) + (Obj.magic (destruct_list a2)) + (fun uu___1 -> + FStar_Tactics_Effect.lift_div_tac + (fun uu___2 -> a1 :: uu___1))) else - FStar_Tactics_Effect.raise - FStar_Tactics_Common.NotAListLiteral)) + Obj.repr + (FStar_Tactics_Effect.raise + FStar_Tactics_Common.NotAListLiteral))) | (FStar_Reflection_Data.Tv_FVar fv, (uu___1, FStar_Reflection_Data.Q_Implicit)::(a1, FStar_Reflection_Data.Q_Explicit):: @@ -3805,20 +3855,22 @@ let rec (destruct_list : (FStar_Reflection_Builtins.inspect_fv fv) = FStar_Reflection_Const.cons_qn then - FStar_Tactics_Effect.tac_bind - (Prims.mk_range "FStar.Tactics.Derived.fst" - (Prims.of_int (851)) (Prims.of_int (17)) - (Prims.of_int (851)) (Prims.of_int (33))) - (Prims.mk_range "FStar.Tactics.Derived.fst" - (Prims.of_int (851)) (Prims.of_int (14)) - (Prims.of_int (851)) (Prims.of_int (16))) - (Obj.magic (destruct_list a2)) - (fun uu___2 -> - FStar_Tactics_Effect.lift_div_tac - (fun uu___3 -> a1 :: uu___2)) + Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "FStar.Tactics.Derived.fst" + (Prims.of_int (851)) (Prims.of_int (17)) + (Prims.of_int (851)) (Prims.of_int (33))) + (Prims.mk_range "FStar.Tactics.Derived.fst" + (Prims.of_int (851)) (Prims.of_int (14)) + (Prims.of_int (851)) (Prims.of_int (16))) + (Obj.magic (destruct_list a2)) + (fun uu___2 -> + FStar_Tactics_Effect.lift_div_tac + (fun uu___3 -> a1 :: uu___2))) else - FStar_Tactics_Effect.raise - FStar_Tactics_Common.NotAListLiteral)) + Obj.repr + (FStar_Tactics_Effect.raise + FStar_Tactics_Common.NotAListLiteral))) | (FStar_Reflection_Data.Tv_FVar fv, uu___1) -> Obj.magic (Obj.repr @@ -3826,13 +3878,11 @@ let rec (destruct_list : (FStar_Reflection_Builtins.inspect_fv fv) = FStar_Reflection_Const.nil_qn then - Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___2 -> [])) + FStar_Tactics_Effect.lift_div_tac + (fun uu___2 -> []) else - Obj.repr - (FStar_Tactics_Effect.raise - FStar_Tactics_Common.NotAListLiteral))) + FStar_Tactics_Effect.raise + FStar_Tactics_Common.NotAListLiteral)) | uu___1 -> Obj.magic (Obj.repr @@ -3860,30 +3910,25 @@ let (get_match_body : (fun uu___1 -> (fun uu___1 -> match uu___1 with - | FStar_Pervasives_Native.None -> Obj.magic (fail "") + | FStar_Pervasives_Native.None -> Obj.magic (Obj.repr (fail "")) | FStar_Pervasives_Native.Some t -> Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "FStar.Tactics.Derived.fst" - (Prims.of_int (863)) (Prims.of_int (20)) - (Prims.of_int (863)) (Prims.of_int (39))) - (Prims.mk_range "FStar.Tactics.Derived.fst" - (Prims.of_int (863)) (Prims.of_int (14)) - (Prims.of_int (865)) (Prims.of_int (46))) - (Obj.magic (inspect_unascribe t)) - (fun uu___2 -> + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "FStar.Tactics.Derived.fst" + (Prims.of_int (863)) (Prims.of_int (20)) + (Prims.of_int (863)) (Prims.of_int (39))) + (Prims.mk_range "FStar.Tactics.Derived.fst" + (Prims.of_int (863)) (Prims.of_int (14)) + (Prims.of_int (865)) (Prims.of_int (46))) + (Obj.magic (inspect_unascribe t)) (fun uu___2 -> match uu___2 with | FStar_Reflection_Data.Tv_Match (sc, uu___3, uu___4) -> - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___5 -> sc))) - | uu___3 -> - Obj.magic - (Obj.repr (fail "Goal is not a match"))) - uu___2))) uu___1) + FStar_Tactics_Effect.lift_div_tac + (fun uu___5 -> sc) + | uu___3 -> fail "Goal is not a match")))) uu___1) let rec last : 'a . 'a Prims.list -> ('a, unit) FStar_Tactics_Effect.tac_repr = fun uu___ -> @@ -4016,26 +4061,17 @@ let (nth_binder : (Prims.of_int (896)) (Prims.of_int (2)) (Prims.of_int (898)) (Prims.of_int (15))) (if k < Prims.int_zero - then - Obj.magic - (Obj.repr (fail "not enough binders")) + then fail "not enough binders" else - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___1 -> k)))) - (fun uu___ -> - (fun k1 -> - match FStar_List_Tot_Base.nth bs k1 with - | FStar_Pervasives_Native.None -> - Obj.magic - (Obj.repr (fail "not enough binders")) - | FStar_Pervasives_Native.Some b -> - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___ -> b)))) uu___))) - uu___))) uu___) + FStar_Tactics_Effect.lift_div_tac + (fun uu___1 -> k)) + (fun k1 -> + match FStar_List_Tot_Base.nth bs k1 with + | FStar_Pervasives_Native.None -> + fail "not enough binders" + | FStar_Pervasives_Native.Some b -> + FStar_Tactics_Effect.lift_div_tac + (fun uu___ -> b)))) uu___))) uu___) exception Appears let (uu___is_Appears : Prims.exn -> Prims.bool) = fun projectee -> match projectee with | Appears -> true | uu___ -> false @@ -4077,15 +4113,10 @@ let (name_appears_in : (if (FStar_Reflection_Builtins.inspect_fv fv) = nm - then - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.raise Appears)) + then FStar_Tactics_Effect.raise Appears else - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___3 -> ())))) + FStar_Tactics_Effect.lift_div_tac + (fun uu___3 -> ())) (fun uu___2 -> FStar_Tactics_Effect.lift_div_tac (fun uu___3 -> t1))) @@ -4126,12 +4157,9 @@ let (name_appears_in : match uu___ with | Appears -> Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___1 -> true))) - | e -> - Obj.magic - (Obj.repr (FStar_Tactics_Effect.raise e))) + (FStar_Tactics_Effect.lift_div_tac + (fun uu___1 -> true)) + | e -> Obj.magic (FStar_Tactics_Effect.raise e)) uu___))) uu___) let rec (mk_abs : FStar_Reflection_Types.binder Prims.list -> diff --git a/ocaml/fstar-lib/generated/FStar_Tactics_Effect.ml b/ocaml/fstar-lib/generated/FStar_Tactics_Effect.ml index 226dda9374b..c20d9427b13 100644 --- a/ocaml/fstar-lib/generated/FStar_Tactics_Effect.ml +++ b/ocaml/fstar-lib/generated/FStar_Tactics_Effect.ml @@ -49,9 +49,9 @@ let __proj__TAC__item__bind = tac_bind type ('a, 'wp, 'uuuuu, 'uuuuu1) lift_div_tac_wp = 'wp let lift_div_tac : 'a 'wp . (unit -> 'a) -> ('a, 'wp) tac_repr = fun f -> fun ps -> FStar_Tactics_Result.Success ((f ()), ps) -let (get : unit -> (FStar_Tactics_Types.proofstate, unit) tac_repr) = +let (get : unit -> (FStar_Tactics_Types.proofstate, Obj.t) tac_repr) = fun uu___ -> fun ps -> FStar_Tactics_Result.Success (ps, ps) -let raise : 'a . Prims.exn -> ('a, unit) tac_repr = +let raise : 'a . Prims.exn -> ('a, Obj.t) tac_repr = fun e -> fun ps -> FStar_Tactics_Result.Failed (e, ps) type ('uuuuu, 'p) with_tactic = 'p let (rewrite_with_tactic : diff --git a/ocaml/fstar-lib/generated/FStar_Tactics_Logic.ml b/ocaml/fstar-lib/generated/FStar_Tactics_Logic.ml index 9848dd07536..453aaf00628 100644 --- a/ocaml/fstar-lib/generated/FStar_Tactics_Logic.ml +++ b/ocaml/fstar-lib/generated/FStar_Tactics_Logic.ml @@ -105,7 +105,10 @@ let (split : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = (FStar_Reflection_Data.Tv_FVar (FStar_Reflection_Builtins.pack_fv ["FStar"; "Tactics"; "Logic"; "split_lem"])))) - (fun uu___1 -> FStar_Tactics_Derived.fail "Could not split goal") + (fun uu___1 -> + (fun uu___1 -> + Obj.magic (FStar_Tactics_Derived.fail "Could not split goal")) + uu___1) let (implies_intro : unit -> (FStar_Reflection_Types.binder, unit) FStar_Tactics_Effect.tac_repr) = @@ -219,12 +222,9 @@ let (pose_lemma : (Prims.of_int (126)) (Prims.of_int (5))) (match FStar_Reflection_Builtins.inspect_comp c with | FStar_Reflection_Data.C_Lemma (pre, post, uu___) -> - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___1 -> (pre, post)))) - | uu___ -> - Obj.magic (Obj.repr (FStar_Tactics_Derived.fail ""))) + FStar_Tactics_Effect.lift_div_tac + (fun uu___1 -> (pre, post)) + | uu___ -> FStar_Tactics_Derived.fail "") (fun uu___ -> (fun uu___ -> match uu___ with @@ -729,74 +729,80 @@ let rec (simplify_eq_implication : match r with | FStar_Pervasives_Native.None -> Obj.magic - (FStar_Tactics_Derived.fail - "Not an equality implication") + (Obj.repr + (FStar_Tactics_Derived.fail + "Not an equality implication")) | FStar_Pervasives_Native.Some (uu___1, rhs) -> Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.Logic.fst" - (Prims.of_int (160)) - (Prims.of_int (19)) - (Prims.of_int (160)) - (Prims.of_int (35))) - (Prims.mk_range - "FStar.Tactics.Logic.fst" - (Prims.of_int (161)) - (Prims.of_int (8)) - (Prims.of_int (163)) - (Prims.of_int (37))) - (Obj.magic (implies_intro ())) - (fun uu___2 -> - (fun eq_h -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.Logic.fst" - (Prims.of_int (161)) - (Prims.of_int (8)) - (Prims.of_int (161)) - (Prims.of_int (20))) - (Prims.mk_range - "FStar.Tactics.Logic.fst" - (Prims.of_int (162)) - (Prims.of_int (8)) - (Prims.of_int (163)) - (Prims.of_int (37))) - (Obj.magic - (FStar_Tactics_Builtins.rewrite - eq_h)) - (fun uu___2 -> + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.Logic.fst" + (Prims.of_int (160)) + (Prims.of_int (19)) + (Prims.of_int (160)) + (Prims.of_int (35))) + (Prims.mk_range + "FStar.Tactics.Logic.fst" + (Prims.of_int (161)) + (Prims.of_int (8)) + (Prims.of_int (163)) + (Prims.of_int (37))) + (Obj.magic (implies_intro ())) + (fun uu___2 -> + (fun eq_h -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.Logic.fst" + (Prims.of_int (161)) + (Prims.of_int (8)) + (Prims.of_int (161)) + (Prims.of_int (20))) + (Prims.mk_range + "FStar.Tactics.Logic.fst" + (Prims.of_int (162)) + (Prims.of_int (8)) + (Prims.of_int (163)) + (Prims.of_int (37))) + (Obj.magic + (FStar_Tactics_Builtins.rewrite + eq_h)) (fun uu___2 -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range + (fun uu___2 -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + ( + Prims.mk_range "FStar.Tactics.Logic.fst" (Prims.of_int (162)) (Prims.of_int (8)) (Prims.of_int (162)) (Prims.of_int (20))) - (Prims.mk_range + ( + Prims.mk_range "FStar.Tactics.Logic.fst" (Prims.of_int (163)) (Prims.of_int (8)) (Prims.of_int (163)) (Prims.of_int (37))) - (Obj.magic ( - FStar_Tactics_Builtins.clear_top + Obj.magic + (FStar_Tactics_Builtins.clear_top ())) - (fun uu___3 - -> + ( + fun + uu___3 -> (fun uu___3 -> Obj.magic (visit simplify_eq_implication)) uu___3))) - uu___2))) uu___2))) - uu___1))) uu___1))) uu___1) + uu___2))) + uu___2)))) uu___1))) + uu___1))) uu___1) let (rewrite_all_equalities : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = fun uu___ -> visit simplify_eq_implication @@ -856,49 +862,52 @@ let rec (unfold_definition_and_simplify_eq : match r with | FStar_Pervasives_Native.None -> Obj.magic - (FStar_Tactics_Derived.fail - "Not an equality implication") + (Obj.repr + (FStar_Tactics_Derived.fail + "Not an equality implication")) | FStar_Pervasives_Native.Some (uu___2, rhs) -> Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.Logic.fst" - (Prims.of_int (180)) - (Prims.of_int (23)) - (Prims.of_int (180)) - (Prims.of_int (39))) - (Prims.mk_range - "FStar.Tactics.Logic.fst" - (Prims.of_int (181)) - (Prims.of_int (12)) - (Prims.of_int (183)) - (Prims.of_int (66))) - (Obj.magic - (implies_intro ())) - (fun uu___3 -> - (fun eq_h -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.Logic.fst" - (Prims.of_int (181)) - (Prims.of_int (12)) - (Prims.of_int (181)) - (Prims.of_int (24))) - (Prims.mk_range - "FStar.Tactics.Logic.fst" - (Prims.of_int (182)) - (Prims.of_int (12)) - (Prims.of_int (183)) - (Prims.of_int (66))) - (Obj.magic - (FStar_Tactics_Builtins.rewrite - eq_h)) - (fun uu___3 -> + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.Logic.fst" + (Prims.of_int (180)) + (Prims.of_int (23)) + (Prims.of_int (180)) + (Prims.of_int (39))) + (Prims.mk_range + "FStar.Tactics.Logic.fst" + (Prims.of_int (181)) + (Prims.of_int (12)) + (Prims.of_int (183)) + (Prims.of_int (66))) + (Obj.magic + (implies_intro ())) + (fun uu___3 -> + (fun eq_h -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.Logic.fst" + (Prims.of_int (181)) + (Prims.of_int (12)) + (Prims.of_int (181)) + (Prims.of_int (24))) + (Prims.mk_range + "FStar.Tactics.Logic.fst" + (Prims.of_int (182)) + (Prims.of_int (12)) + (Prims.of_int (183)) + (Prims.of_int (66))) + (Obj.magic + (FStar_Tactics_Builtins.rewrite + eq_h)) (fun uu___3 -> - Obj.magic + (fun + uu___3 -> + Obj.magic (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.Logic.fst" @@ -926,8 +935,8 @@ let rec (unfold_definition_and_simplify_eq : unfold_definition_and_simplify_eq tm))) uu___4))) - uu___3))) - uu___3))) uu___2)))) + uu___3))) + uu___3)))) uu___2)))) uu___))) uu___) let (unsquash : FStar_Reflection_Types.term -> @@ -1238,7 +1247,10 @@ let (instantiate : (fa, FStar_Reflection_Data.Q_Explicit)))), (x, FStar_Reflection_Data.Q_Explicit))))) (fun uu___1 -> - FStar_Tactics_Derived.fail "could not instantiate")) + (fun uu___1 -> + Obj.magic + (FStar_Tactics_Derived.fail "could not instantiate")) + uu___1)) let (instantiate_as : FStar_Reflection_Types.term -> FStar_Reflection_Types.term -> @@ -1361,19 +1373,12 @@ let rec (sk_binder' : (fun uu___4 -> uu___3 <> Prims.int_one)))) (fun uu___3 -> - (fun uu___3 -> - if uu___3 - then - Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "no")) - else - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___5 -> ())))) - uu___3))) + if uu___3 + then + FStar_Tactics_Derived.fail "no" + else + FStar_Tactics_Effect.lift_div_tac + (fun uu___5 -> ())))) (fun uu___3 -> (fun uu___3 -> Obj.magic @@ -1568,5 +1573,7 @@ let (using_lemma : "lem3_fa"]))), (t, FStar_Reflection_Data.Q_Explicit))))) (fun uu___2 -> - FStar_Tactics_Derived.fail - "using_lemma: failed to instantiate"))) \ No newline at end of file + (fun uu___2 -> + Obj.magic + (FStar_Tactics_Derived.fail + "using_lemma: failed to instantiate")) uu___2))) \ No newline at end of file diff --git a/ocaml/fstar-lib/generated/FStar_Tactics_PatternMatching.ml b/ocaml/fstar-lib/generated/FStar_Tactics_PatternMatching.ml index de9f29a87f5..976c932a28e 100644 --- a/ocaml/fstar-lib/generated/FStar_Tactics_PatternMatching.ml +++ b/ocaml/fstar-lib/generated/FStar_Tactics_PatternMatching.ml @@ -28,73 +28,79 @@ let (fetch_eq_side : | FStar_Reflection_Data.Tv_App (squash, (g1, uu___2)) -> Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.PatternMatching.fst" - (Prims.of_int (66)) (Prims.of_int (11)) - (Prims.of_int (66)) (Prims.of_int (25))) - (Prims.mk_range - "FStar.Tactics.PatternMatching.fst" - (Prims.of_int (66)) (Prims.of_int (4)) - (Prims.of_int (87)) (Prims.of_int (51))) - (Obj.magic - (FStar_Tactics_Builtins.inspect squash)) - (fun uu___3 -> + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.PatternMatching.fst" + (Prims.of_int (66)) (Prims.of_int (11)) + (Prims.of_int (66)) (Prims.of_int (25))) + (Prims.mk_range + "FStar.Tactics.PatternMatching.fst" + (Prims.of_int (66)) (Prims.of_int (4)) + (Prims.of_int (87)) (Prims.of_int (51))) + (Obj.magic + (FStar_Tactics_Builtins.inspect squash)) (fun uu___3 -> - match uu___3 with - | FStar_Reflection_Data.Tv_UInst - (squash1, uu___4) -> - if - (FStar_Reflection_Derived.fv_to_string - squash1) - = - (FStar_Reflection_Derived.flatten_name - FStar_Reflection_Const.squash_qn) - then - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.PatternMatching.fst" - (Prims.of_int (70)) - (Prims.of_int (16)) - (Prims.of_int (70)) - (Prims.of_int (25))) - (Prims.mk_range - "FStar.Tactics.PatternMatching.fst" - (Prims.of_int (70)) - (Prims.of_int (9)) - (Prims.of_int (85)) - (Prims.of_int (48))) - (Obj.magic - (FStar_Tactics_Builtins.inspect - g1)) - (fun uu___5 -> - (fun uu___5 -> - match uu___5 with - | FStar_Reflection_Data.Tv_App - (eq_type_x, - (y, uu___6)) - -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.PatternMatching.fst" - (Prims.of_int (72)) - (Prims.of_int (19)) - (Prims.of_int (72)) - (Prims.of_int (36))) - (Prims.mk_range - "FStar.Tactics.PatternMatching.fst" - (Prims.of_int (72)) - (Prims.of_int (12)) - (Prims.of_int (84)) - (Prims.of_int (39))) - (Obj.magic - (FStar_Tactics_Builtins.inspect + (fun uu___3 -> + match uu___3 with + | FStar_Reflection_Data.Tv_UInst + (squash1, uu___4) -> + Obj.magic + (Obj.repr + (if + (FStar_Reflection_Derived.fv_to_string + squash1) + = + (FStar_Reflection_Derived.flatten_name + FStar_Reflection_Const.squash_qn) + then + Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.PatternMatching.fst" + (Prims.of_int (70)) + (Prims.of_int (16)) + (Prims.of_int (70)) + (Prims.of_int (25))) + (Prims.mk_range + "FStar.Tactics.PatternMatching.fst" + (Prims.of_int (70)) + (Prims.of_int (9)) + (Prims.of_int (85)) + (Prims.of_int (48))) + (Obj.magic + (FStar_Tactics_Builtins.inspect + g1)) + (fun uu___5 -> + (fun uu___5 -> + match uu___5 + with + | FStar_Reflection_Data.Tv_App + (eq_type_x, + (y, + uu___6)) + -> + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.PatternMatching.fst" + (Prims.of_int (72)) + (Prims.of_int (19)) + (Prims.of_int (72)) + (Prims.of_int (36))) + (Prims.mk_range + "FStar.Tactics.PatternMatching.fst" + (Prims.of_int (72)) + (Prims.of_int (12)) + (Prims.of_int (84)) + (Prims.of_int (39))) + (Obj.magic + (FStar_Tactics_Builtins.inspect eq_type_x)) - (fun uu___7 - -> - (fun + (fun + uu___7 -> + (fun uu___7 -> match uu___7 with @@ -105,6 +111,7 @@ let (fetch_eq_side : uu___8)) -> Obj.magic + (Obj.repr (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.PatternMatching.fst" @@ -134,6 +141,7 @@ let (fetch_eq_side : uu___10)) -> Obj.magic + (Obj.repr (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.PatternMatching.fst" @@ -153,9 +161,6 @@ let (fetch_eq_side : (fun uu___11 -> - (fun - uu___11 - -> match uu___11 with | @@ -163,126 +168,123 @@ let (fetch_eq_side : (eq1, uu___12) -> - Obj.magic - (Obj.repr - (if + if (FStar_Reflection_Derived.fv_to_string eq1) = (FStar_Reflection_Derived.flatten_name FStar_Reflection_Const.eq2_qn) then - Obj.repr - (FStar_Tactics_Effect.lift_div_tac + FStar_Tactics_Effect.lift_div_tac (fun uu___13 - -> (x, y))) + -> (x, y)) else - Obj.repr - (FStar_Tactics_Derived.fail - "not an equality"))) + FStar_Tactics_Derived.fail + "not an equality" | FStar_Reflection_Data.Tv_FVar eq1 -> - Obj.magic - (Obj.repr - (if + if (FStar_Reflection_Derived.fv_to_string eq1) = (FStar_Reflection_Derived.flatten_name FStar_Reflection_Const.eq2_qn) then - Obj.repr - (FStar_Tactics_Effect.lift_div_tac + FStar_Tactics_Effect.lift_div_tac (fun uu___12 - -> (x, y))) + -> (x, y)) else - Obj.repr - (FStar_Tactics_Derived.fail - "not an equality"))) + FStar_Tactics_Derived.fail + "not an equality" | uu___12 -> - Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail + FStar_Tactics_Derived.fail "not an app2 of fvar: "))) - uu___11)) | uu___10 -> Obj.magic + (Obj.repr (FStar_Tactics_Derived.fail - "not an app3")) - uu___9)) + "not an app3"))) + uu___9))) | uu___8 -> Obj.magic + (Obj.repr (FStar_Tactics_Derived.fail - "not an app2")) - uu___7)) - | uu___6 -> - Obj.magic - (FStar_Tactics_Derived.fail - "not an app under squash")) - uu___5)) - else - Obj.magic - (FStar_Tactics_Derived.fail - "not a squash") - | FStar_Reflection_Data.Tv_FVar squash1 - -> - if - (FStar_Reflection_Derived.fv_to_string - squash1) - = - (FStar_Reflection_Derived.flatten_name - FStar_Reflection_Const.squash_qn) - then - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.PatternMatching.fst" - (Prims.of_int (70)) - (Prims.of_int (16)) - (Prims.of_int (70)) - (Prims.of_int (25))) - (Prims.mk_range - "FStar.Tactics.PatternMatching.fst" - (Prims.of_int (70)) - (Prims.of_int (9)) - (Prims.of_int (85)) - (Prims.of_int (48))) - (Obj.magic - (FStar_Tactics_Builtins.inspect - g1)) - (fun uu___4 -> - (fun uu___4 -> - match uu___4 with - | FStar_Reflection_Data.Tv_App - (eq_type_x, - (y, uu___5)) - -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.PatternMatching.fst" - (Prims.of_int (72)) - (Prims.of_int (19)) - (Prims.of_int (72)) - (Prims.of_int (36))) - (Prims.mk_range - "FStar.Tactics.PatternMatching.fst" - (Prims.of_int (72)) - (Prims.of_int (12)) - (Prims.of_int (84)) - (Prims.of_int (39))) - (Obj.magic - (FStar_Tactics_Builtins.inspect + "not an app2"))) + uu___7))) + | uu___6 -> + Obj.magic + (Obj.repr + (FStar_Tactics_Derived.fail + "not an app under squash"))) + uu___5)) + else + Obj.repr + (FStar_Tactics_Derived.fail + "not a squash"))) + | FStar_Reflection_Data.Tv_FVar + squash1 -> + Obj.magic + (Obj.repr + (if + (FStar_Reflection_Derived.fv_to_string + squash1) + = + (FStar_Reflection_Derived.flatten_name + FStar_Reflection_Const.squash_qn) + then + Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.PatternMatching.fst" + (Prims.of_int (70)) + (Prims.of_int (16)) + (Prims.of_int (70)) + (Prims.of_int (25))) + (Prims.mk_range + "FStar.Tactics.PatternMatching.fst" + (Prims.of_int (70)) + (Prims.of_int (9)) + (Prims.of_int (85)) + (Prims.of_int (48))) + (Obj.magic + (FStar_Tactics_Builtins.inspect + g1)) + (fun uu___4 -> + (fun uu___4 -> + match uu___4 + with + | FStar_Reflection_Data.Tv_App + (eq_type_x, + (y, + uu___5)) + -> + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.PatternMatching.fst" + (Prims.of_int (72)) + (Prims.of_int (19)) + (Prims.of_int (72)) + (Prims.of_int (36))) + (Prims.mk_range + "FStar.Tactics.PatternMatching.fst" + (Prims.of_int (72)) + (Prims.of_int (12)) + (Prims.of_int (84)) + (Prims.of_int (39))) + (Obj.magic + (FStar_Tactics_Builtins.inspect eq_type_x)) - (fun uu___6 - -> - (fun + (fun + uu___6 -> + (fun uu___6 -> match uu___6 with @@ -293,6 +295,7 @@ let (fetch_eq_side : uu___7)) -> Obj.magic + (Obj.repr (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.PatternMatching.fst" @@ -322,6 +325,7 @@ let (fetch_eq_side : uu___9)) -> Obj.magic + (Obj.repr (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.PatternMatching.fst" @@ -341,9 +345,6 @@ let (fetch_eq_side : (fun uu___10 -> - (fun - uu___10 - -> match uu___10 with | @@ -351,81 +352,76 @@ let (fetch_eq_side : (eq1, uu___11) -> - Obj.magic - (Obj.repr - (if + if (FStar_Reflection_Derived.fv_to_string eq1) = (FStar_Reflection_Derived.flatten_name FStar_Reflection_Const.eq2_qn) then - Obj.repr - (FStar_Tactics_Effect.lift_div_tac + FStar_Tactics_Effect.lift_div_tac (fun uu___12 - -> (x, y))) + -> (x, y)) else - Obj.repr - (FStar_Tactics_Derived.fail - "not an equality"))) + FStar_Tactics_Derived.fail + "not an equality" | FStar_Reflection_Data.Tv_FVar eq1 -> - Obj.magic - (Obj.repr - (if + if (FStar_Reflection_Derived.fv_to_string eq1) = (FStar_Reflection_Derived.flatten_name FStar_Reflection_Const.eq2_qn) then - Obj.repr - (FStar_Tactics_Effect.lift_div_tac + FStar_Tactics_Effect.lift_div_tac (fun uu___11 - -> (x, y))) + -> (x, y)) else - Obj.repr - (FStar_Tactics_Derived.fail - "not an equality"))) + FStar_Tactics_Derived.fail + "not an equality" | uu___11 -> - Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail + FStar_Tactics_Derived.fail "not an app2 of fvar: "))) - uu___10)) | uu___9 -> Obj.magic + (Obj.repr (FStar_Tactics_Derived.fail - "not an app3")) - uu___8)) + "not an app3"))) + uu___8))) | uu___7 -> Obj.magic + (Obj.repr (FStar_Tactics_Derived.fail - "not an app2")) - uu___6)) - | uu___5 -> - Obj.magic - (FStar_Tactics_Derived.fail - "not an app under squash")) - uu___4)) - else - Obj.magic - (FStar_Tactics_Derived.fail - "not a squash") - | uu___4 -> - Obj.magic - (FStar_Tactics_Derived.fail - "not an app of fvar at top level")) - uu___3)) + "not an app2"))) + uu___6))) + | uu___5 -> + Obj.magic + (Obj.repr + (FStar_Tactics_Derived.fail + "not an app under squash"))) + uu___4)) + else + Obj.repr + (FStar_Tactics_Derived.fail + "not a squash"))) + | uu___4 -> + Obj.magic + (Obj.repr + (FStar_Tactics_Derived.fail + "not an app of fvar at top level"))) + uu___3))) | uu___2 -> Obj.magic - (FStar_Tactics_Derived.fail - "not an app at top level")) uu___1))) uu___1) + (Obj.repr + (FStar_Tactics_Derived.fail + "not an app at top level"))) uu___1))) + uu___1) let mustfail : 'a . (unit -> ('a, unit) FStar_Tactics_Effect.tac_repr) -> @@ -441,15 +437,11 @@ let mustfail : (Prims.of_int (130)) (Prims.of_int (4)) (Prims.of_int (132)) (Prims.of_int (16))) (Obj.magic (FStar_Tactics_Derived.trytac t)) (fun uu___ -> - (fun uu___ -> - match uu___ with - | FStar_Pervasives_Native.Some uu___1 -> - Obj.magic (Obj.repr (FStar_Tactics_Derived.fail message)) - | FStar_Pervasives_Native.None -> - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac (fun uu___1 -> ())))) - uu___) + match uu___ with + | FStar_Pervasives_Native.Some uu___1 -> + FStar_Tactics_Derived.fail message + | FStar_Pervasives_Native.None -> + FStar_Tactics_Effect.lift_div_tac (fun uu___1 -> ())) let (implies_intro' : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = fun uu___ -> FStar_Tactics_Effect.tac_bind @@ -1045,10 +1037,8 @@ let lift_exn_tac : (Prims.of_int (268)) (Prims.of_int (18)) (Prims.of_int (268)) (Prims.of_int (61))) (Obj.magic (string_of_match_exception ex)) - (fun uu___ -> - (fun uu___ -> - Obj.magic (FStar_Tactics_Derived.fail uu___)) - uu___)))) uu___1 uu___ + (fun uu___ -> FStar_Tactics_Derived.fail uu___)))) + uu___1 uu___ let lift_exn_tactic : 'a 'b . ('a -> 'b match_res) -> 'a -> ('b, unit) FStar_Tactics_Effect.tac_repr @@ -1073,10 +1063,8 @@ let lift_exn_tactic : (Prims.of_int (273)) (Prims.of_int (18)) (Prims.of_int (273)) (Prims.of_int (61))) (Obj.magic (string_of_match_exception ex)) - (fun uu___ -> - (fun uu___ -> - Obj.magic (FStar_Tactics_Derived.fail uu___)) - uu___)))) uu___1 uu___ + (fun uu___ -> FStar_Tactics_Derived.fail uu___)))) + uu___1 uu___ type bindings = (varname * FStar_Reflection_Types.term) Prims.list let (string_of_bindings : bindings -> (Prims.string, unit) FStar_Tactics_Effect.tac_repr) = @@ -1496,10 +1484,8 @@ let (match_term : (Prims.of_int (338)) (Prims.of_int (20)) (Prims.of_int (338)) (Prims.of_int (63))) (Obj.magic (string_of_match_exception ex)) - (fun uu___1 -> - (fun uu___1 -> - Obj.magic (FStar_Tactics_Derived.fail uu___1)) - uu___1)))) uu___) + (fun uu___1 -> FStar_Tactics_Derived.fail uu___1)))) + uu___) let debug : 'uuuuu . 'uuuuu -> (unit, unit) FStar_Tactics_Effect.tac_repr = fun uu___ -> (fun msg -> @@ -1663,13 +1649,9 @@ let assoc_varname_fail : match FStar_List_Tot_Base.assoc key ls with | FStar_Pervasives_Native.None -> Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - (Prims.strcat "Not found: " key))) + (FStar_Tactics_Derived.fail (Prims.strcat "Not found: " key)) | FStar_Pervasives_Native.Some x -> - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac (fun uu___ -> x)))) + Obj.magic (FStar_Tactics_Effect.lift_div_tac (fun uu___ -> x))) uu___1 uu___ let ms_locate_hyp : 'a . @@ -1712,98 +1694,123 @@ let rec solve_mp_for_single_hyp : (matching_solution -> ('a, unit) FStar_Tactics_Effect.tac_repr) -> matching_solution -> ('a, unit) FStar_Tactics_Effect.tac_repr = - fun name -> - fun pat -> - fun hypotheses -> - fun body -> - fun part_sol -> - match hypotheses with - | [] -> FStar_Tactics_Derived.fail "No matching hypothesis" - | h::hs -> - FStar_Tactics_Derived.or_else - (fun uu___ -> - FStar_Tactics_Effect.tac_bind - (Prims.mk_range "FStar.Tactics.PatternMatching.fst" - (Prims.of_int (448)) (Prims.of_int (15)) - (Prims.of_int (448)) (Prims.of_int (73))) - (Prims.mk_range "FStar.Tactics.PatternMatching.fst" - (Prims.of_int (448)) (Prims.of_int (9)) - (Prims.of_int (453)) (Prims.of_int (73))) - (Obj.magic - (interp_pattern_aux pat part_sol.ms_vars - (FStar_Reflection_Derived.type_of_binder h))) - (fun uu___1 -> - (fun uu___1 -> - match uu___1 with - | Failure ex -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.PatternMatching.fst" - (Prims.of_int (450)) - (Prims.of_int (16)) - (Prims.of_int (450)) - (Prims.of_int (74))) - (Prims.mk_range - "FStar.Tactics.PatternMatching.fst" - (Prims.of_int (450)) - (Prims.of_int (11)) - (Prims.of_int (450)) - (Prims.of_int (74))) - (Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.PatternMatching.fst" - (Prims.of_int (450)) - (Prims.of_int (43)) - (Prims.of_int (450)) - (Prims.of_int (73))) - (Prims.mk_range "prims.fst" - (Prims.of_int (606)) - (Prims.of_int (19)) - (Prims.of_int (606)) - (Prims.of_int (31))) - (Obj.magic - (string_of_match_exception ex)) - (fun uu___2 -> - FStar_Tactics_Effect.lift_div_tac - (fun uu___3 -> - Prims.strcat - "Failed to match hyp: " - uu___2)))) - (fun uu___2 -> - (fun uu___2 -> - Obj.magic - (FStar_Tactics_Derived.fail - uu___2)) uu___2)) - | Success bindings1 -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.PatternMatching.fst" - (Prims.of_int (452)) - (Prims.of_int (35)) - (Prims.of_int (452)) - (Prims.of_int (37))) - (Prims.mk_range - "FStar.Tactics.PatternMatching.fst" - (Prims.of_int (453)) - (Prims.of_int (11)) - (Prims.of_int (453)) - (Prims.of_int (73))) - (FStar_Tactics_Effect.lift_div_tac - (fun uu___2 -> (name, h) :: - (part_sol.ms_hyps))) - (fun uu___2 -> - (fun ms_hyps -> - Obj.magic - (body - { - ms_vars = bindings1; - ms_hyps - })) uu___2))) uu___1)) - (fun uu___ -> - solve_mp_for_single_hyp name pat hs body part_sol) + fun uu___4 -> + fun uu___3 -> + fun uu___2 -> + fun uu___1 -> + fun uu___ -> + (fun name -> + fun pat -> + fun hypotheses -> + fun body -> + fun part_sol -> + match hypotheses with + | [] -> + Obj.magic + (Obj.repr + (FStar_Tactics_Derived.fail + "No matching hypothesis")) + | h::hs -> + Obj.magic + (Obj.repr + (FStar_Tactics_Derived.or_else + (fun uu___ -> + FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.PatternMatching.fst" + (Prims.of_int (448)) + (Prims.of_int (15)) + (Prims.of_int (448)) + (Prims.of_int (73))) + (Prims.mk_range + "FStar.Tactics.PatternMatching.fst" + (Prims.of_int (448)) + (Prims.of_int (9)) + (Prims.of_int (453)) + (Prims.of_int (73))) + (Obj.magic + (interp_pattern_aux pat + part_sol.ms_vars + (FStar_Reflection_Derived.type_of_binder + h))) + (fun uu___1 -> + (fun uu___1 -> + match uu___1 with + | Failure ex -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.PatternMatching.fst" + (Prims.of_int (450)) + (Prims.of_int (16)) + (Prims.of_int (450)) + (Prims.of_int (74))) + (Prims.mk_range + "FStar.Tactics.PatternMatching.fst" + (Prims.of_int (450)) + (Prims.of_int (11)) + (Prims.of_int (450)) + (Prims.of_int (74))) + (Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.PatternMatching.fst" + (Prims.of_int (450)) + (Prims.of_int (43)) + (Prims.of_int (450)) + (Prims.of_int (73))) + (Prims.mk_range + "prims.fst" + (Prims.of_int (606)) + (Prims.of_int (19)) + (Prims.of_int (606)) + (Prims.of_int (31))) + (Obj.magic + (string_of_match_exception + ex)) + (fun uu___2 -> + FStar_Tactics_Effect.lift_div_tac + (fun uu___3 + -> + Prims.strcat + "Failed to match hyp: " + uu___2)))) + (fun uu___2 -> + FStar_Tactics_Derived.fail + uu___2)) + | Success bindings1 -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.PatternMatching.fst" + (Prims.of_int (452)) + (Prims.of_int (35)) + (Prims.of_int (452)) + (Prims.of_int (37))) + (Prims.mk_range + "FStar.Tactics.PatternMatching.fst" + (Prims.of_int (453)) + (Prims.of_int (11)) + (Prims.of_int (453)) + (Prims.of_int (73))) + (FStar_Tactics_Effect.lift_div_tac + (fun uu___2 -> + (name, h) :: + (part_sol.ms_hyps))) + (fun uu___2 -> + (fun ms_hyps -> + Obj.magic + (body + { + ms_vars = + bindings1; + ms_hyps + })) uu___2))) + uu___1)) + (fun uu___ -> + solve_mp_for_single_hyp name pat hs + body part_sol)))) uu___4 uu___3 + uu___2 uu___1 uu___ let rec solve_mp_for_hyps : 'a . (varname * pattern) Prims.list -> @@ -1898,10 +1905,8 @@ let solve_mp : "Failed to match goal: " uu___1)))) (fun uu___1 -> - (fun uu___1 -> - Obj.magic - (FStar_Tactics_Derived.fail - uu___1)) uu___1))) + FStar_Tactics_Derived.fail + uu___1))) | Success bindings1 -> Obj.magic (Obj.repr @@ -2042,10 +2047,8 @@ let (pattern_of_term : (Prims.of_int (532)) (Prims.of_int (20)) (Prims.of_int (532)) (Prims.of_int (63))) (Obj.magic (string_of_match_exception ex)) - (fun uu___1 -> - (fun uu___1 -> - Obj.magic (FStar_Tactics_Derived.fail uu___1)) - uu___1)))) uu___) + (fun uu___1 -> FStar_Tactics_Derived.fail uu___1)))) + uu___) type 'a hyp = FStar_Reflection_Types.binder type 'a pm_goal = unit let (hyp_qn : Prims.string) = "FStar.Tactics.PatternMatching.hyp" @@ -2205,8 +2208,6 @@ let (classify_abspat_binder : ( fun uu___2 -> - (fun - uu___2 -> match uu___2 with | @@ -2214,32 +2215,25 @@ let (classify_abspat_binder : ((uu___3, goal_typ)::[]) -> - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac + FStar_Tactics_Effect.lift_div_tac (fun uu___4 -> (ABKGoal, - goal_typ)))) + goal_typ)) | Success uu___3 -> - Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "classifiy_abspat_binder: impossible (2)")) + FStar_Tactics_Derived.fail + "classifiy_abspat_binder: impossible (2)" | Failure uu___3 -> - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac + FStar_Tactics_Effect.lift_div_tac (fun uu___4 -> ((ABKVar typ), - typ))))) - uu___2)))) + typ)))))) uu___))) uu___))) uu___))) uu___))) uu___) let rec (binders_and_body_of_abs : diff --git a/ocaml/fstar-lib/generated/FStar_Tactics_Simplifier.ml b/ocaml/fstar-lib/generated/FStar_Tactics_Simplifier.ml index d3708421bb0..e6d04f4ca30 100644 --- a/ocaml/fstar-lib/generated/FStar_Tactics_Simplifier.ml +++ b/ocaml/fstar-lib/generated/FStar_Tactics_Simplifier.ml @@ -144,45 +144,55 @@ let (inhabit : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = match FStar_Reflection_Builtins.inspect_ln t with | FStar_Reflection_Data.Tv_FVar fv -> Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "FStar.Tactics.Simplifier.fst" - (Prims.of_int (200)) (Prims.of_int (17)) - (Prims.of_int (200)) (Prims.of_int (30))) - (Prims.mk_range "FStar.Tactics.Simplifier.fst" - (Prims.of_int (201)) (Prims.of_int (13)) - (Prims.of_int (204)) (Prims.of_int (20))) - (FStar_Tactics_Effect.lift_div_tac + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "FStar.Tactics.Simplifier.fst" + (Prims.of_int (200)) (Prims.of_int (17)) + (Prims.of_int (200)) (Prims.of_int (30))) + (Prims.mk_range "FStar.Tactics.Simplifier.fst" + (Prims.of_int (201)) (Prims.of_int (13)) + (Prims.of_int (204)) (Prims.of_int (20))) + (FStar_Tactics_Effect.lift_div_tac + (fun uu___1 -> + FStar_Reflection_Builtins.inspect_fv fv)) (fun uu___1 -> - FStar_Reflection_Builtins.inspect_fv fv)) - (fun uu___1 -> - (fun qn -> - if qn = FStar_Reflection_Const.int_lid - then - Obj.magic - (FStar_Tactics_Derived.exact - (FStar_Reflection_Builtins.pack_ln - (FStar_Reflection_Data.Tv_Const - (FStar_Reflection_Data.C_Int - (Prims.of_int (42)))))) - else - if qn = FStar_Reflection_Const.bool_lid - then - Obj.magic - (FStar_Tactics_Derived.exact - (FStar_Reflection_Builtins.pack_ln - (FStar_Reflection_Data.Tv_Const - FStar_Reflection_Data.C_True))) - else - if qn = FStar_Reflection_Const.unit_lid - then - Obj.magic - (FStar_Tactics_Derived.exact - (FStar_Reflection_Builtins.pack_ln - (FStar_Reflection_Data.Tv_Const - FStar_Reflection_Data.C_Unit))) - else Obj.magic (FStar_Tactics_Derived.fail "")) - uu___1)) - | uu___1 -> Obj.magic (FStar_Tactics_Derived.fail "")) uu___1) + (fun qn -> + if qn = FStar_Reflection_Const.int_lid + then + Obj.magic + (Obj.repr + (FStar_Tactics_Derived.exact + (FStar_Reflection_Builtins.pack_ln + (FStar_Reflection_Data.Tv_Const + (FStar_Reflection_Data.C_Int + (Prims.of_int (42))))))) + else + Obj.magic + (Obj.repr + (if qn = FStar_Reflection_Const.bool_lid + then + Obj.repr + (FStar_Tactics_Derived.exact + (FStar_Reflection_Builtins.pack_ln + (FStar_Reflection_Data.Tv_Const + FStar_Reflection_Data.C_True))) + else + Obj.repr + (if + qn = + FStar_Reflection_Const.unit_lid + then + Obj.repr + (FStar_Tactics_Derived.exact + (FStar_Reflection_Builtins.pack_ln + (FStar_Reflection_Data.Tv_Const + FStar_Reflection_Data.C_Unit))) + else + Obj.repr + (FStar_Tactics_Derived.fail ""))))) + uu___1))) + | uu___1 -> Obj.magic (Obj.repr (FStar_Tactics_Derived.fail ""))) + uu___1) let rec (simplify_point : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = fun uu___ -> @@ -239,29 +249,31 @@ let rec (simplify_point : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) | FStar_Reflection_Formula.Iff (l, r) -> Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.Simplifier.fst" - (Prims.of_int (217)) - (Prims.of_int (20)) - (Prims.of_int (217)) - (Prims.of_int (38))) - (Prims.mk_range - "FStar.Tactics.Simplifier.fst" - (Prims.of_int (217)) - (Prims.of_int (14)) - (Prims.of_int (264)) - (Prims.of_int (22))) - (Obj.magic - (FStar_Reflection_Formula.term_as_formula' - l)) - (fun uu___3 -> + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.Simplifier.fst" + (Prims.of_int (217)) + (Prims.of_int (20)) + (Prims.of_int (217)) + (Prims.of_int (38))) + (Prims.mk_range + "FStar.Tactics.Simplifier.fst" + (Prims.of_int (217)) + (Prims.of_int (14)) + (Prims.of_int (264)) + (Prims.of_int (22))) + (Obj.magic + (FStar_Reflection_Formula.term_as_formula' + l)) (fun uu___3 -> - match uu___3 with - | FStar_Reflection_Formula.And - (p, q) -> - Obj.magic - (FStar_Tactics_Effect.tac_bind + (fun uu___3 -> + match uu___3 + with + | FStar_Reflection_Formula.And + (p, q) -> + Obj.magic + (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.Simplifier.fst" (Prims.of_int (219)) @@ -401,10 +413,10 @@ let rec (simplify_point : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) uu___8))) uu___6))) uu___4)) - | FStar_Reflection_Formula.Or - (p, q) -> - Obj.magic - (FStar_Tactics_Effect.tac_bind + | FStar_Reflection_Formula.Or + (p, q) -> + Obj.magic + (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.Simplifier.fst" (Prims.of_int (226)) @@ -544,10 +556,10 @@ let rec (simplify_point : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) uu___8))) uu___6))) uu___4)) - | FStar_Reflection_Formula.Implies - (p, q) -> - Obj.magic - (FStar_Tactics_Effect.tac_bind + | FStar_Reflection_Formula.Implies + (p, q) -> + Obj.magic + (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.Simplifier.fst" (Prims.of_int (233)) @@ -650,10 +662,10 @@ let rec (simplify_point : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) uu___8))) uu___6))) uu___4)) - | FStar_Reflection_Formula.Forall - (b, p) -> - Obj.magic - (FStar_Tactics_Effect.tac_bind + | FStar_Reflection_Formula.Forall + (b, p) -> + Obj.magic + (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.Simplifier.fst" (Prims.of_int (239)) @@ -748,10 +760,10 @@ let rec (simplify_point : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) (tiff ())) uu___6))) uu___4)) - | FStar_Reflection_Formula.Exists - (b, p) -> - Obj.magic - (FStar_Tactics_Effect.tac_bind + | FStar_Reflection_Formula.Exists + (b, p) -> + Obj.magic + (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.Simplifier.fst" (Prims.of_int (244)) @@ -846,10 +858,10 @@ let rec (simplify_point : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) (tiff ())) uu___6))) uu___4)) - | FStar_Reflection_Formula.Not - p -> - Obj.magic - (FStar_Tactics_Effect.tac_bind + | FStar_Reflection_Formula.Not + p -> + Obj.magic + (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.Simplifier.fst" (Prims.of_int (249)) @@ -918,10 +930,10 @@ let rec (simplify_point : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) (tiff ())) uu___6))) uu___4)) - | FStar_Reflection_Formula.Iff - (p, q) -> - Obj.magic - (FStar_Tactics_Effect.tac_bind + | FStar_Reflection_Formula.Iff + (p, q) -> + Obj.magic + (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.Simplifier.fst" (Prims.of_int (256)) @@ -1104,14 +1116,15 @@ let rec (simplify_point : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) ())) uu___5))) uu___4)) - | uu___4 -> - Obj.magic - (tiff ())) - uu___3)) + | uu___4 -> + Obj.magic + (tiff ())) + uu___3))) | uu___3 -> Obj.magic - (FStar_Tactics_Derived.fail - "simplify_point: failed precondition: goal should be `g <==> ?u`")) + (Obj.repr + (FStar_Tactics_Derived.fail + "simplify_point: failed precondition: goal should be `g <==> ?u`"))) uu___3))) uu___3))) uu___2))) uu___1) and (recurse : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = @@ -1169,31 +1182,33 @@ and (recurse : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = | FStar_Reflection_Formula.Iff (l, r) -> Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "FStar.Tactics.Simplifier.fst" - (Prims.of_int (275)) - (Prims.of_int (20)) - (Prims.of_int (275)) - (Prims.of_int (38))) - (Prims.mk_range - "FStar.Tactics.Simplifier.fst" - (Prims.of_int (275)) - (Prims.of_int (14)) - (Prims.of_int (302)) - (Prims.of_int (22))) - (Obj.magic - (FStar_Reflection_Formula.term_as_formula' - l)) - (fun uu___3 -> + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "FStar.Tactics.Simplifier.fst" + (Prims.of_int (275)) + (Prims.of_int (20)) + (Prims.of_int (275)) + (Prims.of_int (38))) + (Prims.mk_range + "FStar.Tactics.Simplifier.fst" + (Prims.of_int (275)) + (Prims.of_int (14)) + (Prims.of_int (302)) + (Prims.of_int (22))) + (Obj.magic + (FStar_Reflection_Formula.term_as_formula' + l)) (fun uu___3 -> - match uu___3 with - | FStar_Reflection_Formula.And - (uu___4, - uu___5) - -> - Obj.magic - (FStar_Tactics_Derived.seq + (fun uu___3 -> + match uu___3 + with + | FStar_Reflection_Formula.And + (uu___4, + uu___5) + -> + Obj.magic + (FStar_Tactics_Derived.seq (fun uu___6 -> FStar_Tactics_Derived.apply_lemma @@ -1205,12 +1220,12 @@ and (recurse : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = "Simplifier"; "and_cong"])))) simplify_point) - | FStar_Reflection_Formula.Or - (uu___4, - uu___5) - -> - Obj.magic - (FStar_Tactics_Derived.seq + | FStar_Reflection_Formula.Or + (uu___4, + uu___5) + -> + Obj.magic + (FStar_Tactics_Derived.seq (fun uu___6 -> FStar_Tactics_Derived.apply_lemma @@ -1222,12 +1237,12 @@ and (recurse : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = "Simplifier"; "or_cong"])))) simplify_point) - | FStar_Reflection_Formula.Implies - (uu___4, - uu___5) - -> - Obj.magic - (FStar_Tactics_Derived.seq + | FStar_Reflection_Formula.Implies + (uu___4, + uu___5) + -> + Obj.magic + (FStar_Tactics_Derived.seq (fun uu___6 -> FStar_Tactics_Derived.apply_lemma @@ -1239,12 +1254,12 @@ and (recurse : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = "Simplifier"; "imp_cong"])))) simplify_point) - | FStar_Reflection_Formula.Forall - (uu___4, - uu___5) - -> - Obj.magic - (FStar_Tactics_Effect.tac_bind + | FStar_Reflection_Formula.Forall + (uu___4, + uu___5) + -> + Obj.magic + (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.Simplifier.fst" (Prims.of_int (286)) @@ -1296,12 +1311,12 @@ and (recurse : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = ())) uu___7))) uu___6)) - | FStar_Reflection_Formula.Exists - (uu___4, - uu___5) - -> - Obj.magic - (FStar_Tactics_Effect.tac_bind + | FStar_Reflection_Formula.Exists + (uu___4, + uu___5) + -> + Obj.magic + (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.Simplifier.fst" (Prims.of_int (291)) @@ -1353,10 +1368,10 @@ and (recurse : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = ())) uu___7))) uu___6)) - | FStar_Reflection_Formula.Not - uu___4 -> - Obj.magic - (FStar_Tactics_Effect.tac_bind + | FStar_Reflection_Formula.Not + uu___4 -> + Obj.magic + (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.Simplifier.fst" (Prims.of_int (296)) @@ -1386,12 +1401,12 @@ and (recurse : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = (simplify_point ())) uu___5)) - | FStar_Reflection_Formula.Iff - (uu___4, - uu___5) - -> - Obj.magic - (FStar_Tactics_Derived.seq + | FStar_Reflection_Formula.Iff + (uu___4, + uu___5) + -> + Obj.magic + (FStar_Tactics_Derived.seq (fun uu___6 -> FStar_Tactics_Derived.apply_lemma @@ -1403,14 +1418,15 @@ and (recurse : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = "Simplifier"; "iff_cong"])))) simplify_point) - | uu___4 -> - Obj.magic - (tiff ())) - uu___3)) + | uu___4 -> + Obj.magic + (tiff ())) + uu___3))) | uu___3 -> Obj.magic - (FStar_Tactics_Derived.fail - "recurse: failed precondition: goal should be `g <==> ?u`")) + (Obj.repr + (FStar_Tactics_Derived.fail + "recurse: failed precondition: goal should be `g <==> ?u`"))) uu___3))) uu___3))) uu___2))) uu___1) let (simplify : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = diff --git a/ocaml/fstar-lib/generated/FStar_Tactics_SyntaxHelpers.ml b/ocaml/fstar-lib/generated/FStar_Tactics_SyntaxHelpers.ml index 6f6a6fa5759..805016faf18 100644 --- a/ocaml/fstar-lib/generated/FStar_Tactics_SyntaxHelpers.ml +++ b/ocaml/fstar-lib/generated/FStar_Tactics_SyntaxHelpers.ml @@ -123,7 +123,11 @@ let (collect_abs : match uu___ with | (bs, t') -> ((FStar_List_Tot_Base.rev bs), t'))) let fail : 'a . Prims.string -> ('a, unit) FStar_Tactics_Effect.tac_repr = - fun m -> FStar_Tactics_Effect.raise (FStar_Tactics_Common.TacticFailure m) + fun uu___ -> + (fun m -> + Obj.magic + (FStar_Tactics_Effect.raise (FStar_Tactics_Common.TacticFailure m))) + uu___ let rec (mk_arr : FStar_Reflection_Types.binder Prims.list -> FStar_Reflection_Types.comp -> diff --git a/ocaml/fstar-lib/generated/FStar_Tactics_Typeclasses.ml b/ocaml/fstar-lib/generated/FStar_Tactics_Typeclasses.ml index df399efc6e2..b296a249006 100644 --- a/ocaml/fstar-lib/generated/FStar_Tactics_Typeclasses.ml +++ b/ocaml/fstar-lib/generated/FStar_Tactics_Typeclasses.ml @@ -4,13 +4,18 @@ let rec first : ('a -> ('b, unit) FStar_Tactics_Effect.tac_repr) -> 'a Prims.list -> ('b, unit) FStar_Tactics_Effect.tac_repr = - fun f -> - fun l -> - match l with - | [] -> FStar_Tactics_Derived.fail "no cands" - | x::xs -> - FStar_Tactics_Derived.or_else (fun uu___ -> f x) - (fun uu___ -> first f xs) + fun uu___1 -> + fun uu___ -> + (fun f -> + fun l -> + match l with + | [] -> + Obj.magic (Obj.repr (FStar_Tactics_Derived.fail "no cands")) + | x::xs -> + Obj.magic + (Obj.repr + (FStar_Tactics_Derived.or_else (fun uu___ -> f x) + (fun uu___ -> first f xs)))) uu___1 uu___ let rec (tcresolve' : FStar_Reflection_Types.term Prims.list -> Prims.int -> (unit, unit) FStar_Tactics_Effect.tac_repr) @@ -23,10 +28,8 @@ let rec (tcresolve' : (Prims.mk_range "FStar.Tactics.Typeclasses.fst" (Prims.of_int (45)) (Prims.of_int (4)) (Prims.of_int (50)) (Prims.of_int (137))) (if fuel <= Prims.int_zero - then Obj.magic (Obj.repr (FStar_Tactics_Derived.fail "out of fuel")) - else - Obj.magic - (Obj.repr (FStar_Tactics_Effect.lift_div_tac (fun uu___1 -> ())))) + then FStar_Tactics_Derived.fail "out of fuel" + else FStar_Tactics_Effect.lift_div_tac (fun uu___1 -> ())) (fun uu___ -> (fun uu___ -> Obj.magic @@ -72,15 +75,10 @@ let rec (tcresolve' : (FStar_Reflection_Builtins.term_eq g) seen then - Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "loop")) + FStar_Tactics_Derived.fail "loop" else - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___3 -> ())))) + FStar_Tactics_Effect.lift_div_tac + (fun uu___3 -> ())) (fun uu___2 -> (fun uu___2 -> Obj.magic @@ -155,11 +153,7 @@ let rec (tcresolve' : uu___5)))) (fun uu___5 -> - (fun - uu___5 -> - Obj.magic - (FStar_Tactics_Derived.fail - uu___5)) + FStar_Tactics_Derived.fail uu___5))))) uu___3))) uu___2))) uu___2))) uu___1))) uu___) @@ -360,11 +354,15 @@ let (tcresolve : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = (fun uu___2 -> match () with | () -> tcresolve' [] (Prims.of_int (16))) (fun uu___2 -> - match uu___2 with - | FStar_Tactics_Common.TacticFailure s -> - FStar_Tactics_Derived.fail - (Prims.strcat "Typeclass resolution failed: " s) - | e -> FStar_Tactics_Effect.raise e))) uu___1) + (fun uu___2 -> + match uu___2 with + | FStar_Tactics_Common.TacticFailure s -> + Obj.magic + (FStar_Tactics_Derived.fail + (Prims.strcat "Typeclass resolution failed: " + s)) + | e -> Obj.magic (FStar_Tactics_Effect.raise e)) + uu___2))) uu___1) let _ = FStar_Tactics_Native.register_tactic "FStar.Tactics.Typeclasses.tcresolve" (Prims.of_int (2)) @@ -1100,11 +1098,14 @@ let (mk_class : FStar_Pervasives_Native.None -> Obj.magic + (Obj.repr (FStar_Tactics_Derived.fail - "mk_class: proj not found?") + "mk_class: proj not found?")) | FStar_Pervasives_Native.Some se1 -> + Obj.magic + (Obj.repr (match FStar_Reflection_Builtins.inspect_sigelt se1 @@ -1113,7 +1114,7 @@ let (mk_class : FStar_Reflection_Data.Sg_Let (uu___10, lbs) -> - Obj.magic + Obj.repr (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.Typeclasses.fst" @@ -1154,9 +1155,9 @@ let (mk_class : | uu___10 -> - Obj.magic + Obj.repr (FStar_Tactics_Derived.fail - "mk_class: proj not Sg_Let?"))) + "mk_class: proj not Sg_Let?")))) uu___9))) (fun uu___9 -> @@ -1241,12 +1242,14 @@ let (mk_class : | [] -> Obj.magic + (Obj.repr (FStar_Tactics_Derived.fail - "mk_class: impossible, no binders") + "mk_class: impossible, no binders")) | b1::bs' -> Obj.magic + (Obj.repr (FStar_Tactics_Effect.tac_bind (Prims.mk_range "FStar.Tactics.Typeclasses.fst" @@ -1311,7 +1314,7 @@ let (mk_class : :: bs')) cod2)) uu___12))) - uu___11)))) + uu___11))))) uu___10))) uu___9))) (fun diff --git a/ocaml/fstar-lib/generated/Steel_Effect_Common.ml b/ocaml/fstar-lib/generated/Steel_Effect_Common.ml index afb14f087b1..cdd2331bbcf 100644 --- a/ocaml/fstar-lib/generated/Steel_Effect_Common.ml +++ b/ocaml/fstar-lib/generated/Steel_Effect_Common.ml @@ -804,26 +804,15 @@ let (name_appears_in : (Prims.of_int (730)) (Prims.of_int (13))) (Obj.magic (FStar_Tactics_Builtins.inspect t1)) (fun uu___1 -> - (fun uu___1 -> - match uu___1 with - | FStar_Reflection_Data.Tv_FVar fv -> - Obj.magic - (Obj.repr - (if - (FStar_Reflection_Builtins.inspect_fv fv) - = nm - then - Obj.repr - (FStar_Tactics_Effect.raise Appears) - else - Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___3 -> ())))) - | t2 -> - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___2 -> ())))) uu___1))) + match uu___1 with + | FStar_Reflection_Data.Tv_FVar fv -> + if (FStar_Reflection_Builtins.inspect_fv fv) = nm + then FStar_Tactics_Effect.raise Appears + else + FStar_Tactics_Effect.lift_div_tac + (fun uu___3 -> ()) + | t2 -> + FStar_Tactics_Effect.lift_div_tac (fun uu___2 -> ())))) (fun uu___ -> (fun ff -> Obj.magic @@ -858,12 +847,9 @@ let (name_appears_in : match uu___ with | Appears -> Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___1 -> true))) - | e -> - Obj.magic - (Obj.repr (FStar_Tactics_Effect.raise e))) + (FStar_Tactics_Effect.lift_div_tac + (fun uu___1 -> true)) + | e -> Obj.magic (FStar_Tactics_Effect.raise e)) uu___))) uu___) let (term_appears_in : FStar_Reflection_Types.term -> @@ -1312,17 +1298,13 @@ let rec (try_candidates : (select t am) (select hd am))) (fun uu___3 -> - (fun uu___3 -> - if uu___3 - then - Obj.magic - (FStar_Tactics_Effect.raise - Success) - else - Obj.magic - (FStar_Tactics_Effect.raise - Failed)) - uu___3)) + if uu___3 + then + FStar_Tactics_Effect.raise + Success + else + FStar_Tactics_Effect.raise + Failed)) (fun uu___2 -> (fun uu___2 -> Obj.magic @@ -1384,9 +1366,8 @@ let rec (remove_from_list : (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (863)) (Prims.of_int (72)) (Prims.of_int (863)) (Prims.of_int (74))) - (Obj.magic - (FStar_Tactics_Derived.fail - "atom in remove_from_list not found: should not happen")) + (FStar_Tactics_Derived.fail + "atom in remove_from_list not found: should not happen") (fun uu___ -> FStar_Tactics_Effect.lift_div_tac (fun uu___1 -> [])))) @@ -1748,16 +1729,10 @@ let rec (try_unifying_remaining : (FStar_Tactics_Derived.unify u (select hd am))) (fun uu___1 -> - (fun uu___1 -> - if uu___1 - then - Obj.magic - (FStar_Tactics_Effect.raise - Success) - else - Obj.magic - (FStar_Tactics_Effect.raise - Failed)) uu___1)) + if uu___1 + then + FStar_Tactics_Effect.raise Success + else FStar_Tactics_Effect.raise Failed)) (fun uu___ -> match uu___ with | Success -> try_unifying_remaining tl u am @@ -1798,11 +1773,8 @@ let rec (try_unifying_remaining : "could not find candidate for scrutinee " uu___2)))) (fun uu___2 -> - (fun uu___2 -> - Obj.magic - (FStar_Tactics_Derived.fail - uu___2)) uu___2))))) uu___2 - uu___1 uu___ + FStar_Tactics_Derived.fail uu___2))))) + uu___2 uu___1 uu___ let (is_smt_binder : FStar_Reflection_Types.binder -> (Prims.bool, unit) FStar_Tactics_Effect.tac_repr) @@ -2085,13 +2057,11 @@ let rec (new_args_for_smt_attrs : with | FStar_Reflection_Data.C_Total ty2 -> - Obj.magic + FStar_Tactics_Effect.lift_div_tac ( - Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun + fun uu___1 -> - ty2))) + ty2) | FStar_Reflection_Data.C_Eff (uu___1, eff_name, @@ -2099,30 +2069,22 @@ let rec (new_args_for_smt_attrs : uu___2, uu___3) -> - Obj.magic - ( - Obj.repr - (if + if eff_name = ["Prims"; "Tot"] - then - Obj.repr - (FStar_Tactics_Effect.lift_div_tac + then + FStar_Tactics_Effect.lift_div_tac (fun uu___4 -> - ty2)) - else - Obj.repr - (FStar_Tactics_Derived.fail - "computation type not supported in definition of slprops"))) + ty2) + else + FStar_Tactics_Derived.fail + "computation type not supported in definition of slprops" | uu___1 -> - Obj.magic - ( - Obj.repr - (FStar_Tactics_Derived.fail - "computation type not supported in definition of slprops"))) + FStar_Tactics_Derived.fail + "computation type not supported in definition of slprops") (fun uu___1 -> (fun ty2 -> Obj.magic @@ -2360,9 +2322,7 @@ let fail_atoms : (fun uu___1 -> Prims.strcat "could not find a solution for unifying\n" uu___)))) - (fun uu___ -> - (fun uu___ -> Obj.magic (FStar_Tactics_Derived.fail uu___)) - uu___) + (fun uu___ -> FStar_Tactics_Derived.fail uu___) let rec (equivalent_lists_fallback : Prims.nat -> atom Prims.list -> @@ -2471,11 +2431,7 @@ let rec (equivalent_lists_fallback : "could not find candidates for " uu___2)))) (fun uu___2 -> - (fun uu___2 - -> - Obj.magic - (FStar_Tactics_Derived.fail - uu___2)) + FStar_Tactics_Derived.fail uu___2)))) uu___)) | uu___ -> @@ -2517,10 +2473,8 @@ let rec (equivalent_lists_fallback : "could not find candidates for " uu___1)))) (fun uu___1 -> - (fun uu___1 -> - Obj.magic - (FStar_Tactics_Derived.fail - uu___1)) uu___1)))) + FStar_Tactics_Derived.fail + uu___1)))) | uu___ -> Obj.magic (Obj.repr @@ -2767,11 +2721,7 @@ let rec (equivalent_lists' : uu___2)))) (fun uu___2 -> - (fun - uu___2 -> - Obj.magic - (FStar_Tactics_Derived.fail - uu___2)) + FStar_Tactics_Derived.fail uu___2)))) uu___)) | uu___ -> @@ -2813,10 +2763,8 @@ let rec (equivalent_lists' : "could not find candidates for " uu___1)))) (fun uu___1 -> - (fun uu___1 -> - Obj.magic - (FStar_Tactics_Derived.fail - uu___1)) uu___1)))) + FStar_Tactics_Derived.fail + uu___1)))) | uu___ -> Obj.magic (Obj.repr @@ -3171,18 +3119,13 @@ let rec (unifies_with_all_uvars : t hd_t)) (fun uu___3 -> - (fun - uu___3 -> if uu___3 then - Obj.magic - (FStar_Tactics_Effect.raise - Success) + FStar_Tactics_Effect.raise + Success else - Obj.magic - (FStar_Tactics_Effect.raise + FStar_Tactics_Effect.raise Failed)) - uu___3)) (fun uu___2 -> (fun @@ -4343,96 +4286,93 @@ let rec (unify_pr_with_true : | (hd, tl) -> if is_and hd then - (match tl with - | (pr_l, uu___1)::(pr_r, uu___2)::[] -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (1795)) (Prims.of_int (6)) - (Prims.of_int (1795)) (Prims.of_int (29))) - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (1796)) (Prims.of_int (6)) - (Prims.of_int (1796)) (Prims.of_int (29))) - (Obj.magic (unify_pr_with_true pr_l)) - (fun uu___3 -> - (fun uu___3 -> - Obj.magic (unify_pr_with_true pr_r)) uu___3)) - | uu___1 -> - Obj.magic - (FStar_Tactics_Derived.fail - "unify_pr_with_true: ill-formed /\\")) + Obj.magic + (Obj.repr + (match tl with + | (pr_l, uu___1)::(pr_r, uu___2)::[] -> + Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (1795)) (Prims.of_int (6)) + (Prims.of_int (1795)) (Prims.of_int (29))) + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (1796)) (Prims.of_int (6)) + (Prims.of_int (1796)) (Prims.of_int (29))) + (Obj.magic (unify_pr_with_true pr_l)) + (fun uu___3 -> + (fun uu___3 -> + Obj.magic (unify_pr_with_true pr_r)) + uu___3)) + | uu___1 -> + Obj.repr + (FStar_Tactics_Derived.fail + "unify_pr_with_true: ill-formed /\\"))) else Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (1799)) (Prims.of_int (10)) - (Prims.of_int (1799)) (Prims.of_int (30))) - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (1799)) (Prims.of_int (4)) - (Prims.of_int (1811)) (Prims.of_int (73))) - (Obj.magic - (FStar_Tactics_Derived.inspect_unascribe hd)) - (fun uu___2 -> + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (1799)) (Prims.of_int (10)) + (Prims.of_int (1799)) (Prims.of_int (30))) + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (1799)) (Prims.of_int (4)) + (Prims.of_int (1811)) (Prims.of_int (73))) + (Obj.magic + (FStar_Tactics_Derived.inspect_unascribe hd)) (fun uu___2 -> - match uu___2 with - | FStar_Reflection_Data.Tv_Uvar (uu___3, uu___4) - -> - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (1801)) - (Prims.of_int (9)) - (Prims.of_int (1801)) - (Prims.of_int (27))) - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (1801)) - (Prims.of_int (6)) - (Prims.of_int (1805)) - (Prims.of_int (9))) - (Obj.magic - (FStar_Tactics_Derived.unify pr - (FStar_Reflection_Builtins.pack_ln - (FStar_Reflection_Data.Tv_FVar - (FStar_Reflection_Builtins.pack_fv - ["Steel"; - "Effect"; - "Common"; - "true_p"]))))) - (fun uu___5 -> + (fun uu___2 -> + match uu___2 with + | FStar_Reflection_Data.Tv_Uvar + (uu___3, uu___4) -> + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (1801)) + (Prims.of_int (9)) + (Prims.of_int (1801)) + (Prims.of_int (27))) + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (1801)) + (Prims.of_int (6)) + (Prims.of_int (1805)) + (Prims.of_int (9))) + (Obj.magic + (FStar_Tactics_Derived.unify + pr + (FStar_Reflection_Builtins.pack_ln + (FStar_Reflection_Data.Tv_FVar + (FStar_Reflection_Builtins.pack_fv + ["Steel"; + "Effect"; + "Common"; + "true_p"]))))) (fun uu___5 -> if uu___5 then - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___6 -> ()))) + FStar_Tactics_Effect.lift_div_tac + (fun uu___6 -> ()) else - Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "unify_pr_with_true: could not unify SMT prop with True"))) - uu___5))) - | uu___3 -> - Obj.magic - (Obj.repr - (if - (FStar_List_Tot_Base.length - (FStar_Reflection_Builtins.free_uvars - pr)) - = Prims.int_zero - then - Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___4 -> ())) - else - Obj.repr - (FStar_Tactics_Effect.raise + FStar_Tactics_Derived.fail + "unify_pr_with_true: could not unify SMT prop with True"))) + | uu___3 -> + Obj.magic + (Obj.repr + (if + (FStar_List_Tot_Base.length + (FStar_Reflection_Builtins.free_uvars + pr)) + = Prims.int_zero + then + FStar_Tactics_Effect.lift_div_tac + (fun uu___4 -> ()) + else + FStar_Tactics_Effect.raise (Postpone - "unify_pr_with_true: some uvars are still there"))))) - uu___2))) uu___) + "unify_pr_with_true: some uvars are still there")))) + uu___2)))) uu___) let rec (set_abduction_variable_term : FStar_Reflection_Types.term -> (FStar_Reflection_Types.term, unit) FStar_Tactics_Effect.tac_repr) @@ -4451,27 +4391,112 @@ let rec (set_abduction_variable_term : | (hd, tl) -> if is_and hd then - (match tl with - | (pr_l, FStar_Reflection_Data.Q_Explicit)::(pr_r, - FStar_Reflection_Data.Q_Explicit)::[] - -> - if - (FStar_List_Tot_Base.length - (FStar_Reflection_Builtins.free_uvars pr_r)) - = Prims.int_zero - then - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (1838)) (Prims.of_int (18)) - (Prims.of_int (1838)) (Prims.of_int (50))) - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (1839)) (Prims.of_int (8)) - (Prims.of_int (1839)) (Prims.of_int (53))) - (Obj.magic (set_abduction_variable_term pr_l)) - (fun arg -> + Obj.magic + (Obj.repr + (match tl with + | (pr_l, FStar_Reflection_Data.Q_Explicit)::(pr_r, + FStar_Reflection_Data.Q_Explicit)::[] + -> + Obj.repr + (if + (FStar_List_Tot_Base.length + (FStar_Reflection_Builtins.free_uvars + pr_r)) + = Prims.int_zero + then + Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (1838)) + (Prims.of_int (18)) + (Prims.of_int (1838)) + (Prims.of_int (50))) + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (1839)) + (Prims.of_int (8)) + (Prims.of_int (1839)) + (Prims.of_int (53))) + (Obj.magic + (set_abduction_variable_term pr_l)) + (fun arg -> + FStar_Tactics_Effect.lift_div_tac + (fun uu___1 -> + FStar_Reflection_Derived.mk_app + (FStar_Reflection_Builtins.pack_ln + (FStar_Reflection_Data.Tv_FVar + (FStar_Reflection_Builtins.pack_fv + ["Steel"; + "Effect"; + "Common"; + "elim_and_l_squash"]))) + [(arg, + FStar_Reflection_Data.Q_Explicit)]))) + else + Obj.repr + (if + (FStar_List_Tot_Base.length + (FStar_Reflection_Builtins.free_uvars + pr_l)) + = Prims.int_zero + then + Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (1842)) + (Prims.of_int (18)) + (Prims.of_int (1842)) + (Prims.of_int (50))) + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (1843)) + (Prims.of_int (8)) + (Prims.of_int (1843)) + (Prims.of_int (53))) + (Obj.magic + (set_abduction_variable_term + pr_r)) + (fun arg -> + FStar_Tactics_Effect.lift_div_tac + (fun uu___2 -> + FStar_Reflection_Derived.mk_app + (FStar_Reflection_Builtins.pack_ln + (FStar_Reflection_Data.Tv_FVar + (FStar_Reflection_Builtins.pack_fv + ["Steel"; + "Effect"; + "Common"; + "elim_and_r_squash"]))) + [(arg, + FStar_Reflection_Data.Q_Explicit)]))) + else + Obj.repr + (FStar_Tactics_Effect.raise + (Postpone + "set_abduction_variable_term: there are still uvars on both sides of l_and")))) + | uu___1 -> + Obj.repr + (FStar_Tactics_Derived.fail + "set_abduction_variable: ill-formed /\\"))) + else + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (1831)) (Prims.of_int (6)) + (Prims.of_int (1831)) (Prims.of_int (8))) + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (1849)) (Prims.of_int (4)) + (Prims.of_int (1852)) (Prims.of_int (54))) + (Obj.magic (FStar_Tactics_Builtins.inspect hd)) + (fun uu___2 -> + match uu___2 with + | FStar_Reflection_Data.Tv_Uvar (uu___3, uu___4) + -> FStar_Tactics_Effect.lift_div_tac - (fun uu___1 -> + (fun uu___5 -> FStar_Reflection_Derived.mk_app (FStar_Reflection_Builtins.pack_ln (FStar_Reflection_Data.Tv_FVar @@ -4479,83 +4504,15 @@ let rec (set_abduction_variable_term : ["Steel"; "Effect"; "Common"; - "elim_and_l_squash"]))) - [(arg, - FStar_Reflection_Data.Q_Explicit)]))) - else - if - (FStar_List_Tot_Base.length - (FStar_Reflection_Builtins.free_uvars pr_l)) - = Prims.int_zero - then - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (1842)) (Prims.of_int (18)) - (Prims.of_int (1842)) (Prims.of_int (50))) - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (1843)) (Prims.of_int (8)) - (Prims.of_int (1843)) (Prims.of_int (53))) - (Obj.magic (set_abduction_variable_term pr_r)) - (fun arg -> - FStar_Tactics_Effect.lift_div_tac - (fun uu___2 -> - FStar_Reflection_Derived.mk_app - (FStar_Reflection_Builtins.pack_ln - (FStar_Reflection_Data.Tv_FVar - (FStar_Reflection_Builtins.pack_fv - ["Steel"; - "Effect"; - "Common"; - "elim_and_r_squash"]))) - [(arg, - FStar_Reflection_Data.Q_Explicit)]))) - else - Obj.magic - (FStar_Tactics_Effect.raise - (Postpone - "set_abduction_variable_term: there are still uvars on both sides of l_and")) - | uu___1 -> - Obj.magic - (FStar_Tactics_Derived.fail - "set_abduction_variable: ill-formed /\\")) - else - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (1831)) (Prims.of_int (6)) - (Prims.of_int (1831)) (Prims.of_int (8))) - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (1849)) (Prims.of_int (4)) - (Prims.of_int (1852)) (Prims.of_int (54))) - (Obj.magic (FStar_Tactics_Builtins.inspect hd)) - (fun uu___2 -> - (fun uu___2 -> - match uu___2 with - | FStar_Reflection_Data.Tv_Uvar (uu___3, uu___4) - -> - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___5 -> - FStar_Reflection_Derived.mk_app - (FStar_Reflection_Builtins.pack_ln - (FStar_Reflection_Data.Tv_FVar - (FStar_Reflection_Builtins.pack_fv - ["Steel"; - "Effect"; - "Common"; - "_return_squash"]))) - [((FStar_Reflection_Builtins.pack_ln - (FStar_Reflection_Data.Tv_Const - FStar_Reflection_Data.C_Unit)), - FStar_Reflection_Data.Q_Explicit)]))) + "_return_squash"]))) + [((FStar_Reflection_Builtins.pack_ln + (FStar_Reflection_Data.Tv_Const + FStar_Reflection_Data.C_Unit)), + FStar_Reflection_Data.Q_Explicit)]) | uu___3 -> - Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "set_abduction_variable: cannot unify"))) - uu___2))) uu___) + FStar_Tactics_Derived.fail + "set_abduction_variable: cannot unify")))) + uu___) let (set_abduction_variable : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = fun uu___ -> @@ -4581,62 +4538,68 @@ let (set_abduction_variable : match uu___1 with | FStar_Reflection_Data.Tv_Arrow (b, uu___2) -> Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (1858)) (Prims.of_int (18)) - (Prims.of_int (1858)) (Prims.of_int (34))) - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (1858)) (Prims.of_int (4)) - (Prims.of_int (1861)) (Prims.of_int (42))) - (FStar_Tactics_Effect.lift_div_tac - (fun uu___3 -> - FStar_Reflection_Builtins.inspect_binder - b)) - (fun uu___3 -> + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (1858)) + (Prims.of_int (18)) + (Prims.of_int (1858)) + (Prims.of_int (34))) + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (1858)) + (Prims.of_int (4)) + (Prims.of_int (1861)) + (Prims.of_int (42))) + (FStar_Tactics_Effect.lift_div_tac + (fun uu___3 -> + FStar_Reflection_Builtins.inspect_binder + b)) (fun uu___3 -> - match uu___3 with - | (bv, uu___4) -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (1859)) - (Prims.of_int (13)) - (Prims.of_int (1859)) - (Prims.of_int (26))) - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (1860)) - (Prims.of_int (4)) - (Prims.of_int (1861)) - (Prims.of_int (42))) - (FStar_Tactics_Effect.lift_div_tac + (fun uu___3 -> + match uu___3 with + | (bv, uu___4) -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (1859)) + (Prims.of_int (13)) + (Prims.of_int (1859)) + (Prims.of_int (26))) + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (1860)) + (Prims.of_int (4)) + (Prims.of_int (1861)) + (Prims.of_int (42))) + (FStar_Tactics_Effect.lift_div_tac + (fun uu___5 -> + FStar_Reflection_Builtins.inspect_bv + bv)) (fun uu___5 -> - FStar_Reflection_Builtins.inspect_bv - bv)) - (fun uu___5 -> - (fun bv1 -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (1860)) - (Prims.of_int (13)) - (Prims.of_int (1860)) - (Prims.of_int (23))) - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (1861)) - (Prims.of_int (4)) - (Prims.of_int (1861)) - (Prims.of_int (42))) - (FStar_Tactics_Effect.lift_div_tac + (fun bv1 -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (1860)) + (Prims.of_int (13)) + (Prims.of_int (1860)) + (Prims.of_int (23))) + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (1861)) + (Prims.of_int (4)) + (Prims.of_int (1861)) + (Prims.of_int (42))) + (FStar_Tactics_Effect.lift_div_tac + (fun uu___5 + -> + bv1.FStar_Reflection_Data.bv_sort)) (fun uu___5 -> - bv1.FStar_Reflection_Data.bv_sort)) - (fun uu___5 -> - (fun pr -> - Obj.magic - (FStar_Tactics_Effect.tac_bind + (fun pr -> + Obj.magic + (FStar_Tactics_Effect.tac_bind (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (1861)) @@ -4660,12 +4623,13 @@ let (set_abduction_variable : (FStar_Tactics_Derived.exact uu___5)) uu___5))) - uu___5))) - uu___5))) uu___3)) + uu___5))) + uu___5))) uu___3))) | uu___2 -> Obj.magic - (FStar_Tactics_Derived.fail "Not an arrow goal")) - uu___1))) uu___1) + (Obj.repr + (FStar_Tactics_Derived.fail + "Not an arrow goal"))) uu___1))) uu___1) let (canon_l_r : Prims.bool -> FStar_Reflection_Types.term -> @@ -4850,15 +4814,11 @@ let (canon_l_r : (flatten r2_raw) am2)) - (fun - uu___4 -> (fun res -> - Obj.magic - (FStar_Tactics_Effect.raise + FStar_Tactics_Effect.raise (Result res))) - uu___4)) (fun uu___3 -> (fun @@ -4869,24 +4829,21 @@ let (canon_l_r : FStar_Tactics_Common.TacticFailure m1 -> Obj.magic - (Obj.repr (FStar_Tactics_Derived.fail - m1)) + m1) | Result res -> Obj.magic - (Obj.repr (FStar_Tactics_Effect.lift_div_tac (fun uu___4 -> - res))) + res)) | uu___4 -> Obj.magic - (Obj.repr (FStar_Tactics_Derived.fail - "uncaught exception in equivalent_lists"))) + "uncaught exception in equivalent_lists")) uu___3))) (fun uu___3 -> @@ -6577,19 +6534,25 @@ let (canon_l_r : (fun uu___11 -> + (fun + uu___11 + -> match uu___11 with | FStar_Tactics_Common.TacticFailure msg -> - FStar_Tactics_Derived.fail + Obj.magic + (FStar_Tactics_Derived.fail (Prims.strcat "Cannot unify pr with true: " - msg) + msg)) | e -> - FStar_Tactics_Effect.raise - e))) + Obj.magic + (FStar_Tactics_Effect.raise + e)) + uu___11))) | l -> Obj.magic @@ -6832,34 +6795,40 @@ let (canon_monoid : (match rel_xy with | (rel_xy1, uu___2)::[] -> Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2090)) - (Prims.of_int (21)) - (Prims.of_int (2090)) - (Prims.of_int (43))) - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2090)) - (Prims.of_int (7)) - (Prims.of_int (2100)) - (Prims.of_int (8))) - (FStar_Tactics_Effect.lift_div_tac - (fun uu___3 -> - FStar_Reflection_Derived_Lemmas.collect_app_ref - rel_xy1)) - (fun uu___3 -> + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2090)) + (Prims.of_int (21)) + (Prims.of_int (2090)) + (Prims.of_int (43))) + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2090)) + (Prims.of_int (7)) + (Prims.of_int (2100)) + (Prims.of_int (8))) + (FStar_Tactics_Effect.lift_div_tac + (fun uu___3 + -> + FStar_Reflection_Derived_Lemmas.collect_app_ref + rel_xy1)) (fun uu___3 -> - match uu___3 - with - | (rel, xy) + (fun uu___3 + -> + match uu___3 + with + | + (rel, xy) -> if (FStar_List_Tot_Base.length xy) >= (Prims.of_int (2)) then + Obj.magic + (Obj.repr (match ((FStar_List_Tot_Base.index xy @@ -6878,7 +6847,7 @@ let (canon_monoid : (rhs, FStar_Reflection_Data.Q_Explicit)) -> - Obj.magic + Obj.repr (canon_l_r use_smt carrier_t @@ -6888,18 +6857,20 @@ let (canon_monoid : rhs) | uu___4 -> - Obj.magic + Obj.repr (FStar_Tactics_Derived.fail - "Goal should have been an application of a binary relation to 2 explicit arguments")) + "Goal should have been an application of a binary relation to 2 explicit arguments"))) else Obj.magic + (Obj.repr (FStar_Tactics_Derived.fail - "Goal should have been an application of a binary relation to n implicit and 2 explicit arguments")) - uu___3)) + "Goal should have been an application of a binary relation to n implicit and 2 explicit arguments"))) + uu___3))) | uu___2 -> Obj.magic - (FStar_Tactics_Derived.fail - "Goal should be squash applied to a binary relation"))) + (Obj.repr + (FStar_Tactics_Derived.fail + "Goal should be squash applied to a binary relation")))) uu___1))) uu___1))) uu___) let (canon' : Prims.bool -> @@ -7765,13 +7736,15 @@ let rec (extract_contexts : FStar_Pervasives_Native.None -> Obj.magic + (Obj.repr (FStar_Tactics_Derived.fail - "no context on the right either") + "no context on the right either")) | FStar_Pervasives_Native.Some g -> Obj.magic - (g ())) + (Obj.repr + (g ()))) uu___5))))))) uu___1))) uu___1)) | uu___1 -> @@ -8018,15 +7991,9 @@ let (open_existentials : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) "Common"; "solve_can_be_split_lookup"]))) e) then - Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "Tactic disabled: no available lemmas in context")) - else - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___2 -> ())))) + FStar_Tactics_Derived.fail + "Tactic disabled: no available lemmas in context" + else FStar_Tactics_Effect.lift_div_tac (fun uu___2 -> ())) (fun uu___1 -> (fun uu___1 -> Obj.magic @@ -8069,40 +8036,44 @@ let (open_existentials : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) FStar_Reflection_Data.Q_Explicit)::[]) -> Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2309)) - (Prims.of_int (22)) - (Prims.of_int (2309)) - (Prims.of_int (36))) - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2309)) - (Prims.of_int (9)) - (Prims.of_int (2325)) - (Prims.of_int (60))) - (FStar_Tactics_Effect.lift_div_tac - (fun uu___4 -> - FStar_Reflection_Derived.collect_app - t1)) - (fun uu___4 -> + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2309)) + (Prims.of_int (22)) + (Prims.of_int (2309)) + (Prims.of_int (36))) + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2309)) + (Prims.of_int (9)) + (Prims.of_int (2325)) + (Prims.of_int (60))) + (FStar_Tactics_Effect.lift_div_tac + (fun uu___4 -> + FStar_Reflection_Derived.collect_app + t1)) (fun uu___4 -> - match uu___4 with - | (hd, tl) -> - if - FStar_Reflection_Derived.is_fvar + (fun uu___4 -> + match uu___4 + with + | (hd, tl) -> + if + FStar_Reflection_Derived.is_fvar hd "Steel.Effect.Common.can_be_split" - then - (match tl - with - | - uu___5:: + then + Obj.magic + (Obj.repr + (match tl + with + | + uu___5:: (rhs, FStar_Reflection_Data.Q_Explicit)::[] -> - Obj.magic + Obj.repr (FStar_Tactics_Effect.tac_bind (Prims.mk_range "Steel.Effect.Common.fsti" @@ -8129,12 +8100,14 @@ let (open_existentials : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) FStar_Pervasives_Native.None -> Obj.magic + (Obj.repr (FStar_Tactics_Derived.fail - "open_existentials: no context found") + "open_existentials: no context found")) | FStar_Pervasives_Native.Some f -> Obj.magic + (Obj.repr (FStar_Tactics_Effect.tac_bind (Prims.mk_range "Steel.Effect.Common.fsti" @@ -8232,23 +8205,24 @@ let (open_existentials : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) uu___10))) uu___9))) uu___8))) - uu___7))) + uu___7)))) uu___6)) - | - uu___5 -> + | + uu___5 -> + Obj.repr + (FStar_Tactics_Derived.fail + "open_existentials: ill-formed can_be_split"))) + else Obj.magic + (Obj.repr (FStar_Tactics_Derived.fail - "open_existentials: ill-formed can_be_split")) - else - Obj.magic - ( - FStar_Tactics_Derived.fail - "open_existentials: not a can_be_split goal")) - uu___4)) + "open_existentials: not a can_be_split goal"))) + uu___4))) | uu___3 -> Obj.magic - (FStar_Tactics_Derived.fail - "open_existentials: not a squash goal")) + (Obj.repr + (FStar_Tactics_Derived.fail + "open_existentials: not a squash goal"))) uu___3))) uu___2))) uu___1))) uu___1) let (try_open_existentials : @@ -8567,85 +8541,104 @@ let (solve_can_be_split_dep : FStar_Reflection_Data.argv Prims.list -> (Prims.bool, unit) FStar_Tactics_Effect.tac_repr) = - fun args -> - match args with - | (p, uu___)::(t1, uu___1)::(t2, uu___2)::[] -> - FStar_Tactics_Effect.tac_bind - (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2388)) - (Prims.of_int (17)) (Prims.of_int (2388)) (Prims.of_int (36))) - (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2389)) - (Prims.of_int (6)) (Prims.of_int (2420)) (Prims.of_int (18))) - (Obj.magic (slterm_nbr_uvars t1)) - (fun uu___3 -> - (fun lnbr -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2389)) (Prims.of_int (17)) - (Prims.of_int (2389)) (Prims.of_int (36))) - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2390)) (Prims.of_int (6)) - (Prims.of_int (2420)) (Prims.of_int (18))) - (Obj.magic (slterm_nbr_uvars t2)) - (fun uu___3 -> - (fun rnbr -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2391)) (Prims.of_int (8)) - (Prims.of_int (2393)) (Prims.of_int (18))) - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2390)) (Prims.of_int (6)) - (Prims.of_int (2420)) (Prims.of_int (18))) - (if (lnbr + rnbr) <= Prims.int_one - then Obj.magic (Obj.repr (unfold_guard ())) - else - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___4 -> false)))) - (fun uu___3 -> - (fun uu___3 -> - if uu___3 - then - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2396)) - (Prims.of_int (8)) - (Prims.of_int (2416)) - (Prims.of_int (36))) - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2418)) - (Prims.of_int (8)) - (Prims.of_int (2418)) - (Prims.of_int (12))) - (Obj.magic - (FStar_Tactics_Derived.focus - (fun uu___4 -> - FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2397)) - (Prims.of_int (23)) - (Prims.of_int (2397)) - (Prims.of_int (39))) - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2398)) - (Prims.of_int (10)) - (Prims.of_int (2416)) - (Prims.of_int (35))) - (Obj.magic - (FStar_Tactics_Logic.implies_intro - ())) - (fun uu___5 -> - (fun p_bind -> - Obj.magic - (FStar_Tactics_Effect.tac_bind + fun uu___ -> + (fun args -> + match args with + | (p, uu___)::(t1, uu___1)::(t2, uu___2)::[] -> + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2388)) (Prims.of_int (17)) + (Prims.of_int (2388)) (Prims.of_int (36))) + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2389)) (Prims.of_int (6)) + (Prims.of_int (2420)) (Prims.of_int (18))) + (Obj.magic (slterm_nbr_uvars t1)) + (fun uu___3 -> + (fun lnbr -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2389)) (Prims.of_int (17)) + (Prims.of_int (2389)) (Prims.of_int (36))) + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2390)) (Prims.of_int (6)) + (Prims.of_int (2420)) (Prims.of_int (18))) + (Obj.magic (slterm_nbr_uvars t2)) + (fun uu___3 -> + (fun rnbr -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2391)) + (Prims.of_int (8)) + (Prims.of_int (2393)) + (Prims.of_int (18))) + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2390)) + (Prims.of_int (6)) + (Prims.of_int (2420)) + (Prims.of_int (18))) + (if (lnbr + rnbr) <= Prims.int_one + then + Obj.magic + (Obj.repr (unfold_guard ())) + else + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___4 -> false)))) + (fun uu___3 -> + (fun uu___3 -> + if uu___3 + then + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2396)) + (Prims.of_int (8)) + (Prims.of_int (2416)) + (Prims.of_int (36))) + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2418)) + (Prims.of_int (8)) + (Prims.of_int (2418)) + (Prims.of_int (12))) + (Obj.magic + (FStar_Tactics_Derived.focus + (fun uu___4 -> + FStar_Tactics_Effect.tac_bind + ( + Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2397)) + (Prims.of_int (23)) + (Prims.of_int (2397)) + (Prims.of_int (39))) + ( + Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2398)) + (Prims.of_int (10)) + (Prims.of_int (2416)) + (Prims.of_int (35))) + ( + Obj.magic + (FStar_Tactics_Logic.implies_intro + ())) + ( + fun + uu___5 -> + (fun + p_bind -> + Obj.magic + (FStar_Tactics_Effect.tac_bind (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2398)) @@ -8740,17 +8733,13 @@ let (solve_can_be_split_dep : Prims.op_Negation b then - Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "could not unify SMT prop with True")) + FStar_Tactics_Derived.fail + "could not unify SMT prop with True" else - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac + FStar_Tactics_Effect.lift_div_tac (fun uu___9 -> - ())))) + ())) (fun uu___8 -> (fun @@ -8912,121 +8901,144 @@ let (solve_can_be_split_dep : uu___8)))) uu___6))) uu___5))) - uu___5)))) - (fun uu___4 -> - FStar_Tactics_Effect.lift_div_tac - (fun uu___5 -> true)))) - else - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___5 -> false)))) - uu___3))) uu___3))) uu___3) - | uu___ -> FStar_Tactics_Derived.fail "ill-formed can_be_split_dep" + uu___5)))) + (fun uu___4 -> + FStar_Tactics_Effect.lift_div_tac + (fun uu___5 -> + true)))) + else + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___5 -> false)))) + uu___3))) uu___3))) uu___3))) + | uu___ -> + Obj.magic + (Obj.repr + (FStar_Tactics_Derived.fail "ill-formed can_be_split_dep"))) + uu___ let (solve_can_be_split_forall : FStar_Reflection_Data.argv Prims.list -> (Prims.bool, unit) FStar_Tactics_Effect.tac_repr) = - fun args -> - match args with - | uu___::(t1, uu___1)::(t2, uu___2)::[] -> - FStar_Tactics_Effect.tac_bind - (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2432)) - (Prims.of_int (17)) (Prims.of_int (2432)) (Prims.of_int (36))) - (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2433)) - (Prims.of_int (6)) (Prims.of_int (2458)) (Prims.of_int (18))) - (Obj.magic (slterm_nbr_uvars t1)) - (fun uu___3 -> - (fun lnbr -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2433)) (Prims.of_int (17)) - (Prims.of_int (2433)) (Prims.of_int (36))) - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2434)) (Prims.of_int (6)) - (Prims.of_int (2458)) (Prims.of_int (18))) - (Obj.magic (slterm_nbr_uvars t2)) - (fun uu___3 -> - (fun rnbr -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2435)) (Prims.of_int (8)) - (Prims.of_int (2437)) (Prims.of_int (18))) - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2434)) (Prims.of_int (6)) - (Prims.of_int (2458)) (Prims.of_int (18))) - (if (lnbr + rnbr) <= Prims.int_one - then Obj.magic (Obj.repr (unfold_guard ())) - else - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___4 -> false)))) - (fun uu___3 -> - (fun uu___3 -> - if uu___3 - then - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2440)) - (Prims.of_int (8)) - (Prims.of_int (2456)) - (Prims.of_int (46))) - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2457)) - (Prims.of_int (8)) - (Prims.of_int (2457)) - (Prims.of_int (12))) - (Obj.magic - (FStar_Tactics_Derived.focus - (fun uu___4 -> - FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2441)) - (Prims.of_int (10)) - (Prims.of_int (2441)) - (Prims.of_int (33))) - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2442)) - (Prims.of_int (10)) - (Prims.of_int (2456)) - (Prims.of_int (45))) - (Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range + fun uu___ -> + (fun args -> + match args with + | uu___::(t1, uu___1)::(t2, uu___2)::[] -> + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2432)) (Prims.of_int (17)) + (Prims.of_int (2432)) (Prims.of_int (36))) + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2433)) (Prims.of_int (6)) + (Prims.of_int (2458)) (Prims.of_int (18))) + (Obj.magic (slterm_nbr_uvars t1)) + (fun uu___3 -> + (fun lnbr -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2433)) (Prims.of_int (17)) + (Prims.of_int (2433)) (Prims.of_int (36))) + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2434)) (Prims.of_int (6)) + (Prims.of_int (2458)) (Prims.of_int (18))) + (Obj.magic (slterm_nbr_uvars t2)) + (fun uu___3 -> + (fun rnbr -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2435)) + (Prims.of_int (8)) + (Prims.of_int (2437)) + (Prims.of_int (18))) + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2434)) + (Prims.of_int (6)) + (Prims.of_int (2458)) + (Prims.of_int (18))) + (if (lnbr + rnbr) <= Prims.int_one + then + Obj.magic + (Obj.repr (unfold_guard ())) + else + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___4 -> false)))) + (fun uu___3 -> + (fun uu___3 -> + if uu___3 + then + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2440)) + (Prims.of_int (8)) + (Prims.of_int (2456)) + (Prims.of_int (46))) + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2457)) + (Prims.of_int (8)) + (Prims.of_int (2457)) + (Prims.of_int (12))) + (Obj.magic + (FStar_Tactics_Derived.focus + (fun uu___4 -> + FStar_Tactics_Effect.tac_bind + ( + Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2441)) + (Prims.of_int (10)) + (Prims.of_int (2441)) + (Prims.of_int (33))) + ( + Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2442)) + (Prims.of_int (10)) + (Prims.of_int (2456)) + (Prims.of_int (45))) + ( + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2441)) (Prims.of_int (17)) (Prims.of_int (2441)) (Prims.of_int (33))) - (Prims.mk_range + (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2441)) (Prims.of_int (10)) (Prims.of_int (2441)) (Prims.of_int (33))) - (Obj.magic - ( - FStar_Tactics_Logic.forall_intro + (Obj.magic + (FStar_Tactics_Logic.forall_intro ())) - (fun uu___5 - -> + (fun + uu___5 -> FStar_Tactics_Effect.lift_div_tac (fun uu___6 -> ())))) - (fun uu___5 -> - (fun uu___5 -> - Obj.magic - (FStar_Tactics_Effect.tac_bind + ( + fun + uu___5 -> + (fun + uu___5 -> + Obj.magic + (FStar_Tactics_Effect.tac_bind (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2442)) @@ -9225,19 +9237,23 @@ let (solve_can_be_split_forall : uu___9)))) uu___7))) uu___6))) - uu___5)))) - (fun uu___4 -> - FStar_Tactics_Effect.lift_div_tac - (fun uu___5 -> true)))) - else - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___5 -> false)))) - uu___3))) uu___3))) uu___3) - | uu___ -> - FStar_Tactics_Derived.fail - "Ill-formed can_be_split_forall, should not happen" + uu___5)))) + (fun uu___4 -> + FStar_Tactics_Effect.lift_div_tac + (fun uu___5 -> + true)))) + else + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___5 -> false)))) + uu___3))) uu___3))) uu___3))) + | uu___ -> + Obj.magic + (Obj.repr + (FStar_Tactics_Derived.fail + "Ill-formed can_be_split_forall, should not happen"))) + uu___ let (extract_cbs_forall_dep_contexts : FStar_Reflection_Types.term -> ((unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) @@ -9298,15 +9314,9 @@ let (open_existentials_forall_dep : "solve_can_be_split_forall_dep_lookup"]))) e) then - Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "Tactic disabled: no available lemmas in context")) - else - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___2 -> ())))) + FStar_Tactics_Derived.fail + "Tactic disabled: no available lemmas in context" + else FStar_Tactics_Effect.lift_div_tac (fun uu___2 -> ())) (fun uu___1 -> (fun uu___1 -> Obj.magic @@ -9354,40 +9364,44 @@ let (open_existentials_forall_dep : FStar_Reflection_Data.Q_Explicit)::[]) -> Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2491)) - (Prims.of_int (17)) - (Prims.of_int (2491)) - (Prims.of_int (31))) - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2491)) - (Prims.of_int (4)) - (Prims.of_int (2513)) - (Prims.of_int (78))) - (FStar_Tactics_Effect.lift_div_tac - (fun uu___4 -> - FStar_Reflection_Derived.collect_app - t1)) - (fun uu___4 -> + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2491)) + (Prims.of_int (17)) + (Prims.of_int (2491)) + (Prims.of_int (31))) + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2491)) + (Prims.of_int (4)) + (Prims.of_int (2513)) + (Prims.of_int (78))) + (FStar_Tactics_Effect.lift_div_tac + (fun uu___4 -> + FStar_Reflection_Derived.collect_app + t1)) (fun uu___4 -> - match uu___4 with - | (hd, tl) -> - if - FStar_Reflection_Derived.is_fvar + (fun uu___4 -> + match uu___4 + with + | (hd, tl) -> + if + FStar_Reflection_Derived.is_fvar hd "Steel.Effect.Common.can_be_split_forall_dep" - then - (match tl - with - | - uu___5::uu___6:: + then + Obj.magic + (Obj.repr + (match tl + with + | + uu___5::uu___6:: (rhs, FStar_Reflection_Data.Q_Explicit)::[] -> - Obj.magic + Obj.repr (FStar_Tactics_Effect.tac_bind (Prims.mk_range "Steel.Effect.Common.fsti" @@ -9415,6 +9429,7 @@ let (open_existentials_forall_dep : (uu___8, body) -> Obj.magic + (Obj.repr (FStar_Tactics_Effect.tac_bind (Prims.mk_range "Steel.Effect.Common.fsti" @@ -9441,12 +9456,14 @@ let (open_existentials_forall_dep : FStar_Pervasives_Native.None -> Obj.magic + (Obj.repr (FStar_Tactics_Derived.fail - "open_existentials_forall_dep: no candidate") + "open_existentials_forall_dep: no candidate")) | FStar_Pervasives_Native.Some f -> Obj.magic + (Obj.repr (FStar_Tactics_Effect.tac_bind (Prims.mk_range "Steel.Effect.Common.fsti" @@ -9632,21 +9649,22 @@ let (open_existentials_forall_dep : uu___13))) uu___12))) uu___11))) - uu___10))) - uu___9)) + uu___10)))) + uu___9))) | uu___8 -> Obj.magic + (Obj.repr (FStar_Tactics_Derived.fail - "open_existentials_forall_dep : not an abstraction")) + "open_existentials_forall_dep : not an abstraction"))) uu___7)) - | - (uu___5, + | + (uu___5, FStar_Reflection_Data.Q_Implicit)::uu___6::uu___7:: (rhs, FStar_Reflection_Data.Q_Explicit)::[] -> - Obj.magic + Obj.repr (FStar_Tactics_Effect.tac_bind (Prims.mk_range "Steel.Effect.Common.fsti" @@ -9674,6 +9692,7 @@ let (open_existentials_forall_dep : (uu___9, body) -> Obj.magic + (Obj.repr (FStar_Tactics_Effect.tac_bind (Prims.mk_range "Steel.Effect.Common.fsti" @@ -9702,12 +9721,14 @@ let (open_existentials_forall_dep : FStar_Pervasives_Native.None -> Obj.magic + (Obj.repr (FStar_Tactics_Derived.fail - "open_existentials_forall_dep: no candidate") + "open_existentials_forall_dep: no candidate")) | FStar_Pervasives_Native.Some f -> Obj.magic + (Obj.repr (FStar_Tactics_Effect.tac_bind (Prims.mk_range "Steel.Effect.Common.fsti" @@ -9893,29 +9914,31 @@ let (open_existentials_forall_dep : uu___14))) uu___13))) uu___12))) - uu___11))) - uu___10)) + uu___11)))) + uu___10))) | uu___9 -> Obj.magic + (Obj.repr (FStar_Tactics_Derived.fail - "open_existentials_forall_dep : not an abstraction")) + "open_existentials_forall_dep : not an abstraction"))) uu___8)) - | - uu___5 -> + | + uu___5 -> + Obj.repr + (FStar_Tactics_Derived.fail + "open_existentials_forall_dep : wrong number of arguments to can_be_split_forall_dep"))) + else Obj.magic + (Obj.repr (FStar_Tactics_Derived.fail - "open_existentials_forall_dep : wrong number of arguments to can_be_split_forall_dep")) - else - Obj.magic - ( - FStar_Tactics_Derived.fail - "open_existentials_forall_dep : not a can_be_split_forall_dep goal")) - uu___4)) + "open_existentials_forall_dep : not a can_be_split_forall_dep goal"))) + uu___4))) | uu___3 -> Obj.magic - (FStar_Tactics_Derived.fail - "open_existentials_forall_dep : not a squash/auto_squash goal")) + (Obj.repr + (FStar_Tactics_Derived.fail + "open_existentials_forall_dep : not a squash/auto_squash goal"))) uu___3))) uu___2))) uu___1))) uu___1) let (try_open_existentials_forall_dep : @@ -9946,90 +9969,100 @@ let rec (solve_can_be_split_forall_dep : FStar_Reflection_Data.argv Prims.list -> (Prims.bool, unit) FStar_Tactics_Effect.tac_repr) = - fun args -> - match args with - | uu___::(pr, uu___1)::(t1, uu___2)::(t2, uu___3)::[] -> - FStar_Tactics_Effect.tac_bind - (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2530)) - (Prims.of_int (17)) (Prims.of_int (2530)) (Prims.of_int (36))) - (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2531)) - (Prims.of_int (6)) (Prims.of_int (2575)) (Prims.of_int (18))) - (Obj.magic (slterm_nbr_uvars t1)) - (fun uu___4 -> - (fun lnbr -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2531)) (Prims.of_int (17)) - (Prims.of_int (2531)) (Prims.of_int (36))) - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2532)) (Prims.of_int (6)) - (Prims.of_int (2575)) (Prims.of_int (18))) - (Obj.magic (slterm_nbr_uvars t2)) - (fun uu___4 -> - (fun rnbr -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2533)) (Prims.of_int (8)) - (Prims.of_int (2535)) (Prims.of_int (18))) - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2532)) (Prims.of_int (6)) - (Prims.of_int (2575)) (Prims.of_int (18))) - (if (lnbr + rnbr) <= Prims.int_one - then Obj.magic (Obj.repr (unfold_guard ())) - else - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___5 -> false)))) - (fun uu___4 -> - (fun uu___4 -> - if uu___4 - then - Obj.magic - (Obj.repr - (FStar_Tactics_Derived.try_with - (fun uu___5 -> - match () with - | () -> - FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2539)) - (Prims.of_int (9)) - (Prims.of_int (2563)) - (Prims.of_int (37))) - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2565)) - (Prims.of_int (9)) - (Prims.of_int (2565)) - (Prims.of_int (13))) - (Obj.magic - (FStar_Tactics_Derived.focus - (fun uu___6 -> - FStar_Tactics_Effect.tac_bind - ( - Prims.mk_range + fun uu___ -> + (fun args -> + match args with + | uu___::(pr, uu___1)::(t1, uu___2)::(t2, uu___3)::[] -> + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2530)) (Prims.of_int (17)) + (Prims.of_int (2530)) (Prims.of_int (36))) + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2531)) (Prims.of_int (6)) + (Prims.of_int (2575)) (Prims.of_int (18))) + (Obj.magic (slterm_nbr_uvars t1)) + (fun uu___4 -> + (fun lnbr -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2531)) (Prims.of_int (17)) + (Prims.of_int (2531)) (Prims.of_int (36))) + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2532)) (Prims.of_int (6)) + (Prims.of_int (2575)) (Prims.of_int (18))) + (Obj.magic (slterm_nbr_uvars t2)) + (fun uu___4 -> + (fun rnbr -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2533)) + (Prims.of_int (8)) + (Prims.of_int (2535)) + (Prims.of_int (18))) + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2532)) + (Prims.of_int (6)) + (Prims.of_int (2575)) + (Prims.of_int (18))) + (if (lnbr + rnbr) <= Prims.int_one + then + Obj.magic + (Obj.repr (unfold_guard ())) + else + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___5 -> false)))) + (fun uu___4 -> + (fun uu___4 -> + if uu___4 + then + Obj.magic + (Obj.repr + (FStar_Tactics_Derived.try_with + (fun uu___5 -> + match () with + | () -> + FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2540)) - (Prims.of_int (10)) - (Prims.of_int (2540)) - (Prims.of_int (17))) - ( - Prims.mk_range + (Prims.of_int (2539)) + (Prims.of_int (9)) + (Prims.of_int (2563)) + (Prims.of_int (37))) + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2565)) + (Prims.of_int (9)) + (Prims.of_int (2565)) + (Prims.of_int (13))) + (Obj.magic + (FStar_Tactics_Derived.focus + (fun + uu___6 -> + FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2540)) + (Prims.of_int (10)) + (Prims.of_int (2540)) + (Prims.of_int (17))) + (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2541)) (Prims.of_int (10)) (Prims.of_int (2563)) (Prims.of_int (36))) - ( - Obj.magic + (Obj.magic (FStar_Tactics_Builtins.norm [])) - ( - fun + (fun uu___7 -> (fun uu___7 -> @@ -10290,17 +10323,13 @@ let rec (solve_can_be_split_forall_dep : Prims.op_Negation b then - Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "could not unify SMT prop with True")) + FStar_Tactics_Derived.fail + "could not unify SMT prop with True" else - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac + FStar_Tactics_Effect.lift_div_tac (fun uu___12 - -> ())))) + -> ())) (fun uu___11 -> @@ -10477,158 +10506,186 @@ let rec (solve_can_be_split_forall_dep : uu___8))) uu___8))) uu___7)))) - (fun uu___6 -> - FStar_Tactics_Effect.lift_div_tac - (fun uu___7 -> - true))) - (fun uu___5 -> - (fun uu___5 -> - match uu___5 with - | Postpone msg -> - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac (fun uu___6 - -> false))) - | FStar_Tactics_Common.TacticFailure - msg -> - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range + -> + FStar_Tactics_Effect.lift_div_tac + (fun + uu___7 -> + true))) + (fun uu___5 -> + (fun uu___5 -> + match uu___5 + with + | Postpone msg + -> + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun + uu___6 -> + false))) + | FStar_Tactics_Common.TacticFailure + msg -> + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2570)) (Prims.of_int (22)) (Prims.of_int (2570)) (Prims.of_int (57))) - (Prims.mk_range + (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2571)) (Prims.of_int (9)) (Prims.of_int (2573)) (Prims.of_int (22))) - (Obj.magic + (Obj.magic (try_open_existentials_forall_dep ())) - (fun uu___6 - -> + (fun + uu___6 -> (fun opened -> if opened then Obj.magic + (Obj.repr (solve_can_be_split_forall_dep - args) + args)) else Obj.magic + (Obj.repr (FStar_Tactics_Derived.fail - msg)) + msg))) uu___6))) - | uu___6 -> - Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "Unexpected exception in framing tactic"))) - uu___5))) - else - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___6 -> false)))) - uu___4))) uu___4))) uu___4) - | uu___ -> - FStar_Tactics_Derived.fail - "Ill-formed can_be_split_forall_dep, should not happen" + | uu___6 -> + Obj.magic + (Obj.repr + (FStar_Tactics_Derived.fail + "Unexpected exception in framing tactic"))) + uu___5))) + else + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___6 -> false)))) + uu___4))) uu___4))) uu___4))) + | uu___ -> + Obj.magic + (Obj.repr + (FStar_Tactics_Derived.fail + "Ill-formed can_be_split_forall_dep, should not happen"))) + uu___ let (solve_equiv_forall : FStar_Reflection_Data.argv Prims.list -> (Prims.bool, unit) FStar_Tactics_Effect.tac_repr) = - fun args -> - match args with - | uu___::(t1, uu___1)::(t2, uu___2)::[] -> - FStar_Tactics_Effect.tac_bind - (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2583)) - (Prims.of_int (17)) (Prims.of_int (2583)) (Prims.of_int (36))) - (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2584)) - (Prims.of_int (6)) (Prims.of_int (2612)) (Prims.of_int (18))) - (Obj.magic (slterm_nbr_uvars t1)) - (fun uu___3 -> - (fun lnbr -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2584)) (Prims.of_int (17)) - (Prims.of_int (2584)) (Prims.of_int (36))) - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2585)) (Prims.of_int (6)) - (Prims.of_int (2612)) (Prims.of_int (18))) - (Obj.magic (slterm_nbr_uvars t2)) - (fun uu___3 -> - (fun rnbr -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2586)) (Prims.of_int (8)) - (Prims.of_int (2588)) (Prims.of_int (18))) - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2585)) (Prims.of_int (6)) - (Prims.of_int (2612)) (Prims.of_int (18))) - (if (lnbr + rnbr) <= Prims.int_one - then Obj.magic (Obj.repr (unfold_guard ())) - else - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___4 -> false)))) - (fun uu___3 -> - (fun uu___3 -> - if uu___3 - then - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2591)) - (Prims.of_int (8)) - (Prims.of_int (2610)) - (Prims.of_int (62))) - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2611)) - (Prims.of_int (8)) - (Prims.of_int (2611)) - (Prims.of_int (12))) - (Obj.magic - (FStar_Tactics_Derived.focus - (fun uu___4 -> - FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2591)) - (Prims.of_int (24)) - (Prims.of_int (2591)) - (Prims.of_int (56))) - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2592)) - (Prims.of_int (22)) - (Prims.of_int (2610)) - (Prims.of_int (61))) - (Obj.magic - (FStar_Tactics_Derived.apply_lemma - (FStar_Reflection_Builtins.pack_ln + fun uu___ -> + (fun args -> + match args with + | uu___::(t1, uu___1)::(t2, uu___2)::[] -> + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2583)) (Prims.of_int (17)) + (Prims.of_int (2583)) (Prims.of_int (36))) + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2584)) (Prims.of_int (6)) + (Prims.of_int (2612)) (Prims.of_int (18))) + (Obj.magic (slterm_nbr_uvars t1)) + (fun uu___3 -> + (fun lnbr -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2584)) (Prims.of_int (17)) + (Prims.of_int (2584)) (Prims.of_int (36))) + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2585)) (Prims.of_int (6)) + (Prims.of_int (2612)) (Prims.of_int (18))) + (Obj.magic (slterm_nbr_uvars t2)) + (fun uu___3 -> + (fun rnbr -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2586)) + (Prims.of_int (8)) + (Prims.of_int (2588)) + (Prims.of_int (18))) + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2585)) + (Prims.of_int (6)) + (Prims.of_int (2612)) + (Prims.of_int (18))) + (if (lnbr + rnbr) <= Prims.int_one + then + Obj.magic + (Obj.repr (unfold_guard ())) + else + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___4 -> false)))) + (fun uu___3 -> + (fun uu___3 -> + if uu___3 + then + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2591)) + (Prims.of_int (8)) + (Prims.of_int (2610)) + (Prims.of_int (62))) + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2611)) + (Prims.of_int (8)) + (Prims.of_int (2611)) + (Prims.of_int (12))) + (Obj.magic + (FStar_Tactics_Derived.focus + (fun uu___4 -> + FStar_Tactics_Effect.tac_bind + ( + Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2591)) + (Prims.of_int (24)) + (Prims.of_int (2591)) + (Prims.of_int (56))) + ( + Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2592)) + (Prims.of_int (22)) + (Prims.of_int (2610)) + (Prims.of_int (61))) ( - FStar_Reflection_Data.Tv_FVar + Obj.magic + (FStar_Tactics_Derived.apply_lemma + (FStar_Reflection_Builtins.pack_ln + (FStar_Reflection_Data.Tv_FVar (FStar_Reflection_Builtins.pack_fv ["Steel"; "Effect"; "Common"; "equiv_forall_elim"]))))) - (fun uu___5 -> - (fun uu___5 -> - Obj.magic - (FStar_Tactics_Effect.tac_bind + ( + fun + uu___5 -> + (fun + uu___5 -> + Obj.magic + (FStar_Tactics_Effect.tac_bind (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2592)) @@ -10880,114 +10937,134 @@ let (solve_equiv_forall : uu___9))) uu___8)))) uu___6))) - uu___5)))) - (fun uu___4 -> - FStar_Tactics_Effect.lift_div_tac - (fun uu___5 -> true)))) - else - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___5 -> false)))) - uu___3))) uu___3))) uu___3) - | uu___ -> - FStar_Tactics_Derived.fail - "Ill-formed equiv_forall, should not happen" + uu___5)))) + (fun uu___4 -> + FStar_Tactics_Effect.lift_div_tac + (fun uu___5 -> + true)))) + else + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___5 -> false)))) + uu___3))) uu___3))) uu___3))) + | uu___ -> + Obj.magic + (Obj.repr + (FStar_Tactics_Derived.fail + "Ill-formed equiv_forall, should not happen"))) uu___ let (solve_equiv : FStar_Reflection_Data.argv Prims.list -> (Prims.bool, unit) FStar_Tactics_Effect.tac_repr) = - fun args -> - match args with - | (t1, uu___)::(t2, uu___1)::[] -> - FStar_Tactics_Effect.tac_bind - (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2620)) - (Prims.of_int (17)) (Prims.of_int (2620)) (Prims.of_int (36))) - (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2621)) - (Prims.of_int (6)) (Prims.of_int (2645)) (Prims.of_int (18))) - (Obj.magic (slterm_nbr_uvars t1)) - (fun uu___2 -> - (fun lnbr -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2621)) (Prims.of_int (17)) - (Prims.of_int (2621)) (Prims.of_int (36))) - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2622)) (Prims.of_int (6)) - (Prims.of_int (2645)) (Prims.of_int (18))) - (Obj.magic (slterm_nbr_uvars t2)) - (fun uu___2 -> - (fun rnbr -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2623)) (Prims.of_int (8)) - (Prims.of_int (2625)) (Prims.of_int (18))) - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2622)) (Prims.of_int (6)) - (Prims.of_int (2645)) (Prims.of_int (18))) - (if (lnbr + rnbr) <= Prims.int_one - then Obj.magic (Obj.repr (unfold_guard ())) - else - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___3 -> false)))) - (fun uu___2 -> - (fun uu___2 -> - if uu___2 - then - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2628)) - (Prims.of_int (8)) - (Prims.of_int (2642)) - (Prims.of_int (48))) - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2643)) - (Prims.of_int (8)) - (Prims.of_int (2643)) - (Prims.of_int (12))) - (Obj.magic - (FStar_Tactics_Derived.focus - (fun uu___3 -> - FStar_Tactics_Derived.or_else - (fun uu___4 -> - FStar_Tactics_Derived.apply_lemma - (FStar_Reflection_Builtins.pack_ln - (FStar_Reflection_Data.Tv_FVar + fun uu___ -> + (fun args -> + match args with + | (t1, uu___)::(t2, uu___1)::[] -> + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2620)) (Prims.of_int (17)) + (Prims.of_int (2620)) (Prims.of_int (36))) + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2621)) (Prims.of_int (6)) + (Prims.of_int (2645)) (Prims.of_int (18))) + (Obj.magic (slterm_nbr_uvars t1)) + (fun uu___2 -> + (fun lnbr -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2621)) (Prims.of_int (17)) + (Prims.of_int (2621)) (Prims.of_int (36))) + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2622)) (Prims.of_int (6)) + (Prims.of_int (2645)) (Prims.of_int (18))) + (Obj.magic (slterm_nbr_uvars t2)) + (fun uu___2 -> + (fun rnbr -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2623)) + (Prims.of_int (8)) + (Prims.of_int (2625)) + (Prims.of_int (18))) + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2622)) + (Prims.of_int (6)) + (Prims.of_int (2645)) + (Prims.of_int (18))) + (if (lnbr + rnbr) <= Prims.int_one + then + Obj.magic + (Obj.repr (unfold_guard ())) + else + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___3 -> false)))) + (fun uu___2 -> + (fun uu___2 -> + if uu___2 + then + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2628)) + (Prims.of_int (8)) + (Prims.of_int (2642)) + (Prims.of_int (48))) + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2643)) + (Prims.of_int (8)) + (Prims.of_int (2643)) + (Prims.of_int (12))) + (Obj.magic + (FStar_Tactics_Derived.focus + (fun uu___3 -> + FStar_Tactics_Derived.or_else + ( + fun + uu___4 -> + FStar_Tactics_Derived.apply_lemma + (FStar_Reflection_Builtins.pack_ln + (FStar_Reflection_Data.Tv_FVar (FStar_Reflection_Builtins.pack_fv ["Steel"; "Effect"; "Common"; "equiv_refl"])))) - (fun uu___4 -> - FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2632)) - (Prims.of_int (14)) - (Prims.of_int (2632)) - (Prims.of_int (68))) - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2633)) - (Prims.of_int (14)) - (Prims.of_int (2642)) - (Prims.of_int (46))) - (if - (lnbr <> + ( + fun + uu___4 -> + FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2632)) + (Prims.of_int (14)) + (Prims.of_int (2632)) + (Prims.of_int (68))) + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2633)) + (Prims.of_int (14)) + (Prims.of_int (2642)) + (Prims.of_int (46))) + (if + (lnbr <> Prims.int_zero) && (rnbr = Prims.int_zero) - then - Obj.magic + then + Obj.magic (Obj.repr (FStar_Tactics_Derived.apply_lemma (FStar_Reflection_Builtins.pack_ln @@ -10997,16 +11074,16 @@ let (solve_equiv : "Effect"; "Common"; "equiv_sym"]))))) - else - Obj.magic + else + Obj.magic (Obj.repr (FStar_Tactics_Effect.lift_div_tac (fun uu___6 -> ())))) - (fun uu___5 - -> - (fun + (fun + uu___5 -> + (fun uu___5 -> Obj.magic (FStar_Tactics_Effect.tac_bind @@ -11075,121 +11152,144 @@ let (solve_equiv : (fun uu___7 -> (fun - uu___7 -> + uu___7 -> + Obj.magic + (canon' + false + (FStar_Reflection_Builtins.pack_ln + (FStar_Reflection_Data.Tv_FVar + (FStar_Reflection_Builtins.pack_fv + ["Steel"; + "Effect"; + "Common"; + "true_p"]))) + (FStar_Reflection_Builtins.pack_ln + (FStar_Reflection_Data.Tv_FVar + (FStar_Reflection_Builtins.pack_fv + ["Steel"; + "Effect"; + "Common"; + "true_p"]))))) + uu___7))) + uu___6))) + uu___5))))) + (fun uu___3 -> + FStar_Tactics_Effect.lift_div_tac + (fun uu___4 -> + true)))) + else + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___4 -> false)))) + uu___2))) uu___2))) uu___2))) + | uu___ -> + Obj.magic + (Obj.repr + (FStar_Tactics_Derived.fail + "Ill-formed equiv, should not happen"))) uu___ +let (solve_can_be_split_post : + FStar_Reflection_Data.argv Prims.list -> + (Prims.bool, unit) FStar_Tactics_Effect.tac_repr) + = + fun uu___ -> + (fun args -> + match args with + | uu___::uu___1::(t1, uu___2)::(t2, uu___3)::[] -> + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2653)) (Prims.of_int (17)) + (Prims.of_int (2653)) (Prims.of_int (36))) + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2654)) (Prims.of_int (6)) + (Prims.of_int (2685)) (Prims.of_int (18))) + (Obj.magic (slterm_nbr_uvars t1)) + (fun uu___4 -> + (fun lnbr -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2654)) (Prims.of_int (17)) + (Prims.of_int (2654)) (Prims.of_int (36))) + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2655)) (Prims.of_int (6)) + (Prims.of_int (2685)) (Prims.of_int (18))) + (Obj.magic (slterm_nbr_uvars t2)) + (fun uu___4 -> + (fun rnbr -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2656)) + (Prims.of_int (8)) + (Prims.of_int (2658)) + (Prims.of_int (18))) + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2655)) + (Prims.of_int (6)) + (Prims.of_int (2685)) + (Prims.of_int (18))) + (if (lnbr + rnbr) <= Prims.int_one + then + Obj.magic + (Obj.repr (unfold_guard ())) + else + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___5 -> false)))) + (fun uu___4 -> + (fun uu___4 -> + if uu___4 + then + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2661)) + (Prims.of_int (8)) + (Prims.of_int (2683)) + (Prims.of_int (62))) + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2684)) + (Prims.of_int (8)) + (Prims.of_int (2684)) + (Prims.of_int (12))) + (Obj.magic + (FStar_Tactics_Derived.focus + (fun uu___5 -> + FStar_Tactics_Effect.tac_bind + ( + Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2661)) + (Prims.of_int (24)) + (Prims.of_int (2661)) + (Prims.of_int (30))) + ( + Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2662)) + (Prims.of_int (22)) + (Prims.of_int (2683)) + (Prims.of_int (61))) + ( + Obj.magic + (FStar_Tactics_Builtins.norm + [])) + ( + fun + uu___6 -> + (fun + uu___6 -> Obj.magic - (canon' - false - (FStar_Reflection_Builtins.pack_ln - (FStar_Reflection_Data.Tv_FVar - (FStar_Reflection_Builtins.pack_fv - ["Steel"; - "Effect"; - "Common"; - "true_p"]))) - (FStar_Reflection_Builtins.pack_ln - (FStar_Reflection_Data.Tv_FVar - (FStar_Reflection_Builtins.pack_fv - ["Steel"; - "Effect"; - "Common"; - "true_p"]))))) - uu___7))) - uu___6))) - uu___5))))) - (fun uu___3 -> - FStar_Tactics_Effect.lift_div_tac - (fun uu___4 -> true)))) - else - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___4 -> false)))) - uu___2))) uu___2))) uu___2) - | uu___ -> - FStar_Tactics_Derived.fail "Ill-formed equiv, should not happen" -let (solve_can_be_split_post : - FStar_Reflection_Data.argv Prims.list -> - (Prims.bool, unit) FStar_Tactics_Effect.tac_repr) - = - fun args -> - match args with - | uu___::uu___1::(t1, uu___2)::(t2, uu___3)::[] -> - FStar_Tactics_Effect.tac_bind - (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2653)) - (Prims.of_int (17)) (Prims.of_int (2653)) (Prims.of_int (36))) - (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2654)) - (Prims.of_int (6)) (Prims.of_int (2685)) (Prims.of_int (18))) - (Obj.magic (slterm_nbr_uvars t1)) - (fun uu___4 -> - (fun lnbr -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2654)) (Prims.of_int (17)) - (Prims.of_int (2654)) (Prims.of_int (36))) - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2655)) (Prims.of_int (6)) - (Prims.of_int (2685)) (Prims.of_int (18))) - (Obj.magic (slterm_nbr_uvars t2)) - (fun uu___4 -> - (fun rnbr -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2656)) (Prims.of_int (8)) - (Prims.of_int (2658)) (Prims.of_int (18))) - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2655)) (Prims.of_int (6)) - (Prims.of_int (2685)) (Prims.of_int (18))) - (if (lnbr + rnbr) <= Prims.int_one - then Obj.magic (Obj.repr (unfold_guard ())) - else - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___5 -> false)))) - (fun uu___4 -> - (fun uu___4 -> - if uu___4 - then - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2661)) - (Prims.of_int (8)) - (Prims.of_int (2683)) - (Prims.of_int (62))) - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2684)) - (Prims.of_int (8)) - (Prims.of_int (2684)) - (Prims.of_int (12))) - (Obj.magic - (FStar_Tactics_Derived.focus - (fun uu___5 -> - FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2661)) - (Prims.of_int (24)) - (Prims.of_int (2661)) - (Prims.of_int (30))) - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2662)) - (Prims.of_int (22)) - (Prims.of_int (2683)) - (Prims.of_int (61))) - (Obj.magic - (FStar_Tactics_Builtins.norm - [])) - (fun uu___6 -> - (fun uu___6 -> - Obj.magic - (FStar_Tactics_Effect.tac_bind + (FStar_Tactics_Effect.tac_bind (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2662)) @@ -11538,17 +11638,22 @@ let (solve_can_be_split_post : uu___8))) uu___7))) uu___7))) - uu___6)))) - (fun uu___5 -> - FStar_Tactics_Effect.lift_div_tac - (fun uu___6 -> true)))) - else - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___6 -> false)))) - uu___4))) uu___4))) uu___4) - | uu___ -> FStar_Tactics_Derived.fail "ill-formed can_be_split_post" + uu___6)))) + (fun uu___5 -> + FStar_Tactics_Effect.lift_div_tac + (fun uu___6 -> + true)))) + else + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.lift_div_tac + (fun uu___6 -> false)))) + uu___4))) uu___4))) uu___4))) + | uu___ -> + Obj.magic + (Obj.repr + (FStar_Tactics_Derived.fail "ill-formed can_be_split_post"))) + uu___ let (is_return_eq : FStar_Reflection_Types.term -> FStar_Reflection_Types.term -> @@ -11924,116 +12029,118 @@ let (goal_to_equiv : match f with | FStar_Reflection_Formula.App (hd0, t1) -> Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2746)) (Prims.of_int (6)) - (Prims.of_int (2747)) (Prims.of_int (70))) - (Prims.mk_range "Steel.Effect.Common.fsti" - (Prims.of_int (2748)) (Prims.of_int (6)) - (Prims.of_int (2768)) (Prims.of_int (51))) - (if - Prims.op_Negation - (FStar_Reflection_Derived.is_fvar hd0 - "Prims.squash") - then - Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - (Prims.strcat loc - " unexpected non-squash goal in goal_to_equiv"))) - else - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___1 -> ())))) - (fun uu___ -> + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2746)) + (Prims.of_int (6)) + (Prims.of_int (2747)) + (Prims.of_int (70))) + (Prims.mk_range "Steel.Effect.Common.fsti" + (Prims.of_int (2748)) + (Prims.of_int (6)) + (Prims.of_int (2768)) + (Prims.of_int (51))) + (if + Prims.op_Negation + (FStar_Reflection_Derived.is_fvar hd0 + "Prims.squash") + then + FStar_Tactics_Derived.fail + (Prims.strcat loc + " unexpected non-squash goal in goal_to_equiv") + else + FStar_Tactics_Effect.lift_div_tac + (fun uu___1 -> ())) (fun uu___ -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2748)) - (Prims.of_int (21)) - (Prims.of_int (2748)) - (Prims.of_int (34))) - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2748)) - (Prims.of_int (6)) - (Prims.of_int (2768)) - (Prims.of_int (51))) - (FStar_Tactics_Effect.lift_div_tac - (fun uu___1 -> - FStar_Reflection_Derived.collect_app - t1)) - (fun uu___1 -> + (fun uu___ -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2748)) + (Prims.of_int (21)) + (Prims.of_int (2748)) + (Prims.of_int (34))) + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2748)) + (Prims.of_int (6)) + (Prims.of_int (2768)) + (Prims.of_int (51))) + (FStar_Tactics_Effect.lift_div_tac + (fun uu___1 -> + FStar_Reflection_Derived.collect_app + t1)) (fun uu___1 -> - match uu___1 with - | (hd, args) -> - if - FStar_Reflection_Derived.is_fvar - hd - "Steel.Effect.Common.can_be_split" - then - Obj.magic - (FStar_Tactics_Derived.apply_lemma - (FStar_Reflection_Builtins.pack_ln - (FStar_Reflection_Data.Tv_FVar - (FStar_Reflection_Builtins.pack_fv + (fun uu___1 -> + match uu___1 with + | (hd, args) -> + if + FStar_Reflection_Derived.is_fvar + hd + "Steel.Effect.Common.can_be_split" + then + Obj.magic + (Obj.repr + (FStar_Tactics_Derived.apply_lemma + (FStar_Reflection_Builtins.pack_ln + (FStar_Reflection_Data.Tv_FVar + (FStar_Reflection_Builtins.pack_fv ["Steel"; "Effect"; "Common"; - "equiv_can_be_split"])))) - else - if - FStar_Reflection_Derived.is_fvar - hd - "Steel.Effect.Common.can_be_split_forall" - then - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2752)) - (Prims.of_int (8)) - (Prims.of_int (2752)) - (Prims.of_int (32))) - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2753)) - (Prims.of_int (8)) - (Prims.of_int (2753)) - (Prims.of_int (41))) - (Obj.magic - (FStar_Tactics_Effect.tac_bind + "equiv_can_be_split"]))))) + else + Obj.magic + (Obj.repr + (if + FStar_Reflection_Derived.is_fvar + hd + "Steel.Effect.Common.can_be_split_forall" + then + Obj.repr ( - Prims.mk_range + FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2752)) + (Prims.of_int (8)) + (Prims.of_int (2752)) + (Prims.of_int (32))) + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2753)) + (Prims.of_int (8)) + (Prims.of_int (2753)) + (Prims.of_int (41))) + (Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2752)) (Prims.of_int (15)) (Prims.of_int (2752)) (Prims.of_int (32))) - ( - Prims.mk_range + (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2752)) (Prims.of_int (8)) (Prims.of_int (2752)) (Prims.of_int (32))) - ( - Obj.magic + (Obj.magic (FStar_Tactics_Logic.forall_intro ())) - ( - fun + (fun uu___3 -> FStar_Tactics_Effect.lift_div_tac (fun uu___4 -> ())))) - (fun uu___3 -> - (fun uu___3 - -> + (fun + uu___3 -> + (fun + uu___3 -> Obj.magic (FStar_Tactics_Derived.apply_lemma (FStar_Reflection_Builtins.pack_ln @@ -12043,29 +12150,31 @@ let (goal_to_equiv : "Effect"; "Common"; "equiv_can_be_split"]))))) - uu___3)) - else - if - FStar_Reflection_Derived.is_fvar - hd - "Steel.Effect.Common.equiv_forall" - then - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2755)) - (Prims.of_int (8)) - (Prims.of_int (2755)) - (Prims.of_int (40))) - (Prims.mk_range - "Steel.Effect.Common.fsti" - (Prims.of_int (2756)) - (Prims.of_int (8)) - (Prims.of_int (2756)) - (Prims.of_int (32))) - (Obj.magic - (FStar_Tactics_Derived.apply_lemma + uu___3)) + else + Obj.repr + ( + if + FStar_Reflection_Derived.is_fvar + hd + "Steel.Effect.Common.equiv_forall" + then + Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2755)) + (Prims.of_int (8)) + (Prims.of_int (2755)) + (Prims.of_int (40))) + (Prims.mk_range + "Steel.Effect.Common.fsti" + (Prims.of_int (2756)) + (Prims.of_int (8)) + (Prims.of_int (2756)) + (Prims.of_int (32))) + (Obj.magic + (FStar_Tactics_Derived.apply_lemma (FStar_Reflection_Builtins.pack_ln (FStar_Reflection_Data.Tv_FVar (FStar_Reflection_Builtins.pack_fv @@ -12073,9 +12182,9 @@ let (goal_to_equiv : "Effect"; "Common"; "equiv_forall_elim"]))))) - (fun uu___4 - -> - (fun + (fun + uu___4 -> + (fun uu___4 -> Obj.magic (FStar_Tactics_Effect.tac_bind @@ -12101,27 +12210,28 @@ let (goal_to_equiv : uu___6 -> ())))) uu___4)) - else - if - FStar_Reflection_Derived.is_fvar - hd - "Steel.Effect.Common.can_be_split_post" - then - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range + else + Obj.repr + (if + FStar_Reflection_Derived.is_fvar + hd + "Steel.Effect.Common.can_be_split_post" + then + Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2758)) (Prims.of_int (8)) (Prims.of_int (2758)) (Prims.of_int (45))) - (Prims.mk_range + (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (2759)) (Prims.of_int (8)) (Prims.of_int (2761)) (Prims.of_int (32))) - (Obj.magic + (Obj.magic (FStar_Tactics_Derived.apply_lemma (FStar_Reflection_Builtins.pack_ln (FStar_Reflection_Data.Tv_FVar @@ -12130,8 +12240,8 @@ let (goal_to_equiv : "Effect"; "Common"; "can_be_split_post_elim"]))))) - (fun uu___5 - -> + (fun + uu___5 -> (fun uu___5 -> Obj.magic @@ -12222,41 +12332,39 @@ let (goal_to_equiv : uu___7))) uu___6))) uu___5)) - else - if - FStar_Reflection_Derived.is_fvar - hd - "Steel.Effect.Common.can_be_split_dep" - then - Obj.magic - (FStar_Tactics_Derived.fail - ( - Prims.strcat + else + Obj.repr + (if + FStar_Reflection_Derived.is_fvar + hd + "Steel.Effect.Common.can_be_split_dep" + then + FStar_Tactics_Derived.fail + (Prims.strcat "can_be_split_dep not supported in " - loc)) - else - if - FStar_Reflection_Derived.is_fvar - hd - "Steel.Effect.Common.can_be_split_forall_dep" - then - Obj.magic - (FStar_Tactics_Derived.fail + loc) + else + if + FStar_Reflection_Derived.is_fvar + hd + "Steel.Effect.Common.can_be_split_forall_dep" + then + FStar_Tactics_Derived.fail (Prims.strcat "can_be_split_forall_dep not supported in " - loc)) - else - Obj.magic - (FStar_Tactics_Derived.fail + loc) + else + FStar_Tactics_Derived.fail (Prims.strcat loc - " goal in unexpected position"))) - uu___1))) uu___)) + " goal in unexpected position"))))))) + uu___1))) uu___))) | uu___ -> Obj.magic - (FStar_Tactics_Derived.fail - (Prims.strcat loc " unexpected goal"))) uu___))) - uu___) + (Obj.repr + (FStar_Tactics_Derived.fail + (Prims.strcat loc " unexpected goal")))) + uu___))) uu___) let rec term_dict_assoc : 'a . FStar_Reflection_Types.term -> @@ -13311,11 +13419,13 @@ let rec (resolve_tac : if uu___2 then Obj.magic - (resolve_tac dict) + (Obj.repr + (resolve_tac dict)) else Obj.magic - (FStar_Tactics_Derived.fail - "Could not make progress, no solvable goal found")) + (Obj.repr + (FStar_Tactics_Derived.fail + "Could not make progress, no solvable goal found"))) uu___2))) uu___2))) uu___1)))) uu___) let rec (pick_next_logical : @@ -13604,14 +13714,16 @@ let (is_true : match uu___1 with | FStar_Reflection_Formula.True_ -> Obj.magic - (FStar_Tactics_Derived.exact - (FStar_Reflection_Builtins.pack_ln - (FStar_Reflection_Data.Tv_Const - FStar_Reflection_Data.C_Unit))) + (Obj.repr + (FStar_Tactics_Derived.exact + (FStar_Reflection_Builtins.pack_ln + (FStar_Reflection_Data.Tv_Const + FStar_Reflection_Data.C_Unit)))) | uu___2 -> Obj.magic - (FStar_Tactics_Effect.raise - FStar_Tactics_Derived.Goal_not_trivial)) uu___1) + (Obj.repr + (FStar_Tactics_Effect.raise + FStar_Tactics_Derived.Goal_not_trivial))) uu___1) let rec (solve_maybe_emps : Prims.nat -> (unit, unit) FStar_Tactics_Effect.tac_repr) = fun uu___ -> @@ -14529,11 +14641,13 @@ let (ite_soundness_tac : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) match uu___4 with | [] -> Obj.magic - (FStar_Tactics_Derived.fail - "should not happen") + (Obj.repr + (FStar_Tactics_Derived.fail + "should not happen")) | uu___5::tl -> Obj.magic - (FStar_Tactics_Effect.tac_bind + (Obj.repr + (FStar_Tactics_Effect.tac_bind (Prims.mk_range "Steel.Effect.Common.fsti" (Prims.of_int (3153)) @@ -14751,7 +14865,7 @@ let (ite_soundness_tac : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) uu___9))) uu___8))) uu___7))) - uu___6))) + uu___6)))) uu___4))) uu___4))) uu___3))) uu___2))) uu___1) let (vc_norm : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = diff --git a/ocaml/fstar-lib/generated/Steel_ST_GenElim_Base.ml b/ocaml/fstar-lib/generated/Steel_ST_GenElim_Base.ml index ad134920e48..2ca503ea1dc 100644 --- a/ocaml/fstar-lib/generated/Steel_ST_GenElim_Base.ml +++ b/ocaml/fstar-lib/generated/Steel_ST_GenElim_Base.ml @@ -240,24 +240,28 @@ let rec (solve_gen_unit_elim : (t2, FStar_Reflection_Data.Q_Explicit)::[] -> - FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.ST.GenElim.Base.fsti" - (Prims.of_int (335)) - (Prims.of_int (20)) - (Prims.of_int (335)) - (Prims.of_int (42))) - (Prims.mk_range - "Steel.ST.GenElim.Base.fsti" - (Prims.of_int (336)) - (Prims.of_int (10)) - (Prims.of_int (337)) - (Prims.of_int (68))) - (Obj.magic - (solve_gen_unit_elim + Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.ST.GenElim.Base.fsti" + (Prims.of_int (335)) + (Prims.of_int (20)) + (Prims.of_int (335)) + (Prims.of_int (42))) + (Prims.mk_range + "Steel.ST.GenElim.Base.fsti" + (Prims.of_int (336)) + (Prims.of_int (10)) + (Prims.of_int (337)) + (Prims.of_int (68))) + (Obj.magic + ( + solve_gen_unit_elim t1)) - (fun uu___5 -> - (fun t1' -> + (fun uu___5 + -> + (fun t1' + -> Obj.magic (FStar_Tactics_Effect.tac_bind (Prims.mk_range @@ -294,10 +298,11 @@ let rec (solve_gen_unit_elim : FStar_Reflection_Data.Q_Explicit); (t2', FStar_Reflection_Data.Q_Explicit)])))) - uu___5) + uu___5)) | uu___5 -> - FStar_Tactics_Derived.fail - "ill-formed star")) + Obj.repr + (FStar_Tactics_Derived.fail + "ill-formed star"))) else Obj.magic (Obj.repr @@ -429,25 +434,19 @@ let rec (solve_gen_elim : (body, FStar_Reflection_Data.Q_Explicit)::[] -> - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___3 -> - ([(ty, - FStar_Reflection_Data.Q_Implicit)], - body)))) + FStar_Tactics_Effect.lift_div_tac + (fun uu___3 -> + ([(ty, + FStar_Reflection_Data.Q_Implicit)], + body)) | (body, FStar_Reflection_Data.Q_Explicit)::[] -> - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___3 -> ([], body)))) + FStar_Tactics_Effect.lift_div_tac + (fun uu___3 -> ([], body)) | uu___3 -> - Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "ill-formed exists_"))) + FStar_Tactics_Derived.fail + "ill-formed exists_") (fun uu___3 -> (fun uu___3 -> match uu___3 with @@ -805,35 +804,37 @@ let rec (solve_gen_elim : (tr, FStar_Reflection_Data.Q_Explicit)::[] -> - FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.ST.GenElim.Base.fsti" - (Prims.of_int (383)) - (Prims.of_int (15)) - (Prims.of_int (383)) - (Prims.of_int (42))) - (Prims.mk_range - "Steel.ST.GenElim.Base.fsti" - (Prims.of_int (383)) - (Prims.of_int (12)) - (Prims.of_int (396)) - (Prims.of_int (72))) - (Obj.magic - (term_has_head tl - (FStar_Reflection_Builtins.pack_ln - (FStar_Reflection_Data.Tv_FVar + Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.ST.GenElim.Base.fsti" + (Prims.of_int (383)) + (Prims.of_int (15)) + (Prims.of_int (383)) + (Prims.of_int (42))) + (Prims.mk_range + "Steel.ST.GenElim.Base.fsti" + (Prims.of_int (383)) + (Prims.of_int (12)) + (Prims.of_int (396)) + (Prims.of_int (72))) + (Obj.magic + (term_has_head + tl + (FStar_Reflection_Builtins.pack_ln ( - FStar_Reflection_Builtins.pack_fv + FStar_Reflection_Data.Tv_FVar + (FStar_Reflection_Builtins.pack_fv ["Steel"; "ST"; "Util"; "exists_"]))))) - (fun uu___5 -> (fun uu___5 -> - if uu___5 - then - Obj.magic - (FStar_Tactics_Effect.tac_bind + (fun uu___5 -> + if uu___5 + then + Obj.magic + (FStar_Tactics_Effect.tac_bind (Prims.mk_range "Steel.ST.GenElim.Base.fsti" (Prims.of_int (385)) @@ -958,9 +959,9 @@ let rec (solve_gen_elim : FStar_Reflection_Data.Q_Explicit)])))) uu___6))) uu___6)) - else - Obj.magic - (FStar_Tactics_Effect.tac_bind + else + Obj.magic + (FStar_Tactics_Effect.tac_bind (Prims.mk_range "Steel.ST.GenElim.Base.fsti" (Prims.of_int (394)) @@ -1017,10 +1018,11 @@ let rec (solve_gen_elim : (tr', FStar_Reflection_Data.Q_Explicit)])))) uu___7))) - uu___5) + uu___5)) | uu___5 -> - FStar_Tactics_Derived.fail - "ill-formed star")) + Obj.repr + (FStar_Tactics_Derived.fail + "ill-formed star"))) else Obj.magic (Obj.repr @@ -1782,78 +1784,70 @@ let (solve_gen_elim_prop : FStar_Tactics_Effect.lift_div_tac (fun uu___3 -> Prims.op_Negation uu___2)))) (fun uu___2 -> - (fun uu___2 -> - if uu___2 - then - Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "not a squash goal")) - else - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___4 -> ())))) uu___2))) + if uu___2 + then + FStar_Tactics_Derived.fail + "not a squash goal" + else + FStar_Tactics_Effect.lift_div_tac + (fun uu___4 -> ())))) (fun uu___2 -> (fun uu___2 -> match tl with | (body1, FStar_Reflection_Data.Q_Explicit)::[] -> Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.ST.GenElim.Base.fsti" - (Prims.of_int (571)) - (Prims.of_int (21)) - (Prims.of_int (571)) - (Prims.of_int (40))) - (Prims.mk_range - "Steel.ST.GenElim.Base.fsti" - (Prims.of_int (571)) - (Prims.of_int (4)) - (Prims.of_int (603)) - (Prims.of_int (7))) - (FStar_Tactics_Effect.lift_div_tac - (fun uu___3 -> - FStar_Reflection_Derived.collect_app - body1)) - (fun uu___3 -> + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.ST.GenElim.Base.fsti" + (Prims.of_int (571)) + (Prims.of_int (21)) + (Prims.of_int (571)) + (Prims.of_int (40))) + (Prims.mk_range + "Steel.ST.GenElim.Base.fsti" + (Prims.of_int (571)) + (Prims.of_int (4)) + (Prims.of_int (603)) + (Prims.of_int (7))) + (FStar_Tactics_Effect.lift_div_tac + (fun uu___3 -> + FStar_Reflection_Derived.collect_app + body1)) (fun uu___3 -> - match uu___3 with - | (hd1, tl1) -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.ST.GenElim.Base.fsti" - (Prims.of_int (572)) - (Prims.of_int (4)) - (Prims.of_int (573)) - (Prims.of_int (42))) - (Prims.mk_range - "Steel.ST.GenElim.Base.fsti" - (Prims.of_int (574)) - (Prims.of_int (10)) - (Prims.of_int (602)) - (Prims.of_int (44))) - (if - Prims.op_Negation - (is_fvar hd1 - "Steel.ST.GenElim.Base.gen_elim_prop") - then - Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "not a gen_elim_prop goal")) - else - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___5 -> - ())))) - (fun uu___4 -> + (fun uu___3 -> + match uu___3 with + | (hd1, tl1) -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.ST.GenElim.Base.fsti" + (Prims.of_int (572)) + (Prims.of_int (4)) + (Prims.of_int (573)) + (Prims.of_int (42))) + (Prims.mk_range + "Steel.ST.GenElim.Base.fsti" + (Prims.of_int (574)) + (Prims.of_int (10)) + (Prims.of_int (602)) + (Prims.of_int (44))) + (if + Prims.op_Negation + (is_fvar hd1 + "Steel.ST.GenElim.Base.gen_elim_prop") + then + FStar_Tactics_Derived.fail + "not a gen_elim_prop goal" + else + FStar_Tactics_Effect.lift_div_tac + (fun uu___5 -> ())) (fun uu___4 -> - match FStar_List_Tot_Base.filter - (fun uu___5 - -> + (fun uu___4 -> + match FStar_List_Tot_Base.filter + ( + fun + uu___5 -> match uu___5 with | @@ -1861,35 +1855,37 @@ let (solve_gen_elim_prop : x) -> FStar_Reflection_Data.uu___is_Q_Explicit x) tl1 - with - | (enable_nondep_opt_tm, - uu___5):: - (p, uu___6):: - (a, uu___7):: - (q, uu___8):: - (post, uu___9)::[] - -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range + with + | (enable_nondep_opt_tm, + uu___5):: + (p, uu___6):: + (a, uu___7):: + (q, uu___8):: + (post, + uu___9)::[] + -> + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.ST.GenElim.Base.fsti" (Prims.of_int (576)) (Prims.of_int (30)) (Prims.of_int (576)) (Prims.of_int (74))) - (Prims.mk_range + (Prims.mk_range "Steel.ST.GenElim.Base.fsti" (Prims.of_int (577)) (Prims.of_int (6)) (Prims.of_int (601)) (Prims.of_int (44))) - (Obj.magic + (Obj.magic (FStar_Tactics_Builtins.term_eq_old enable_nondep_opt_tm (FStar_Reflection_Builtins.pack_ln (FStar_Reflection_Data.Tv_Const FStar_Reflection_Data.C_True)))) - (fun + (fun uu___10 -> (fun @@ -2416,16 +2412,19 @@ let (solve_gen_elim_prop : uu___10))) uu___10))) uu___10))) - uu___10)) - | uu___5 -> - Obj.magic - (FStar_Tactics_Derived.fail - "ill formed gen_elim_prop")) - uu___4))) uu___3)) + uu___10))) + | uu___5 -> + Obj.magic + (Obj.repr + (FStar_Tactics_Derived.fail + "ill formed gen_elim_prop"))) + uu___4))) uu___3))) | uu___3 -> Obj.magic - (FStar_Tactics_Derived.fail - "ill-formed squash")) uu___2))) uu___1) + (Obj.repr + (FStar_Tactics_Derived.fail + "ill-formed squash"))) uu___2))) + uu___1) let (solve_gen_elim_prop_placeholder : unit -> (Prims.bool, unit) FStar_Tactics_Effect.tac_repr) = fun uu___ -> @@ -2477,78 +2476,70 @@ let (solve_gen_elim_prop_placeholder : FStar_Tactics_Effect.lift_div_tac (fun uu___3 -> Prims.op_Negation uu___2)))) (fun uu___2 -> - (fun uu___2 -> - if uu___2 - then - Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "not a squash goal")) - else - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___4 -> ())))) uu___2))) + if uu___2 + then + FStar_Tactics_Derived.fail + "not a squash goal" + else + FStar_Tactics_Effect.lift_div_tac + (fun uu___4 -> ())))) (fun uu___2 -> (fun uu___2 -> match tl with | (body1, FStar_Reflection_Data.Q_Explicit)::[] -> Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.ST.GenElim.Base.fsti" - (Prims.of_int (615)) - (Prims.of_int (21)) - (Prims.of_int (615)) - (Prims.of_int (40))) - (Prims.mk_range - "Steel.ST.GenElim.Base.fsti" - (Prims.of_int (615)) - (Prims.of_int (4)) - (Prims.of_int (643)) - (Prims.of_int (7))) - (FStar_Tactics_Effect.lift_div_tac - (fun uu___3 -> - FStar_Reflection_Derived.collect_app - body1)) - (fun uu___3 -> + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.ST.GenElim.Base.fsti" + (Prims.of_int (615)) + (Prims.of_int (21)) + (Prims.of_int (615)) + (Prims.of_int (40))) + (Prims.mk_range + "Steel.ST.GenElim.Base.fsti" + (Prims.of_int (615)) + (Prims.of_int (4)) + (Prims.of_int (643)) + (Prims.of_int (7))) + (FStar_Tactics_Effect.lift_div_tac + (fun uu___3 -> + FStar_Reflection_Derived.collect_app + body1)) (fun uu___3 -> - match uu___3 with - | (hd1, tl1) -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range - "Steel.ST.GenElim.Base.fsti" - (Prims.of_int (616)) - (Prims.of_int (4)) - (Prims.of_int (617)) - (Prims.of_int (54))) - (Prims.mk_range - "Steel.ST.GenElim.Base.fsti" - (Prims.of_int (618)) - (Prims.of_int (10)) - (Prims.of_int (642)) - (Prims.of_int (56))) - (if - Prims.op_Negation - (is_fvar hd1 - "Steel.ST.GenElim.Base.gen_elim_prop_placeholder") - then - Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "not a gen_elim_prop_placeholder goal")) - else - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac - (fun uu___5 -> - ())))) - (fun uu___4 -> + (fun uu___3 -> + match uu___3 with + | (hd1, tl1) -> + Obj.magic + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range + "Steel.ST.GenElim.Base.fsti" + (Prims.of_int (616)) + (Prims.of_int (4)) + (Prims.of_int (617)) + (Prims.of_int (54))) + (Prims.mk_range + "Steel.ST.GenElim.Base.fsti" + (Prims.of_int (618)) + (Prims.of_int (10)) + (Prims.of_int (642)) + (Prims.of_int (56))) + (if + Prims.op_Negation + (is_fvar hd1 + "Steel.ST.GenElim.Base.gen_elim_prop_placeholder") + then + FStar_Tactics_Derived.fail + "not a gen_elim_prop_placeholder goal" + else + FStar_Tactics_Effect.lift_div_tac + (fun uu___5 -> ())) (fun uu___4 -> - match FStar_List_Tot_Base.filter - (fun uu___5 - -> + (fun uu___4 -> + match FStar_List_Tot_Base.filter + ( + fun + uu___5 -> match uu___5 with | @@ -2556,29 +2547,31 @@ let (solve_gen_elim_prop_placeholder : x) -> FStar_Reflection_Data.uu___is_Q_Explicit x) tl1 - with - | (enable_nondep_opt_tm, - uu___5):: - (p, uu___6):: - (a, uu___7):: - (q, uu___8):: - (post, uu___9)::[] - -> - Obj.magic - (FStar_Tactics_Effect.tac_bind - (Prims.mk_range + with + | (enable_nondep_opt_tm, + uu___5):: + (p, uu___6):: + (a, uu___7):: + (q, uu___8):: + (post, + uu___9)::[] + -> + Obj.magic + (Obj.repr + (FStar_Tactics_Effect.tac_bind + (Prims.mk_range "Steel.ST.GenElim.Base.fsti" (Prims.of_int (620)) (Prims.of_int (6)) (Prims.of_int (621)) (Prims.of_int (47))) - (Prims.mk_range + (Prims.mk_range "Steel.ST.GenElim.Base.fsti" (Prims.of_int (622)) (Prims.of_int (6)) (Prims.of_int (641)) (Prims.of_int (10))) - (Obj.magic + (Obj.magic (FStar_Tactics_Effect.tac_bind (Prims.mk_range "Steel.ST.GenElim.Base.fsti" @@ -2622,25 +2615,17 @@ let (solve_gen_elim_prop_placeholder : (fun uu___10 -> - (fun - uu___10 - -> if uu___10 then - Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "pre-resource not solved yet")) + FStar_Tactics_Derived.fail + "pre-resource not solved yet" else - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac + FStar_Tactics_Effect.lift_div_tac (fun uu___12 -> ())))) - uu___10))) - (fun + (fun uu___10 -> (fun @@ -2806,17 +2791,13 @@ let (solve_gen_elim_prop_placeholder : && post_is_uvar) then - Obj.magic - (Obj.repr - (FStar_Tactics_Derived.fail - "gen_elim_prop_placeholder is already solved")) + FStar_Tactics_Derived.fail + "gen_elim_prop_placeholder is already solved" else - Obj.magic - (Obj.repr - (FStar_Tactics_Effect.lift_div_tac + FStar_Tactics_Effect.lift_div_tac (fun uu___12 - -> ())))) + -> ())) (fun uu___11 -> @@ -3215,16 +3196,19 @@ let (solve_gen_elim_prop_placeholder : uu___11))) uu___11))) uu___11))) - uu___10)) - | uu___5 -> - Obj.magic - (FStar_Tactics_Derived.fail - "ill-formed gen_elim_prop_placeholder")) - uu___4))) uu___3)) + uu___10))) + | uu___5 -> + Obj.magic + (Obj.repr + (FStar_Tactics_Derived.fail + "ill-formed gen_elim_prop_placeholder"))) + uu___4))) uu___3))) | uu___3 -> Obj.magic - (FStar_Tactics_Derived.fail - "ill-formed squash")) uu___2))) uu___1) + (Obj.repr + (FStar_Tactics_Derived.fail + "ill-formed squash"))) uu___2))) + uu___1) let (init_resolve_tac : unit -> (unit, unit) FStar_Tactics_Effect.tac_repr) = fun uu___ -> Steel_Effect_Common.init_resolve_tac' From a7c7f78b528c70af6d2a17af134af6365a1bedac Mon Sep 17 00:00:00 2001 From: Nikhil Swamy Date: Tue, 21 Mar 2023 15:28:40 -0700 Subject: [PATCH 26/48] handle cases where the file does not end with a new line --- ocaml/fstar-lib/FStar_Parser_ParseIt.ml | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/ocaml/fstar-lib/FStar_Parser_ParseIt.ml b/ocaml/fstar-lib/FStar_Parser_ParseIt.ml index 554dac2d99d..29428f6632e 100644 --- a/ocaml/fstar-lib/FStar_Parser_ParseIt.ml +++ b/ocaml/fstar-lib/FStar_Parser_ParseIt.ml @@ -148,7 +148,7 @@ let parse fn = let lines = U.splitlines contents in fun (start_pos:Lexing.position) (end_pos:Lexing.position) -> let suffix = FStar_Compiler_Util.nth_tail (Z.of_int (if start_pos.pos_lnum > 0 then start_pos.pos_lnum - 1 else 0)) lines in - let text, _ = FStar_Compiler_Util.first_N (Z.of_int (end_pos.pos_lnum - start_pos.pos_lnum)) suffix in + let text, _ = FStar_Compiler_Util.first_N (Z.of_int (if end_pos.pos_lnum = start_pos.pos_lnum then 1 else end_pos.pos_lnum - start_pos.pos_lnum)) suffix in let range = range_of_positions start_pos end_pos in { range; code = FStar_String.concat "\n" text } From cd233a86d76bcdd6eac2fde6f1997d07e7b2b839 Mon Sep 17 00:00:00 2001 From: Nikhil Swamy Date: Tue, 21 Mar 2023 21:23:04 -0700 Subject: [PATCH 27/48] better handling of ReloadDeps; and better handling of file that do not end with a newline --- ocaml/fstar-lib/FStar_Parser_ParseIt.ml | 22 ++- .../generated/FStar_Interactive_Ide_Types.ml | 2 +- .../FStar_Interactive_Incremental.ml | 30 +++-- .../generated/FStar_Interactive_ReplState.ml | 1 - src/Makefile.boot | 5 +- src/fstar/FStar.Interactive.Ide.fst | 1 - src/fstar/FStar.Interactive.Incremental.fst | 12 +- src/fstar/FStar.Interactive.Incremental.fsti | 1 - src/fstar/FStar.Interactive.QueryHelper.fst | 1 - src/fstar/FStar.Interactive.ReplState.fst | 127 ------------------ 10 files changed, 52 insertions(+), 150 deletions(-) delete mode 100644 ocaml/fstar-lib/generated/FStar_Interactive_ReplState.ml delete mode 100755 src/fstar/FStar.Interactive.ReplState.fst diff --git a/ocaml/fstar-lib/FStar_Parser_ParseIt.ml b/ocaml/fstar-lib/FStar_Parser_ParseIt.ml index 29428f6632e..0f9f543598d 100644 --- a/ocaml/fstar-lib/FStar_Parser_ParseIt.ml +++ b/ocaml/fstar-lib/FStar_Parser_ParseIt.ml @@ -148,7 +148,27 @@ let parse fn = let lines = U.splitlines contents in fun (start_pos:Lexing.position) (end_pos:Lexing.position) -> let suffix = FStar_Compiler_Util.nth_tail (Z.of_int (if start_pos.pos_lnum > 0 then start_pos.pos_lnum - 1 else 0)) lines in - let text, _ = FStar_Compiler_Util.first_N (Z.of_int (if end_pos.pos_lnum = start_pos.pos_lnum then 1 else end_pos.pos_lnum - start_pos.pos_lnum)) suffix in + let text, rest = FStar_Compiler_Util.first_N (Z.of_int (end_pos.pos_lnum - start_pos.pos_lnum)) suffix in + let text = + match rest with + | last::_ -> + let col = end_pos.pos_cnum - end_pos.pos_bol in + if col > 0 + then ( + let chars = FStar_String.list_of_string last in + if col <= List.length chars + then ( + let chars, _ = FStar_Compiler_Util.first_N (Z.of_int col) chars in + let last = FStar_String.string_of_list chars in + text@[last] + ) + else ( + text + ) + ) + else text + | _ -> text + in let range = range_of_positions start_pos end_pos in { range; code = FStar_String.concat "\n" text } diff --git a/ocaml/fstar-lib/generated/FStar_Interactive_Ide_Types.ml b/ocaml/fstar-lib/generated/FStar_Interactive_Ide_Types.ml index 295c26eb7e2..0a2ba4e0a3b 100644 --- a/ocaml/fstar-lib/generated/FStar_Interactive_Ide_Types.ml +++ b/ocaml/fstar-lib/generated/FStar_Interactive_Ide_Types.ml @@ -437,7 +437,7 @@ let (repl_state_to_string : repl_state -> Prims.string) = uu___3 :: uu___4 in uu___1 :: uu___2 in FStar_Compiler_Util.format - "{\n\trepl_line=%s;\n\trepl_column=%s;\n\trepl_fname=%s;\n\trepl_cur_mod=%s;\n\t\\ \n repl_deps_stack={%s}\n}" + "{\n\trepl_line=%s;\n\trepl_column=%s;\n\trepl_fname=%s;\n\trepl_cur_mod=%s;\n\t\\ \r\n repl_deps_stack={%s}\n}" uu___ let (push_query_to_string : push_query -> Prims.string) = fun pq -> diff --git a/ocaml/fstar-lib/generated/FStar_Interactive_Incremental.ml b/ocaml/fstar-lib/generated/FStar_Interactive_Incremental.ml index 21b0424bc6c..debb99a2338 100644 --- a/ocaml/fstar-lib/generated/FStar_Interactive_Incremental.ml +++ b/ocaml/fstar-lib/generated/FStar_Interactive_Incremental.ml @@ -368,16 +368,28 @@ let (run_full_buffer : match parse_result with | FStar_Parser_ParseIt.IncrementalFragment (decls, uu___, err_opt) -> - let decls1 = filter_decls decls in - (match request_type with - | FStar_Interactive_Ide_Types.ReloadDeps -> - let uu___1 = - let uu___2 = - FStar_Compiler_Effect.op_Bang - FStar_Interactive_PushHelper.repl_stack in - reload_deps uu___2 in - run_qst uu___1 qid1 + (match (request_type, decls) with + | (FStar_Interactive_Ide_Types.ReloadDeps, d::uu___1) -> + let uu___2 = + let uu___3 = + let uu___4 = + FStar_Compiler_Effect.op_Bang + FStar_Interactive_PushHelper.repl_stack in + reload_deps uu___4 in + op_let_Bang uu___3 + (fun queries -> + let uu___4 = + push_decl + FStar_Interactive_Ide_Types.FullCheck + write_full_buffer_fragment_progress d in + op_let_Bang uu___4 + (fun push_mod -> + return + (FStar_Compiler_List.op_At queries + push_mod))) in + run_qst uu___2 qid1 | uu___1 -> + let decls1 = filter_decls decls in let push_kind = match request_type with | FStar_Interactive_Ide_Types.LaxToPosition uu___2 diff --git a/ocaml/fstar-lib/generated/FStar_Interactive_ReplState.ml b/ocaml/fstar-lib/generated/FStar_Interactive_ReplState.ml deleted file mode 100644 index e8306abedb2..00000000000 --- a/ocaml/fstar-lib/generated/FStar_Interactive_ReplState.ml +++ /dev/null @@ -1 +0,0 @@ -open Prims \ No newline at end of file diff --git a/src/Makefile.boot b/src/Makefile.boot index 915a9f098cd..c7e009099e3 100644 --- a/src/Makefile.boot +++ b/src/Makefile.boot @@ -41,11 +41,10 @@ EXTRACT_MODULES=FStar.Pervasives FStar.Common FStar.Compiler.Range FStar.Thunk FStar.Interactive.JsonHelper FStar.Interactive.QueryHelper \ FStar.Interactive.PushHelper FStar.Interactive.Lsp \ FStar.Interactive.Ide FStar.Interactive.Ide.Types \ - FStar.Interactive.ReplState \ - FStar.Interactive.Incremental FStar.Interactive.Legacy \ + FStar.Interactive.Incremental FStar.Interactive.Legacy \ FStar.CheckedFiles FStar.Universal FStar.Prettyprint \ FStar.Main FStar.Compiler.List FStar.Compiler.Option \ - FStar.Compiler.Dyn + FStar.Compiler.Dyn # And there are a few specific files that should not be extracted at # all, despite being in one of the EXTRACT_NAMESPACES diff --git a/src/fstar/FStar.Interactive.Ide.fst b/src/fstar/FStar.Interactive.Ide.fst index 94c30fb0dfc..0ade7cb0da4 100644 --- a/src/fstar/FStar.Interactive.Ide.fst +++ b/src/fstar/FStar.Interactive.Ide.fst @@ -28,7 +28,6 @@ open FStar.Errors open FStar.Interactive.JsonHelper open FStar.Interactive.QueryHelper open FStar.Interactive.PushHelper -open FStar.Interactive.ReplState open FStar.Interactive.Ide.Types module BU = FStar.Compiler.Util diff --git a/src/fstar/FStar.Interactive.Incremental.fst b/src/fstar/FStar.Interactive.Incremental.fst index f194c493d25..5ee5775c6e8 100755 --- a/src/fstar/FStar.Interactive.Incremental.fst +++ b/src/fstar/FStar.Interactive.Incremental.fst @@ -39,7 +39,6 @@ module TcErr = FStar.TypeChecker.Err module TcEnv = FStar.TypeChecker.Env module CTable = FStar.Interactive.CompletionTable open FStar.Interactive.Ide.Types -open FStar.Interactive.ReplState module P = FStar.Parser.ParseIt module BU = FStar.Compiler.Util open FStar.Parser.AST @@ -250,12 +249,15 @@ let run_full_buffer (st:repl_state) let qs = match parse_result with | IncrementalFragment (decls, _, err_opt) -> ( - let decls = filter_decls decls in - match request_type with - | ReloadDeps -> - run_qst (reload_deps (!repl_stack)) qid + match request_type, decls with + | ReloadDeps, d::_ -> + run_qst (let! queries = reload_deps (!repl_stack) in + let! push_mod = push_decl FullCheck write_full_buffer_fragment_progress d in + return (queries @ push_mod)) + qid | _ -> + let decls = filter_decls decls in let push_kind = match request_type with | LaxToPosition _ -> LaxCheck diff --git a/src/fstar/FStar.Interactive.Incremental.fsti b/src/fstar/FStar.Interactive.Incremental.fsti index 8918f269377..f45f49a9d51 100755 --- a/src/fstar/FStar.Interactive.Incremental.fsti +++ b/src/fstar/FStar.Interactive.Incremental.fsti @@ -20,7 +20,6 @@ open FStar.Compiler.Effect open FStar.Compiler.List open FStar open FStar.Compiler -open FStar.Interactive.ReplState open FStar.Parser.AST open FStar.Errors open FStar.Interactive.Ide.Types diff --git a/src/fstar/FStar.Interactive.QueryHelper.fst b/src/fstar/FStar.Interactive.QueryHelper.fst index 3983ccf0aa3..ce5d202c4f2 100644 --- a/src/fstar/FStar.Interactive.QueryHelper.fst +++ b/src/fstar/FStar.Interactive.QueryHelper.fst @@ -28,7 +28,6 @@ open FStar.Compiler.Util open FStar.TypeChecker.Env open FStar.TypeChecker.Common open FStar.Interactive.JsonHelper -open FStar.Interactive.ReplState open FStar.Interactive.CompletionTable module U = FStar.Compiler.Util diff --git a/src/fstar/FStar.Interactive.ReplState.fst b/src/fstar/FStar.Interactive.ReplState.fst deleted file mode 100755 index 362f1a890a7..00000000000 --- a/src/fstar/FStar.Interactive.ReplState.fst +++ /dev/null @@ -1,127 +0,0 @@ -(* - Copyright 2019 Microsoft Research - - Licensed under the Apache License, Version 2.0 (the "License"); - you may not use this file except in compliance with the License. - You may obtain a copy of the License at - - http://www.apache.org/licenses/LICENSE-2.0 - - Unless required by applicable law or agreed to in writing, software - distributed under the License is distributed on an "AS IS" BASIS, - WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or implied. - See the License for the specific language governing permissions and - limitations under the License. -*) - -module FStar.Interactive.ReplState -open FStar -open FStar.Compiler -open FStar.Errors -open FStar.Compiler.Util -open FStar.Compiler.Range -open FStar.TypeChecker.Env -(* -module U = FStar.Compiler.Util -module TcEnv = FStar.TypeChecker.Env -module PI = FStar.Parser.ParseIt -module CTable = FStar.Interactive.CompletionTable - - -(* Types concerning repl *) -type repl_depth_t = TcEnv.tcenv_depth_t * int -type optmod_t = option Syntax.Syntax.modul - -type timed_fname = - { tf_fname: string; - tf_modtime: time } - -(** Every snapshot pushed in the repl stack is annotated with one of these. The -``LD``-prefixed (“Load Dependency”) onces are useful when loading or updating -dependencies, as they carry enough information to determine whether a dependency -is stale. **) -type repl_task = - | LDInterleaved of timed_fname * timed_fname (* (interface * implementation) *) - | LDSingle of timed_fname (* interface or implementation *) - | LDInterfaceOfCurrentFile of timed_fname (* interface *) - | PushFragment of either PI.input_frag FStar.Parser.AST.decl (* code fragment *) - | Noop (* Used by compute *) - -type repl_state (a:Type) = { - repl_line: int; - repl_column: int; - repl_fname: string; - repl_deps_stack: repl_stack_t a; - repl_curmod: optmod_t; - repl_env: TcEnv.env; - repl_stdin: stream_reader; - repl_names: CTable.table; - repl_buffered_input_queries: list a -} -and repl_stack_t a = list (repl_stack_entry_t a) -and repl_stack_entry_t a = repl_depth_t * (repl_task * repl_state a) - -// Global repl_state, keeping state of different buffers -type grepl_state = { grepl_repls: U.psmap (repl_state unit); grepl_stdin: stream_reader } - - -(*************************) -(* REPL tasks and states *) -(*************************) - -let t0 = Util.now () - -(** Create a timed_fname with a dummy modtime **) -let dummy_tf_of_fname fname = - { tf_fname = fname; - tf_modtime = t0 } - -let string_of_timed_fname { tf_fname = fname; tf_modtime = modtime } = - if modtime = t0 then Util.format1 "{ %s }" fname - else Util.format2 "{ %s; %s }" fname (string_of_time modtime) - -let string_of_repl_task = function - | LDInterleaved (intf, impl) -> - Util.format2 "LDInterleaved (%s, %s)" (string_of_timed_fname intf) (string_of_timed_fname impl) - | LDSingle intf_or_impl -> - Util.format1 "LDSingle %s" (string_of_timed_fname intf_or_impl) - | LDInterfaceOfCurrentFile intf -> - Util.format1 "LDInterfaceOfCurrentFile %s" (string_of_timed_fname intf) - | PushFragment (Inl frag) -> - Util.format1 "PushFragment { code = %s }" frag.frag_text - | PushFragment (Inr d) -> - Util.format1 "PushFragment { decl = %s }" (FStar.Parser.AST.decl_to_string d) - | Noop -> "Noop {}" - -module BU = FStar.Compiler.Util - -let string_of_repl_stack_entry - : repl_stack_entry_t _ -> string - = fun ((depth, i), (task, state)) -> - BU.format "{depth=%s; task=%s}" - [string_of_int i; - string_of_repl_task task] - - -let string_of_repl_stack s = - String.concat ";\n\t\t" - (List.map string_of_repl_stack_entry s) - -let repl_state_to_string (r:repl_state 'a) - : string - = BU.format - "{\n\t\ - repl_line=%s;\n\t\ - repl_column=%s;\n\t\ - repl_fname=%s;\n\t\ - repl_cur_mod=%s;\n\t\ - repl_deps_stack={%s}\n\ - }" - [string_of_int r.repl_line; - string_of_int r.repl_column; - r.repl_fname; - (match r.repl_curmod with - | None -> "None" - | Some m -> Ident.string_of_lid m.name); - string_of_repl_stack r.repl_deps_stack] -*) \ No newline at end of file From fc2174426a997227de08ea2c0b7b81c5003d9fbd Mon Sep 17 00:00:00 2001 From: Nikhil Swamy Date: Wed, 22 Mar 2023 16:09:10 +0000 Subject: [PATCH 28/48] snap --- ocaml/fstar-lib/generated/FStar_Interactive_Ide_Types.ml | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/ocaml/fstar-lib/generated/FStar_Interactive_Ide_Types.ml b/ocaml/fstar-lib/generated/FStar_Interactive_Ide_Types.ml index 0a2ba4e0a3b..295c26eb7e2 100644 --- a/ocaml/fstar-lib/generated/FStar_Interactive_Ide_Types.ml +++ b/ocaml/fstar-lib/generated/FStar_Interactive_Ide_Types.ml @@ -437,7 +437,7 @@ let (repl_state_to_string : repl_state -> Prims.string) = uu___3 :: uu___4 in uu___1 :: uu___2 in FStar_Compiler_Util.format - "{\n\trepl_line=%s;\n\trepl_column=%s;\n\trepl_fname=%s;\n\trepl_cur_mod=%s;\n\t\\ \r\n repl_deps_stack={%s}\n}" + "{\n\trepl_line=%s;\n\trepl_column=%s;\n\trepl_fname=%s;\n\trepl_cur_mod=%s;\n\t\\ \n repl_deps_stack={%s}\n}" uu___ let (push_query_to_string : push_query -> Prims.string) = fun pq -> From a1db6a28b74a97b7489a676fc9894bfc3e8cf206 Mon Sep 17 00:00:00 2001 From: Nikhil Swamy Date: Wed, 22 Mar 2023 16:54:19 +0000 Subject: [PATCH 29/48] Add a few code comments --- ocaml/fstar-lib/FStar_Parser_ParseIt.ml | 26 ++++++++++++++---- .../FStar_Interactive_Incremental.ml | 22 --------------- src/fstar/FStar.Interactive.Ide.Types.fst | 9 +++++++ src/fstar/FStar.Interactive.Incremental.fst | 27 ++++++++++--------- src/fstar/FStar.Interactive.Incremental.fsti | 4 ++- src/fstar/FStar.Universal.fst | 1 + src/parser/FStar.Parser.AST.Comparison.fsti | 2 ++ 7 files changed, 51 insertions(+), 40 deletions(-) diff --git a/ocaml/fstar-lib/FStar_Parser_ParseIt.ml b/ocaml/fstar-lib/FStar_Parser_ParseIt.ml index 0f9f543598d..0b65ff74ed6 100644 --- a/ocaml/fstar-lib/FStar_Parser_ParseIt.ml +++ b/ocaml/fstar-lib/FStar_Parser_ParseIt.ml @@ -146,15 +146,24 @@ let parse fn = let parse_one_decl = MenhirLib.Convert.Simplified.traditional2revised FStar_Parser_Parse.oneDeclOrEOF in let contents_at = let lines = U.splitlines contents in + (* Find the raw content of the input from the line of the start_pos to the end_pos. + This is used by Interactive.Incremental to record exactly the raw content of the + fragment that was checked *) fun (start_pos:Lexing.position) (end_pos:Lexing.position) -> + (* discard all lines until the start line *) let suffix = FStar_Compiler_Util.nth_tail (Z.of_int (if start_pos.pos_lnum > 0 then start_pos.pos_lnum - 1 else 0)) lines in + (* Take all the lines between the start and end lines *) let text, rest = FStar_Compiler_Util.first_N (Z.of_int (end_pos.pos_lnum - start_pos.pos_lnum)) suffix in let text = + (* For the last line itself, take the prefix of it up to the character of the end_pos *) match rest with | last::_ -> let col = end_pos.pos_cnum - end_pos.pos_bol in if col > 0 then ( + (* Don't index directly into the string, since this is a UTF-8 string. + Convert first to a list of charaters, index into that, and then convert + back to a string *) let chars = FStar_String.list_of_string last in if col <= List.length chars then ( @@ -175,10 +184,11 @@ let parse fn = in let open FStar_Pervasives in let rec parse decls = - let _, _, r = err_of_parse_error () in let start_pos = current_pos lexbuf in let d = try + (* Reset the gensym between decls, to ensure determinism, + otherwise, every _ is parsed as different name *) FStar_Ident.reset_gensym(); Inl (parse_one_decl lexer) with @@ -191,6 +201,13 @@ let parse fn = match d with | Inl None -> List.rev decls, None | Inl (Some d) -> + (* The parser may advance the lexer beyond the decls last token. + E.g., in `let f x = 0 let g = 1`, we will have parsed the decl for `f` + but the lexer will have advanced to `let ^ g ...` since the + parser will have looked ahead. + Rollback the lexer one token for declarations whose syntax + requires such lookahead to complete a production. + *) if not (FStar_Parser_AST.decl_syntax_is_delimited d) then rollback lexbuf; let end_pos = current_pos lexbuf in @@ -213,10 +230,9 @@ let parse fn = match fn with | Filename _ | Toplevel _ -> begin - let fileOrFragment = - if FStar_Options.debug_any() - then parse_incremental_fragment () - else MenhirLib.Convert.Simplified.traditional2revised FStar_Parser_Parse.inputFragment lexer in + let fileOrFragment = + MenhirLib.Convert.Simplified.traditional2revised FStar_Parser_Parse.inputFragment lexer + in let frags = match fileOrFragment with | FStar_Pervasives.Inl modul -> if has_extension filename interface_extensions diff --git a/ocaml/fstar-lib/generated/FStar_Interactive_Incremental.ml b/ocaml/fstar-lib/generated/FStar_Interactive_Incremental.ml index debb99a2338..d2068acf6d9 100644 --- a/ocaml/fstar-lib/generated/FStar_Interactive_Incremental.ml +++ b/ocaml/fstar-lib/generated/FStar_Interactive_Incremental.ml @@ -140,28 +140,6 @@ let (pop_entries : FStar_Interactive_Ide_Types.repl_stack_entry_t Prims.list -> FStar_Interactive_Ide_Types.query Prims.list qst) = fun e -> map (fun uu___ -> as_query FStar_Interactive_Ide_Types.Pop) e -let (response_success : - FStar_Parser_AST.decl -> FStar_Compiler_Util.json qst) = - fun d -> - op_let_Bang get_qid - (fun uu___ -> - match uu___ with - | (q, uu___1) -> - let contents = - let uu___2 = - let uu___3 = - let uu___4 = - FStar_Compiler_Range.json_of_def_range - d.FStar_Parser_AST.drange in - ("ranges", uu___4) in - [uu___3] in - FStar_Compiler_Util.JsonAssoc uu___2 in - return - (FStar_Compiler_Util.JsonAssoc - [("kind", (FStar_Compiler_Util.JsonStr "response")); - ("query-id", (FStar_Compiler_Util.JsonStr q)); - ("status", (FStar_Compiler_Util.JsonStr "success")); - ("contents", contents)])) let repl_task : 'uuuuu 'uuuuu1 'uuuuu2 . ('uuuuu * ('uuuuu1 * 'uuuuu2)) -> 'uuuuu1 = fun uu___ -> match uu___ with | (uu___1, (p, uu___2)) -> p diff --git a/src/fstar/FStar.Interactive.Ide.Types.fst b/src/fstar/FStar.Interactive.Ide.Types.fst index 50a6969c69d..7bf22b125af 100755 --- a/src/fstar/FStar.Interactive.Ide.Types.fst +++ b/src/fstar/FStar.Interactive.Ide.Types.fst @@ -119,9 +119,18 @@ type query' = | Search of string | GenericError of string | ProtocolViolation of string +// FullBuffer: To check the full contents of a document. +// FStar.Interactive.Incremental parses it into chunks and turns this into several Push queries | FullBuffer of string & full_buffer_request_kind +// Callback: This is an internal query, it cannot be raised by a client. +// It is useful to inject operations into the query stream. +// e.g., Incremental uses it print progress messages to the client in between +// processing a stream of Pushes that result from a chunking a FullBuffer | Callback of callback_t +// Format: pretty-print the F* code in the selection | Format of string +// Cancel: Cancel any remaining pushes that are at or beyond the provided position. +// Cancel all requests if the position is None | Cancel of option position and query = { qq: query'; qid: string } and callback_t = repl_state -> (query_status * list json) * either repl_state int diff --git a/src/fstar/FStar.Interactive.Incremental.fst b/src/fstar/FStar.Interactive.Incremental.fst index 5ee5775c6e8..0a967aa2156 100755 --- a/src/fstar/FStar.Interactive.Incremental.fst +++ b/src/fstar/FStar.Interactive.Incremental.fst @@ -86,6 +86,9 @@ let as_query (q:query') qid=qid_prefix ^ "." ^ string_of_int i } +(* Push a decl for checking, and before it runs, + print a progress message "fragment-started" + for the decl that is about to run *) let push_decl (push_kind:push_kind) (write_full_buffer_fragment_progress: fragment_progress -> unit) (ds:decl * code_fragment) @@ -118,22 +121,17 @@ let pop_entries (e:list repl_stack_entry_t) : qst (list query) = map (fun _ -> as_query Pop) e -let response_success (d:decl) - : qst json - = let! (q, _) = get_qid in - let contents = JsonAssoc (["ranges", json_of_def_range d.drange]) in - return (JsonAssoc [("kind", JsonStr "response"); - ("query-id", JsonStr q); - ("status", JsonStr "success"); - ("contents", contents)]) - let repl_task (_, (p, _)) = p +(* Find a prefix of the repl stack that matche a prefix of the decls ds, + pop the rest of the stack + and push the remaining suffix of decls +*) let inspect_repl_stack (s:repl_stack_t) (ds:list (decl & code_fragment)) (push_kind : push_kind) (write_full_buffer_fragment_progress: fragment_progress -> unit) - : qst (list query) // & list json) + : qst (list query) = let entries = List.rev s in let push_decls = push_decls push_kind write_full_buffer_fragment_progress in match BU.prefix_until @@ -182,7 +180,9 @@ let inspect_repl_stack (s:repl_stack_t) return pops in matching_prefix entries ds - + +(* A reload_deps request just pops away the entire stack of PushFragments. + We also push on just the `module A` declaration after popping. That's done below. *) let reload_deps repl_stack = let pop_until_deps entries : qst (list query) @@ -199,6 +199,7 @@ let reload_deps repl_stack = in pop_until_deps repl_stack +(* A utility to parse a chunk, used both in full_buffer and formatting *) let parse_code (code:string) = P.parse (Incremental { frag_fname = Range.file_of_range initial_range; @@ -207,7 +208,7 @@ let parse_code (code:string) = frag_col = Range.col_of_pos (Range.start_of_range initial_range); }) - +(* Format FStar.Errors.error into a JSON error message *) let syntax_issue (raw_error, msg, range) = let _, _, num = FStar.Errors.lookup raw_error in let issue = { @@ -219,6 +220,7 @@ let syntax_issue (raw_error, msg, range) = } in issue +(* See comment in the interface file *) let run_full_buffer (st:repl_state) (qid:string) (code:string) @@ -284,6 +286,7 @@ let run_full_buffer (st:repl_state) in qs +(* See comment in interface file *) let format_code (st:repl_state) (code:string) = let parse_result = parse_code code in match parse_result with diff --git a/src/fstar/FStar.Interactive.Incremental.fsti b/src/fstar/FStar.Interactive.Incremental.fsti index f45f49a9d51..35494cdaec7 100755 --- a/src/fstar/FStar.Interactive.Incremental.fsti +++ b/src/fstar/FStar.Interactive.Incremental.fsti @@ -24,6 +24,7 @@ open FStar.Parser.AST open FStar.Errors open FStar.Interactive.Ide.Types +(* Various kinds of progress messages to print back to the client *) type fragment_progress = | FragmentStarted of decl | FragmentSuccess of (decl & FStar.Parser.ParseIt.code_fragment & push_kind) @@ -47,7 +48,8 @@ val run_full_buffer (st:repl_state) (write_full_buffer_fragment_progress: fragment_progress -> unit) : list query +(* Pretty-print the code for reformatting, or return a syntax error *) val format_code (st:repl_state) (code:string) : either string (list issue) - \ No newline at end of file + diff --git a/src/fstar/FStar.Universal.fst b/src/fstar/FStar.Universal.fst index 24d51ebdcf1..36e580819d2 100644 --- a/src/fstar/FStar.Universal.fst +++ b/src/fstar/FStar.Universal.fst @@ -222,6 +222,7 @@ let tc_one_fragment curmod (env:TcEnv.env_t) frag = in match frag with | Inr d -> ( + //We already have a parsed decl, usually from FStar.Interactive.Incremental match d.d with | FStar.Parser.AST.TopLevelModule lid -> check_module_name_declaration (FStar.Parser.AST.Module(lid, [d])) diff --git a/src/parser/FStar.Parser.AST.Comparison.fsti b/src/parser/FStar.Parser.AST.Comparison.fsti index e87dee1a846..71a4e2e91b6 100755 --- a/src/parser/FStar.Parser.AST.Comparison.fsti +++ b/src/parser/FStar.Parser.AST.Comparison.fsti @@ -21,4 +21,6 @@ open FStar.Compiler.Effect open FStar.Compiler.List open FStar.Parser.AST +(* Check if two decls are equal, ignoring range metadata. + Used in FStar.Interactive.Incremental *) val eq_decl (d1 d2:decl) : bool From cf23b6648096c0a558eae13ca636650ccbe8b383 Mon Sep 17 00:00:00 2001 From: Nikhil Swamy Date: Wed, 22 Mar 2023 12:47:32 -0700 Subject: [PATCH 30/48] echo back the search term in auto-completion requests, for correlation in fstar-vscode-assistant --- ocaml/fstar-lib/generated/FStar_Interactive_Ide.ml | 12 +++++++++++- src/fstar/FStar.Interactive.Ide.fst | 7 ++++++- 2 files changed, 17 insertions(+), 2 deletions(-) diff --git a/ocaml/fstar-lib/generated/FStar_Interactive_Ide.ml b/ocaml/fstar-lib/generated/FStar_Interactive_Ide.ml index 7456aecb260..ae4dec14410 100644 --- a/ocaml/fstar-lib/generated/FStar_Interactive_Ide.ml +++ b/ocaml/fstar-lib/generated/FStar_Interactive_Ide.ml @@ -1941,9 +1941,19 @@ let run_code_autocomplete : fun st -> fun search_term -> let result = FStar_Interactive_QueryHelper.ck_completion st search_term in + let result_correlator = + { + FStar_Interactive_CompletionTable.completion_match_length = + Prims.int_zero; + FStar_Interactive_CompletionTable.completion_candidate = + search_term; + FStar_Interactive_CompletionTable.completion_annotation = + "" + } in let js = FStar_Compiler_List.map - FStar_Interactive_CompletionTable.json_of_completion_result result in + FStar_Interactive_CompletionTable.json_of_completion_result + (FStar_Compiler_List.op_At result [result_correlator]) in ((FStar_Interactive_Ide_Types.QueryOK, (FStar_Compiler_Util.JsonList js)), (FStar_Pervasives.Inl st)) let run_module_autocomplete : diff --git a/src/fstar/FStar.Interactive.Ide.fst b/src/fstar/FStar.Interactive.Ide.fst index 0ade7cb0da4..a59a3dac3bd 100644 --- a/src/fstar/FStar.Interactive.Ide.fst +++ b/src/fstar/FStar.Interactive.Ide.fst @@ -738,7 +738,12 @@ let run_lookup st symbol context pos_opt requested_info symrange = let run_code_autocomplete st search_term = let result = QH.ck_completion st search_term in - let js = List.map CTable.json_of_completion_result result in + let result_correlator : CTable.completion_result = { + completion_match_length = 0; + completion_annotation = ""; + completion_candidate = search_term + } in + let js = List.map CTable.json_of_completion_result (result@[result_correlator]) in ((QueryOK, JsonList js), Inl st) let run_module_autocomplete st search_term modules namespaces = From f0cb3ed1543113cd6ae1f148cea4cd98e25fa6eb Mon Sep 17 00:00:00 2001 From: Nikhil Swamy Date: Wed, 22 Mar 2023 20:50:19 +0000 Subject: [PATCH 31/48] update emacs test cases to account for echoed search term in auto-complete requestss --- .../generated/FStar_Interactive_Ide.ml | 26 +++++++++++-------- src/fstar/FStar.Interactive.Ide.fst | 18 ++++++++----- .../backtracking.refinements.out.expected | 4 +-- .../minimal.autocomplete-global.out.expected | 2 +- .../minimal.autocomplete-local.out.expected | 2 +- 5 files changed, 31 insertions(+), 21 deletions(-) diff --git a/ocaml/fstar-lib/generated/FStar_Interactive_Ide.ml b/ocaml/fstar-lib/generated/FStar_Interactive_Ide.ml index ae4dec14410..7dc38acef1b 100644 --- a/ocaml/fstar-lib/generated/FStar_Interactive_Ide.ml +++ b/ocaml/fstar-lib/generated/FStar_Interactive_Ide.ml @@ -1941,19 +1941,23 @@ let run_code_autocomplete : fun st -> fun search_term -> let result = FStar_Interactive_QueryHelper.ck_completion st search_term in - let result_correlator = - { - FStar_Interactive_CompletionTable.completion_match_length = - Prims.int_zero; - FStar_Interactive_CompletionTable.completion_candidate = - search_term; - FStar_Interactive_CompletionTable.completion_annotation = - "" - } in + let results = + match result with + | [] -> result + | uu___ -> + let result_correlator = + { + FStar_Interactive_CompletionTable.completion_match_length = + Prims.int_zero; + FStar_Interactive_CompletionTable.completion_candidate = + search_term; + FStar_Interactive_CompletionTable.completion_annotation = + "" + } in + FStar_Compiler_List.op_At result [result_correlator] in let js = FStar_Compiler_List.map - FStar_Interactive_CompletionTable.json_of_completion_result - (FStar_Compiler_List.op_At result [result_correlator]) in + FStar_Interactive_CompletionTable.json_of_completion_result results in ((FStar_Interactive_Ide_Types.QueryOK, (FStar_Compiler_Util.JsonList js)), (FStar_Pervasives.Inl st)) let run_module_autocomplete : diff --git a/src/fstar/FStar.Interactive.Ide.fst b/src/fstar/FStar.Interactive.Ide.fst index a59a3dac3bd..65e0f0f9509 100644 --- a/src/fstar/FStar.Interactive.Ide.fst +++ b/src/fstar/FStar.Interactive.Ide.fst @@ -738,12 +738,18 @@ let run_lookup st symbol context pos_opt requested_info symrange = let run_code_autocomplete st search_term = let result = QH.ck_completion st search_term in - let result_correlator : CTable.completion_result = { - completion_match_length = 0; - completion_annotation = ""; - completion_candidate = search_term - } in - let js = List.map CTable.json_of_completion_result (result@[result_correlator]) in + let results = + match result with + | [] -> result + | _ -> + let result_correlator : CTable.completion_result = { + completion_match_length = 0; + completion_annotation = ""; + completion_candidate = search_term + } in + result@[result_correlator] + in + let js = List.map CTable.json_of_completion_result results in ((QueryOK, JsonList js), Inl st) let run_module_autocomplete st search_term modules namespaces = diff --git a/tests/ide/emacs/backtracking.refinements.out.expected b/tests/ide/emacs/backtracking.refinements.out.expected index e05397cd0b7..a256e3daf2f 100644 --- a/tests/ide/emacs/backtracking.refinements.out.expected +++ b/tests/ide/emacs/backtracking.refinements.out.expected @@ -6,7 +6,7 @@ {"kind": "response", "query-id": "5", "response": [], "status": "success"} {"kind": "response", "query-id": "6", "response": [], "status": "success"} {"kind": "response", "query-id": "7", "response": null, "status": "success"} -{"kind": "response", "query-id": "8", "response": [[3, "Prims", "nat"]], "status": "success"} +{"kind": "response", "query-id": "8", "response": [[3, "Prims", "nat"], [0, "", "nat"]], "status": "success"} {"kind": "response", "query-id": "9", "response": [], "status": "success"} {"kind": "response", "query-id": "10", "response": [], "status": "success"} {"kind": "response", "query-id": "11", "response": null, "status": "success"} @@ -35,7 +35,7 @@ {"kind": "response", "query-id": "34", "response": [], "status": "success"} {"kind": "response", "query-id": "35", "response": [{"level": "error", "message": "Identifier not found: [b]", "number": 72, "ranges": [{"beg": [5, 7], "end": [5, 8], "fname": ""}]}], "status": "success"} {"kind": "response", "query-id": "36", "response": [{"level": "error", "message": "Identifier not found: [b]", "number": 72, "ranges": [{"beg": [5, 7], "end": [5, 8], "fname": ""}]}], "status": "success"} -{"kind": "response", "query-id": "37", "response": [[3, "Prims", "nat"]], "status": "success"} +{"kind": "response", "query-id": "37", "response": [[3, "Prims", "nat"], [0, "", "nat"]], "status": "success"} {"kind": "response", "query-id": "38", "response": [], "status": "success"} {"kind": "response", "query-id": "39", "response": [{"level": "error", "message": "Syntax error", "number": 168, "ranges": [{"beg": [5, 15], "end": [5, 15], "fname": ""}]}], "status": "success"} {"kind": "response", "query-id": "40", "response": [{"level": "error", "message": "Syntax error", "number": 168, "ranges": [{"beg": [5, 19], "end": [5, 19], "fname": ""}]}], "status": "success"} diff --git a/tests/ide/emacs/minimal.autocomplete-global.out.expected b/tests/ide/emacs/minimal.autocomplete-global.out.expected index 09f947962d6..704d62e25d8 100644 --- a/tests/ide/emacs/minimal.autocomplete-global.out.expected +++ b/tests/ide/emacs/minimal.autocomplete-global.out.expected @@ -1,3 +1,3 @@ {"kind": "protocol-info", "rest": "[...]"} {"kind": "response", "query-id": "1", "response": [], "status": "success"} -{"kind": "response", "query-id": "2", "response": [[14, "", "Prims.Mkdtuple2"]], "status": "success"} +{"kind": "response", "query-id": "2", "response": [[14, "", "Prims.Mkdtuple2"], [0, "", "Pr.Mkdtuple"]], "status": "success"} diff --git a/tests/ide/emacs/minimal.autocomplete-local.out.expected b/tests/ide/emacs/minimal.autocomplete-local.out.expected index d2a86847e87..02f5d2fc901 100644 --- a/tests/ide/emacs/minimal.autocomplete-local.out.expected +++ b/tests/ide/emacs/minimal.autocomplete-local.out.expected @@ -1,4 +1,4 @@ {"kind": "protocol-info", "rest": "[...]"} {"kind": "response", "query-id": "1", "response": [], "status": "success"} -{"kind": "response", "query-id": "2", "response": [[3, "", "hello"]], "status": "success"} +{"kind": "response", "query-id": "2", "response": [[3, "", "hello"], [0, "", "hel"]], "status": "success"} {"kind": "response", "query-id": "3", "response": {"defined-at": {"beg": [2, 4], "end": [2, 9], "fname": ""}, "definition": null, "documentation": null, "kind": "symbol", "name": "Minimal.hello", "type": "string"}, "status": "success"} From 34addfcff56a2e1ab58ac88c685f72caa893d7fc Mon Sep 17 00:00:00 2001 From: Nikhil Swamy Date: Wed, 22 Mar 2023 18:06:15 -0700 Subject: [PATCH 32/48] a missing type promotion in the the id_info_map: perhaps a cause of a long-standing bug --- src/typechecker/FStar.TypeChecker.Common.fst | 3 ++- 1 file changed, 2 insertions(+), 1 deletion(-) diff --git a/src/typechecker/FStar.TypeChecker.Common.fst b/src/typechecker/FStar.TypeChecker.Common.fst index 109923dd1d8..272338470f5 100644 --- a/src/typechecker/FStar.TypeChecker.Common.fst +++ b/src/typechecker/FStar.TypeChecker.Common.fst @@ -109,7 +109,8 @@ let id_info__insert ty_map db info = let use_range = Range.set_def_range range (Range.use_range range) in let id_ty = match info.identifier with - | Inr _ -> info.identifier_ty + | Inr _ -> + ty_map info.identifier_ty | Inl x -> // BU.print1 "id_info__insert: %s\n" // (print_identifier_info info); From f5b496112d39d4fe50a40f3493e6b86dbe5591ec Mon Sep 17 00:00:00 2001 From: Nikhil Swamy Date: Wed, 22 Mar 2023 18:06:50 -0700 Subject: [PATCH 33/48] pushing the symbol lookups when checking full-buffer fragments to preload the extension symbol table --- ocaml/fstar-lib/FStar_List.ml | 1 + .../generated/FStar_Interactive_Ide.ml | 2 + .../generated/FStar_Interactive_Ide_Types.ml | 16 +- .../FStar_Interactive_Incremental.ml | 351 +++++++++++------- .../FStar_Interactive_QueryHelper.ml | 6 +- ocaml/fstar-lib/generated/FStar_Main.ml | 39 +- .../generated/FStar_Parser_AST_Comparison.ml | 249 ++++++++++++- .../generated/FStar_TypeChecker_Common.ml | 2 +- src/fstar/FStar.Interactive.Ide.Types.fst | 10 +- src/fstar/FStar.Interactive.Ide.fst | 3 +- src/fstar/FStar.Interactive.Incremental.fst | 56 ++- src/fstar/FStar.Interactive.QueryHelper.fst | 6 +- src/fstar/FStar.Main.fst | 8 +- src/parser/FStar.Parser.AST.Comparison.fst | 67 ++++ src/parser/FStar.Parser.AST.Comparison.fsti | 2 + 15 files changed, 640 insertions(+), 178 deletions(-) diff --git a/ocaml/fstar-lib/FStar_List.ml b/ocaml/fstar-lib/FStar_List.ml index 36e4a581930..d1b052be031 100644 --- a/ocaml/fstar-lib/FStar_List.ml +++ b/ocaml/fstar-lib/FStar_List.ml @@ -56,6 +56,7 @@ let tryFindT = tryFind let find = tryFind let tryPick f l = try f (BatList.find (fun x -> f x <> None) l) with | Not_found -> None let flatten = BatList.flatten +let concat = flatten let split = unzip let choose = BatList.filter_map let existsb f l = BatList.exists f l diff --git a/ocaml/fstar-lib/generated/FStar_Interactive_Ide.ml b/ocaml/fstar-lib/generated/FStar_Interactive_Ide.ml index ae4dec14410..76fe8b5448f 100644 --- a/ocaml/fstar-lib/generated/FStar_Interactive_Ide.ml +++ b/ocaml/fstar-lib/generated/FStar_Interactive_Ide.ml @@ -358,6 +358,8 @@ let (unpack_interactive_query : let parse_full_buffer_kind uu___1 = match uu___1 with | "full" -> FStar_Interactive_Ide_Types.Full + | "full-with-symbols" -> + FStar_Interactive_Ide_Types.FullBufferWithSymbols | "cache" -> FStar_Interactive_Ide_Types.Cache | "reload-deps" -> FStar_Interactive_Ide_Types.ReloadDeps | "verify-to-position" -> diff --git a/ocaml/fstar-lib/generated/FStar_Interactive_Ide_Types.ml b/ocaml/fstar-lib/generated/FStar_Interactive_Ide_Types.ml index 295c26eb7e2..2ac131e733b 100644 --- a/ocaml/fstar-lib/generated/FStar_Interactive_Ide_Types.ml +++ b/ocaml/fstar-lib/generated/FStar_Interactive_Ide_Types.ml @@ -149,12 +149,17 @@ let (uu___is_Noop : repl_task -> Prims.bool) = fun projectee -> match projectee with | Noop -> true | uu___ -> false type full_buffer_request_kind = | Full + | FullBufferWithSymbols | Cache | ReloadDeps | VerifyToPosition of position | LaxToPosition of position let (uu___is_Full : full_buffer_request_kind -> Prims.bool) = fun projectee -> match projectee with | Full -> true | uu___ -> false +let (uu___is_FullBufferWithSymbols : full_buffer_request_kind -> Prims.bool) + = + fun projectee -> + match projectee with | FullBufferWithSymbols -> true | uu___ -> false let (uu___is_Cache : full_buffer_request_kind -> Prims.bool) = fun projectee -> match projectee with | Cache -> true | uu___ -> false let (uu___is_ReloadDeps : full_buffer_request_kind -> Prims.bool) = @@ -478,7 +483,16 @@ let (query_to_string : query -> Prims.string) = Prims.op_Hat "(Push " uu___ | VfsAdd uu___ -> "VfsAdd" | AutoComplete uu___ -> "AutoComplete" - | Lookup uu___ -> "Lookup" + | Lookup (s, _lc, pos, features, _sr) -> + let uu___ = + match pos with + | FStar_Pervasives_Native.None -> "None" + | FStar_Pervasives_Native.Some (f, i, j) -> + let uu___1 = FStar_Compiler_Util.string_of_int i in + let uu___2 = FStar_Compiler_Util.string_of_int j in + FStar_Compiler_Util.format3 "(%s, %s, %s)" f uu___1 uu___2 in + FStar_Compiler_Util.format3 "(Lookup %s %s [%s])" s uu___ + (FStar_String.concat "; " features) | Compute uu___ -> "Compute" | Search uu___ -> "Search" | GenericError uu___ -> "GenericError" diff --git a/ocaml/fstar-lib/generated/FStar_Interactive_Incremental.ml b/ocaml/fstar-lib/generated/FStar_Interactive_Incremental.ml index d2068acf6d9..cbca0a2208c 100644 --- a/ocaml/fstar-lib/generated/FStar_Interactive_Incremental.ml +++ b/ocaml/fstar-lib/generated/FStar_Interactive_Incremental.ml @@ -83,59 +83,131 @@ let (as_query : FStar_Interactive_Ide_Types.qid = uu___2 } in return uu___1) +let (dump_symbols_of_term : + FStar_Parser_AST.term -> FStar_Interactive_Ide_Types.query Prims.list qst) + = + fun t -> + let lids = FStar_Parser_AST_Comparison.lidents_of_term t in + (let uu___1 = FStar_Parser_AST.term_to_string t in + let uu___2 = + let uu___3 = FStar_Compiler_List.map FStar_Ident.string_of_lid lids in + FStar_String.concat ", " uu___3 in + FStar_Compiler_Util.print2 "Dumping symbols for %s { %s }\n" uu___1 + uu___2); + (let lookup_lid l = + let r = FStar_Ident.range_of_lid l in + let start_pos = FStar_Compiler_Range.start_of_range r in + let end_pos = FStar_Compiler_Range.end_of_range r in + let start_line = FStar_Compiler_Range.line_of_pos start_pos in + let start_col = FStar_Compiler_Range.col_of_pos start_pos in + let end_line = FStar_Compiler_Range.line_of_pos end_pos in + let end_col = FStar_Compiler_Range.col_of_pos end_pos in + let position = ("", start_line, start_col) in + let uu___1 = + let uu___2 = + let uu___3 = FStar_Ident.string_of_lid l in + (uu___3, FStar_Interactive_Ide_Types.LKCode, + (FStar_Pervasives_Native.Some position), + ["type"; "documentation"; "defined-at"], + (FStar_Pervasives_Native.Some + (FStar_Compiler_Util.JsonAssoc + [("fname", (FStar_Compiler_Util.JsonStr "")); + ("beg", + (FStar_Compiler_Util.JsonList + [FStar_Compiler_Util.JsonInt start_line; + FStar_Compiler_Util.JsonInt start_col])); + ("end", + (FStar_Compiler_Util.JsonList + [FStar_Compiler_Util.JsonInt end_line; + FStar_Compiler_Util.JsonInt end_col]))]))) in + FStar_Interactive_Ide_Types.Lookup uu___2 in + as_query uu___1 in + map lookup_lid lids) +let (dump_symbols : + FStar_Parser_AST.decl -> FStar_Interactive_Ide_Types.query Prims.list qst) + = + fun d -> + match d.FStar_Parser_AST.d with + | FStar_Parser_AST.TopLevelLet (uu___, defs) -> + let uu___1 = + map + (fun uu___2 -> + match uu___2 with | (uu___3, t) -> dump_symbols_of_term t) + defs in + op_let_Bang uu___1 + (fun queries -> return (FStar_Compiler_List.concat queries)) + | uu___ -> return [] let (push_decl : FStar_Interactive_Ide_Types.push_kind -> - (fragment_progress -> unit) -> - (FStar_Parser_AST.decl * FStar_Parser_ParseIt.code_fragment) -> - FStar_Interactive_Ide_Types.query Prims.list qst) + Prims.bool -> + (fragment_progress -> unit) -> + (FStar_Parser_AST.decl * FStar_Parser_ParseIt.code_fragment) -> + FStar_Interactive_Ide_Types.query Prims.list qst) = fun push_kind -> - fun write_full_buffer_fragment_progress -> - fun ds -> - let uu___ = ds in - match uu___ with - | (d, s) -> - let pq = - let uu___1 = + fun with_symbols -> + fun write_full_buffer_fragment_progress -> + fun ds -> + let uu___ = ds in + match uu___ with + | (d, s) -> + let pq = + let uu___1 = + let uu___2 = + FStar_Compiler_Range.start_of_range + d.FStar_Parser_AST.drange in + FStar_Compiler_Range.line_of_pos uu___2 in let uu___2 = - FStar_Compiler_Range.start_of_range - d.FStar_Parser_AST.drange in - FStar_Compiler_Range.line_of_pos uu___2 in - let uu___2 = - let uu___3 = - FStar_Compiler_Range.start_of_range - d.FStar_Parser_AST.drange in - FStar_Compiler_Range.col_of_pos uu___3 in - { - FStar_Interactive_Ide_Types.push_kind = push_kind; - FStar_Interactive_Ide_Types.push_line = uu___1; - FStar_Interactive_Ide_Types.push_column = uu___2; - FStar_Interactive_Ide_Types.push_peek_only = false; - FStar_Interactive_Ide_Types.push_code_or_decl = - (FStar_Pervasives.Inr ds) - } in - let progress st = - write_full_buffer_fragment_progress (FragmentStarted d); - ((FStar_Interactive_Ide_Types.QueryOK, []), - (FStar_Pervasives.Inl st)) in - let uu___1 = - as_query (FStar_Interactive_Ide_Types.Callback progress) in - op_let_Bang uu___1 - (fun cb -> - let uu___2 = as_query (FStar_Interactive_Ide_Types.Push pq) in - op_let_Bang uu___2 (fun push -> return [cb; push])) + let uu___3 = + FStar_Compiler_Range.start_of_range + d.FStar_Parser_AST.drange in + FStar_Compiler_Range.col_of_pos uu___3 in + { + FStar_Interactive_Ide_Types.push_kind = push_kind; + FStar_Interactive_Ide_Types.push_line = uu___1; + FStar_Interactive_Ide_Types.push_column = uu___2; + FStar_Interactive_Ide_Types.push_peek_only = false; + FStar_Interactive_Ide_Types.push_code_or_decl = + (FStar_Pervasives.Inr ds) + } in + let progress st = + write_full_buffer_fragment_progress (FragmentStarted d); + ((FStar_Interactive_Ide_Types.QueryOK, []), + (FStar_Pervasives.Inl st)) in + let uu___1 = + as_query (FStar_Interactive_Ide_Types.Callback progress) in + op_let_Bang uu___1 + (fun cb -> + let uu___2 = + as_query (FStar_Interactive_Ide_Types.Push pq) in + op_let_Bang uu___2 + (fun push -> + if with_symbols + then + let uu___3 = dump_symbols d in + op_let_Bang uu___3 + (fun lookups -> + return + (FStar_Compiler_List.op_At [cb; push] + lookups)) + else return [cb; push])) let (push_decls : FStar_Interactive_Ide_Types.push_kind -> - (fragment_progress -> unit) -> - (FStar_Parser_AST.decl * FStar_Parser_ParseIt.code_fragment) Prims.list - -> FStar_Interactive_Ide_Types.query Prims.list qst) + Prims.bool -> + (fragment_progress -> unit) -> + (FStar_Parser_AST.decl * FStar_Parser_ParseIt.code_fragment) + Prims.list -> FStar_Interactive_Ide_Types.query Prims.list qst) = fun push_kind -> - fun write_full_buffer_fragment_progress -> - fun ds -> - let uu___ = - map (push_decl push_kind write_full_buffer_fragment_progress) ds in - op_let_Bang uu___ (fun qs -> return (FStar_Compiler_List.flatten qs)) + fun with_symbols -> + fun write_full_buffer_fragment_progress -> + fun ds -> + let uu___ = + map + (push_decl push_kind with_symbols + write_full_buffer_fragment_progress) ds in + op_let_Bang uu___ + (fun qs -> return (FStar_Compiler_List.flatten qs)) let (pop_entries : FStar_Interactive_Ide_Types.repl_stack_entry_t Prims.list -> FStar_Interactive_Ide_Types.query Prims.list qst) @@ -148,79 +220,82 @@ let (inspect_repl_stack : (FStar_Parser_AST.decl * FStar_Parser_ParseIt.code_fragment) Prims.list -> FStar_Interactive_Ide_Types.push_kind -> - (fragment_progress -> unit) -> - FStar_Interactive_Ide_Types.query Prims.list qst) + Prims.bool -> + (fragment_progress -> unit) -> + FStar_Interactive_Ide_Types.query Prims.list qst) = fun s -> fun ds -> fun push_kind -> - fun write_full_buffer_fragment_progress -> - let entries = FStar_Compiler_List.rev s in - let push_decls1 = - push_decls push_kind write_full_buffer_fragment_progress in - let uu___ = - FStar_Compiler_Util.prefix_until - (fun uu___1 -> - match uu___1 with - | (uu___2, - (FStar_Interactive_Ide_Types.PushFragment uu___3, uu___4)) - -> true - | uu___2 -> false) entries in - match uu___ with - | FStar_Pervasives_Native.None -> - let uu___1 = push_decls1 ds in - op_let_Bang uu___1 (fun ds1 -> return ds1) - | FStar_Pervasives_Native.Some (prefix, first_push, rest) -> - let entries1 = first_push :: rest in - let repl_task1 uu___1 = - match uu___1 with | (uu___2, (p, uu___3)) -> p in - let rec matching_prefix entries2 ds1 = - match (entries2, ds1) with - | ([], []) -> return [] - | (e::entries3, d::ds2) -> - (match repl_task1 e with - | FStar_Interactive_Ide_Types.Noop -> - matching_prefix entries3 (d :: ds2) - | FStar_Interactive_Ide_Types.PushFragment - (FStar_Pervasives.Inl frag, uu___1) -> - let uu___2 = pop_entries (e :: entries3) in - op_let_Bang uu___2 - (fun pops -> - let uu___3 = push_decls1 (d :: ds2) in + fun with_symbols -> + fun write_full_buffer_fragment_progress -> + let entries = FStar_Compiler_List.rev s in + let push_decls1 = + push_decls push_kind with_symbols + write_full_buffer_fragment_progress in + let uu___ = + FStar_Compiler_Util.prefix_until + (fun uu___1 -> + match uu___1 with + | (uu___2, + (FStar_Interactive_Ide_Types.PushFragment uu___3, + uu___4)) -> true + | uu___2 -> false) entries in + match uu___ with + | FStar_Pervasives_Native.None -> + let uu___1 = push_decls1 ds in + op_let_Bang uu___1 (fun ds1 -> return ds1) + | FStar_Pervasives_Native.Some (prefix, first_push, rest) -> + let entries1 = first_push :: rest in + let repl_task1 uu___1 = + match uu___1 with | (uu___2, (p, uu___3)) -> p in + let rec matching_prefix entries2 ds1 = + match (entries2, ds1) with + | ([], []) -> return [] + | (e::entries3, d::ds2) -> + (match repl_task1 e with + | FStar_Interactive_Ide_Types.Noop -> + matching_prefix entries3 (d :: ds2) + | FStar_Interactive_Ide_Types.PushFragment + (FStar_Pervasives.Inl frag, uu___1) -> + let uu___2 = pop_entries (e :: entries3) in + op_let_Bang uu___2 + (fun pops -> + let uu___3 = push_decls1 (d :: ds2) in + op_let_Bang uu___3 + (fun pushes -> + return + (FStar_Compiler_List.op_At pops pushes))) + | FStar_Interactive_Ide_Types.PushFragment + (FStar_Pervasives.Inr d', pk) -> + let uu___1 = + FStar_Parser_AST_Comparison.eq_decl + (FStar_Pervasives_Native.fst d) d' in + if uu___1 + then + let uu___2 = d in + (match uu___2 with + | (d1, s1) -> + (write_full_buffer_fragment_progress + (FragmentSuccess (d1, s1, pk)); + matching_prefix entries3 ds2)) + else + (let uu___3 = pop_entries (e :: entries3) in op_let_Bang uu___3 - (fun pushes -> - return - (FStar_Compiler_List.op_At pops pushes))) - | FStar_Interactive_Ide_Types.PushFragment - (FStar_Pervasives.Inr d', pk) -> - let uu___1 = - FStar_Parser_AST_Comparison.eq_decl - (FStar_Pervasives_Native.fst d) d' in - if uu___1 - then - let uu___2 = d in - (match uu___2 with - | (d1, s1) -> - (write_full_buffer_fragment_progress - (FragmentSuccess (d1, s1, pk)); - matching_prefix entries3 ds2)) - else - (let uu___3 = pop_entries (e :: entries3) in - op_let_Bang uu___3 - (fun pops -> - let uu___4 = push_decls1 (d :: ds2) in - op_let_Bang uu___4 - (fun pushes -> - return - (FStar_Compiler_List.op_At pops - pushes))))) - | ([], ds2) -> - let uu___1 = push_decls1 ds2 in - op_let_Bang uu___1 (fun pushes -> return pushes) - | (es, []) -> - let uu___1 = pop_entries es in - op_let_Bang uu___1 (fun pops -> return pops) in - matching_prefix entries1 ds + (fun pops -> + let uu___4 = push_decls1 (d :: ds2) in + op_let_Bang uu___4 + (fun pushes -> + return + (FStar_Compiler_List.op_At pops + pushes))))) + | ([], ds2) -> + let uu___1 = push_decls1 ds2 in + op_let_Bang uu___1 (fun pushes -> return pushes) + | (es, []) -> + let uu___1 = pop_entries es in + op_let_Bang uu___1 (fun pops -> return pops) in + matching_prefix entries1 ds let reload_deps : 'uuuuu 'uuuuu1 . ('uuuuu * (FStar_Interactive_Ide_Types.repl_task * 'uuuuu1)) Prims.list @@ -315,33 +390,32 @@ let (run_full_buffer : match request_type with | FStar_Interactive_Ide_Types.VerifyToPosition (uu___, line, _col) -> - ((let uu___2 = FStar_Compiler_Util.string_of_int line in - FStar_Compiler_Util.print1 "Got to-position: %s" uu___2); - FStar_Compiler_List.filter - (fun uu___2 -> - match uu___2 with - | (d, uu___3) -> - let start = - FStar_Compiler_Range.start_of_range - d.FStar_Parser_AST.drange in - let start_line = - FStar_Compiler_Range.line_of_pos start in - start_line <= line) decls) + FStar_Compiler_List.filter + (fun uu___1 -> + match uu___1 with + | (d, uu___2) -> + let start = + FStar_Compiler_Range.start_of_range + d.FStar_Parser_AST.drange in + let start_line = + FStar_Compiler_Range.line_of_pos start in + start_line <= line) decls | FStar_Interactive_Ide_Types.LaxToPosition (uu___, line, _col) -> - ((let uu___2 = FStar_Compiler_Util.string_of_int line in - FStar_Compiler_Util.print1 "Got to-position: %s" uu___2); - FStar_Compiler_List.filter - (fun uu___2 -> - match uu___2 with - | (d, uu___3) -> - let start = - FStar_Compiler_Range.start_of_range - d.FStar_Parser_AST.drange in - let start_line = - FStar_Compiler_Range.line_of_pos start in - start_line <= line) decls) + FStar_Compiler_List.filter + (fun uu___1 -> + match uu___1 with + | (d, uu___2) -> + let start = + FStar_Compiler_Range.start_of_range + d.FStar_Parser_AST.drange in + let start_line = + FStar_Compiler_Range.line_of_pos start in + start_line <= line) decls | uu___ -> decls in + let with_symbols = + request_type = + FStar_Interactive_Ide_Types.FullBufferWithSymbols in let qs = match parse_result with | FStar_Parser_ParseIt.IncrementalFragment @@ -359,6 +433,7 @@ let (run_full_buffer : let uu___4 = push_decl FStar_Interactive_Ide_Types.FullCheck + with_symbols write_full_buffer_fragment_progress d in op_let_Bang uu___4 (fun push_mod -> @@ -372,6 +447,8 @@ let (run_full_buffer : match request_type with | FStar_Interactive_Ide_Types.LaxToPosition uu___2 -> FStar_Interactive_Ide_Types.LaxCheck + | FStar_Interactive_Ide_Types.FullBufferWithSymbols + -> FStar_Interactive_Ide_Types.LaxCheck | uu___2 -> FStar_Interactive_Ide_Types.FullCheck in let queries = let uu___2 = @@ -379,7 +456,7 @@ let (run_full_buffer : FStar_Compiler_Effect.op_Bang FStar_Interactive_PushHelper.repl_stack in inspect_repl_stack uu___3 decls1 push_kind - write_full_buffer_fragment_progress in + with_symbols write_full_buffer_fragment_progress in run_qst uu___2 qid1 in (if request_type = FStar_Interactive_Ide_Types.Full then log_syntax_issues err_opt diff --git a/ocaml/fstar-lib/generated/FStar_Interactive_QueryHelper.ml b/ocaml/fstar-lib/generated/FStar_Interactive_QueryHelper.ml index ffe40f25332..97ce93e4c96 100644 --- a/ocaml/fstar-lib/generated/FStar_Interactive_QueryHelper.ml +++ b/ocaml/fstar-lib/generated/FStar_Interactive_QueryHelper.ml @@ -44,7 +44,7 @@ let (term_to_string : fun tcenv -> fun t -> with_printed_effect_args - (fun uu___ -> FStar_TypeChecker_Normalize.term_to_string tcenv t) + (fun uu___ -> FStar_Syntax_Print.term_to_string t) let (sigelt_to_string : FStar_Syntax_Syntax.sigelt -> Prims.string) = fun se -> with_printed_effect_args @@ -108,6 +108,10 @@ let (symlookup : match name_or_lid with | FStar_Pervasives.Inl name1 -> name1 | FStar_Pervasives.Inr lid -> FStar_Ident.string_of_lid lid in + let str_of_opt uu___ = + match uu___ with + | FStar_Pervasives_Native.None -> "" + | FStar_Pervasives_Native.Some s -> s in let typ_str = if FStar_Compiler_List.mem "type" requested_info then diff --git a/ocaml/fstar-lib/generated/FStar_Main.ml b/ocaml/fstar-lib/generated/FStar_Main.ml index 8ad6b788389..dff0561b53c 100644 --- a/ocaml/fstar-lib/generated/FStar_Main.ml +++ b/ocaml/fstar-lib/generated/FStar_Main.ml @@ -320,26 +320,19 @@ let (handle_error : Prims.exn -> unit) = report_errors [] let main : 'uuuuu . unit -> 'uuuuu = fun uu___ -> - try - (fun uu___1 -> - match () with - | () -> - (setup_hooks (); - (let uu___3 = FStar_Compiler_Util.record_time go in - match uu___3 with - | (uu___4, time) -> - ((let uu___6 = FStar_Options.query_stats () in - if uu___6 - then - let uu___7 = FStar_Compiler_Util.string_of_int time in - let uu___8 = - let uu___9 = FStar_Getopt.cmdline () in - FStar_String.concat " " uu___9 in - FStar_Compiler_Util.print2_error - "TOTAL TIME %s ms: %s\n" uu___7 uu___8 - else ()); - cleanup (); - FStar_Compiler_Effect.exit Prims.int_zero)))) () - with - | uu___1 -> - (handle_error uu___1; FStar_Compiler_Effect.exit Prims.int_one) \ No newline at end of file + setup_hooks (); + (let uu___2 = FStar_Compiler_Util.record_time go in + match uu___2 with + | (uu___3, time) -> + ((let uu___5 = FStar_Options.query_stats () in + if uu___5 + then + let uu___6 = FStar_Compiler_Util.string_of_int time in + let uu___7 = + let uu___8 = FStar_Getopt.cmdline () in + FStar_String.concat " " uu___8 in + FStar_Compiler_Util.print2_error "TOTAL TIME %s ms: %s\n" uu___6 + uu___7 + else ()); + cleanup (); + FStar_Compiler_Effect.exit Prims.int_zero)) \ No newline at end of file diff --git a/ocaml/fstar-lib/generated/FStar_Parser_AST_Comparison.ml b/ocaml/fstar-lib/generated/FStar_Parser_AST_Comparison.ml index 1f13b917564..fb250b57829 100644 --- a/ocaml/fstar-lib/generated/FStar_Parser_AST_Comparison.ml +++ b/ocaml/fstar-lib/generated/FStar_Parser_AST_Comparison.ml @@ -604,4 +604,251 @@ and (eq_decl : FStar_Parser_AST.decl -> FStar_Parser_AST.decl -> Prims.bool) (eq_list eq_qualifier d1.FStar_Parser_AST.quals d2.FStar_Parser_AST.quals)) && - (eq_list eq_term d1.FStar_Parser_AST.attrs d2.FStar_Parser_AST.attrs) \ No newline at end of file + (eq_list eq_term d1.FStar_Parser_AST.attrs d2.FStar_Parser_AST.attrs) +let concat_map : + 'uuuuu 'uuuuu1 . + unit -> + ('uuuuu -> 'uuuuu1 Prims.list) -> + 'uuuuu Prims.list -> 'uuuuu1 Prims.list + = fun uu___ -> FStar_Compiler_List.collect +let opt_map : + 'uuuuu 'a . + ('a -> 'uuuuu Prims.list) -> + 'a FStar_Pervasives_Native.option -> 'uuuuu Prims.list + = + fun f -> + fun x -> + match x with + | FStar_Pervasives_Native.None -> [] + | FStar_Pervasives_Native.Some x1 -> f x1 +let rec (lidents_of_term : + FStar_Parser_AST.term -> FStar_Ident.lident Prims.list) = + fun t -> lidents_of_term' t.FStar_Parser_AST.tm +and (lidents_of_term' : + FStar_Parser_AST.term' -> FStar_Ident.lident Prims.list) = + fun t -> + match t with + | FStar_Parser_AST.Wild -> [] + | FStar_Parser_AST.Const uu___ -> [] + | FStar_Parser_AST.Op (s, ts) -> (concat_map ()) lidents_of_term ts + | FStar_Parser_AST.Tvar uu___ -> [] + | FStar_Parser_AST.Uvar uu___ -> [] + | FStar_Parser_AST.Var lid -> [lid] + | FStar_Parser_AST.Name lid -> [lid] + | FStar_Parser_AST.Projector (lid, uu___) -> [lid] + | FStar_Parser_AST.Construct (lid, ts) -> + let uu___ = + (concat_map ()) + (fun uu___1 -> + match uu___1 with | (t1, uu___2) -> lidents_of_term t1) ts in + lid :: uu___ + | FStar_Parser_AST.Abs (ps, t1) -> + let uu___ = (concat_map ()) lidents_of_pattern ps in + let uu___1 = lidents_of_term t1 in + FStar_Compiler_List.op_At uu___ uu___1 + | FStar_Parser_AST.App (t1, t2, uu___) -> + let uu___1 = lidents_of_term t1 in + let uu___2 = lidents_of_term t2 in + FStar_Compiler_List.op_At uu___1 uu___2 + | FStar_Parser_AST.Let (uu___, lbs, t1) -> + let uu___1 = + (concat_map ()) + (fun uu___2 -> + match uu___2 with + | (uu___3, (uu___4, t2)) -> lidents_of_term t2) lbs in + let uu___2 = lidents_of_term t1 in + FStar_Compiler_List.op_At uu___1 uu___2 + | FStar_Parser_AST.LetOperator (lbs, t1) -> + let uu___ = + (concat_map ()) + (fun uu___1 -> + match uu___1 with | (uu___2, uu___3, t2) -> lidents_of_term t2) + lbs in + let uu___1 = lidents_of_term t1 in + FStar_Compiler_List.op_At uu___ uu___1 + | FStar_Parser_AST.LetOpen (lid, t1) -> + let uu___ = lidents_of_term t1 in lid :: uu___ + | FStar_Parser_AST.LetOpenRecord (t1, t2, t3) -> + let uu___ = lidents_of_term t1 in + let uu___1 = + let uu___2 = lidents_of_term t2 in + let uu___3 = lidents_of_term t3 in + FStar_Compiler_List.op_At uu___2 uu___3 in + FStar_Compiler_List.op_At uu___ uu___1 + | FStar_Parser_AST.Seq (t1, t2) -> + let uu___ = lidents_of_term t1 in + let uu___1 = lidents_of_term t2 in + FStar_Compiler_List.op_At uu___ uu___1 + | FStar_Parser_AST.Bind (uu___, t1, t2) -> + let uu___1 = lidents_of_term t1 in + let uu___2 = lidents_of_term t2 in + FStar_Compiler_List.op_At uu___1 uu___2 + | FStar_Parser_AST.If (t1, uu___, uu___1, t2, t3) -> + let uu___2 = lidents_of_term t1 in + let uu___3 = + let uu___4 = lidents_of_term t2 in + let uu___5 = lidents_of_term t3 in + FStar_Compiler_List.op_At uu___4 uu___5 in + FStar_Compiler_List.op_At uu___2 uu___3 + | FStar_Parser_AST.Match (t1, uu___, uu___1, bs) -> + let uu___2 = lidents_of_term t1 in + let uu___3 = (concat_map ()) lidents_of_branch bs in + FStar_Compiler_List.op_At uu___2 uu___3 + | FStar_Parser_AST.TryWith (t1, bs) -> + let uu___ = lidents_of_term t1 in + let uu___1 = (concat_map ()) lidents_of_branch bs in + FStar_Compiler_List.op_At uu___ uu___1 + | FStar_Parser_AST.Ascribed (t1, t2, uu___, uu___1) -> + let uu___2 = lidents_of_term t1 in + let uu___3 = lidents_of_term t2 in + FStar_Compiler_List.op_At uu___2 uu___3 + | FStar_Parser_AST.Record (t1, ts) -> + let uu___ = + (concat_map ()) + (fun uu___1 -> + match uu___1 with | (uu___2, t2) -> lidents_of_term t2) ts in + let uu___1 = opt_map lidents_of_term t1 in + FStar_Compiler_List.op_At uu___ uu___1 + | FStar_Parser_AST.Project (t1, uu___) -> lidents_of_term t1 + | FStar_Parser_AST.Product (ts, t1) -> + let uu___ = (concat_map ()) lidents_of_binder ts in + let uu___1 = lidents_of_term t1 in + FStar_Compiler_List.op_At uu___ uu___1 + | FStar_Parser_AST.Sum (ts, t1) -> + let uu___ = + (concat_map ()) + (fun uu___1 -> + match uu___1 with + | FStar_Pervasives.Inl b -> lidents_of_binder b + | FStar_Pervasives.Inr t2 -> lidents_of_term t2) ts in + let uu___1 = lidents_of_term t1 in + FStar_Compiler_List.op_At uu___ uu___1 + | FStar_Parser_AST.QForall (bs, _pats, t1) -> lidents_of_term t1 + | FStar_Parser_AST.QExists (bs, _pats, t1) -> lidents_of_term t1 + | FStar_Parser_AST.Refine (b, t1) -> lidents_of_term t1 + | FStar_Parser_AST.NamedTyp (i, t1) -> lidents_of_term t1 + | FStar_Parser_AST.Paren t1 -> lidents_of_term t1 + | FStar_Parser_AST.Requires (t1, uu___) -> lidents_of_term t1 + | FStar_Parser_AST.Ensures (t1, uu___) -> lidents_of_term t1 + | FStar_Parser_AST.LexList ts -> (concat_map ()) lidents_of_term ts + | FStar_Parser_AST.WFOrder (t1, t2) -> + let uu___ = lidents_of_term t1 in + let uu___1 = lidents_of_term t2 in + FStar_Compiler_List.op_At uu___ uu___1 + | FStar_Parser_AST.Decreases (t1, uu___) -> lidents_of_term t1 + | FStar_Parser_AST.Labeled (t1, uu___, uu___1) -> lidents_of_term t1 + | FStar_Parser_AST.Discrim lid -> [lid] + | FStar_Parser_AST.Attributes ts -> (concat_map ()) lidents_of_term ts + | FStar_Parser_AST.Antiquote t1 -> lidents_of_term t1 + | FStar_Parser_AST.Quote (t1, uu___) -> lidents_of_term t1 + | FStar_Parser_AST.VQuote t1 -> lidents_of_term t1 + | FStar_Parser_AST.CalcProof (t1, t2, ts) -> + let uu___ = lidents_of_term t1 in + let uu___1 = + let uu___2 = lidents_of_term t2 in + let uu___3 = (concat_map ()) lidents_of_calc_step ts in + FStar_Compiler_List.op_At uu___2 uu___3 in + FStar_Compiler_List.op_At uu___ uu___1 + | FStar_Parser_AST.IntroForall (bs, t1, t2) -> + let uu___ = lidents_of_term t1 in + let uu___1 = lidents_of_term t2 in + FStar_Compiler_List.op_At uu___ uu___1 + | FStar_Parser_AST.IntroExists (bs, t1, ts, t2) -> + let uu___ = lidents_of_term t1 in + let uu___1 = + let uu___2 = (concat_map ()) lidents_of_term ts in + let uu___3 = lidents_of_term t2 in + FStar_Compiler_List.op_At uu___2 uu___3 in + FStar_Compiler_List.op_At uu___ uu___1 + | FStar_Parser_AST.IntroImplies (t1, t2, b, t3) -> + let uu___ = lidents_of_term t1 in + let uu___1 = + let uu___2 = lidents_of_term t2 in + let uu___3 = lidents_of_term t3 in + FStar_Compiler_List.op_At uu___2 uu___3 in + FStar_Compiler_List.op_At uu___ uu___1 + | FStar_Parser_AST.IntroOr (b, t1, t2, t3) -> + let uu___ = lidents_of_term t1 in + let uu___1 = + let uu___2 = lidents_of_term t2 in + let uu___3 = lidents_of_term t3 in + FStar_Compiler_List.op_At uu___2 uu___3 in + FStar_Compiler_List.op_At uu___ uu___1 + | FStar_Parser_AST.IntroAnd (t1, t2, t3, t4) -> + let uu___ = lidents_of_term t1 in + let uu___1 = + let uu___2 = lidents_of_term t2 in + let uu___3 = + let uu___4 = lidents_of_term t3 in + let uu___5 = lidents_of_term t4 in + FStar_Compiler_List.op_At uu___4 uu___5 in + FStar_Compiler_List.op_At uu___2 uu___3 in + FStar_Compiler_List.op_At uu___ uu___1 + | FStar_Parser_AST.ElimForall (bs, t1, ts) -> + let uu___ = lidents_of_term t1 in + let uu___1 = (concat_map ()) lidents_of_term ts in + FStar_Compiler_List.op_At uu___ uu___1 + | FStar_Parser_AST.ElimExists (bs, t1, t2, b, t3) -> + let uu___ = lidents_of_term t1 in + let uu___1 = + let uu___2 = lidents_of_term t2 in + let uu___3 = lidents_of_term t3 in + FStar_Compiler_List.op_At uu___2 uu___3 in + FStar_Compiler_List.op_At uu___ uu___1 + | FStar_Parser_AST.ElimImplies (t1, t2, t3) -> + let uu___ = lidents_of_term t1 in + let uu___1 = + let uu___2 = lidents_of_term t2 in + let uu___3 = lidents_of_term t3 in + FStar_Compiler_List.op_At uu___2 uu___3 in + FStar_Compiler_List.op_At uu___ uu___1 + | FStar_Parser_AST.ElimOr (t1, t2, t3, b1, t4, b2, t5) -> + let uu___ = lidents_of_term t1 in + let uu___1 = + let uu___2 = lidents_of_term t2 in + let uu___3 = + let uu___4 = lidents_of_term t3 in + let uu___5 = + let uu___6 = lidents_of_term t4 in + let uu___7 = lidents_of_term t5 in + FStar_Compiler_List.op_At uu___6 uu___7 in + FStar_Compiler_List.op_At uu___4 uu___5 in + FStar_Compiler_List.op_At uu___2 uu___3 in + FStar_Compiler_List.op_At uu___ uu___1 + | FStar_Parser_AST.ElimAnd (t1, t2, t3, b1, b2, t4) -> + let uu___ = lidents_of_term t1 in + let uu___1 = + let uu___2 = lidents_of_term t2 in + let uu___3 = + let uu___4 = lidents_of_term t3 in + let uu___5 = lidents_of_term t4 in + FStar_Compiler_List.op_At uu___4 uu___5 in + FStar_Compiler_List.op_At uu___2 uu___3 in + FStar_Compiler_List.op_At uu___ uu___1 +and (lidents_of_branch : + (FStar_Parser_AST.pattern * FStar_Parser_AST.term + FStar_Pervasives_Native.option * FStar_Parser_AST.term) -> + FStar_Ident.lident Prims.list) + = + fun uu___ -> + match uu___ with + | (p, uu___1, t) -> + let uu___2 = lidents_of_pattern p in + let uu___3 = lidents_of_term t in + FStar_Compiler_List.op_At uu___2 uu___3 +and (lidents_of_calc_step : + FStar_Parser_AST.calc_step -> FStar_Ident.lident Prims.list) = + fun uu___ -> + match uu___ with + | FStar_Parser_AST.CalcStep (t1, t2, t3) -> + let uu___1 = lidents_of_term t1 in + let uu___2 = + let uu___3 = lidents_of_term t2 in + let uu___4 = lidents_of_term t3 in + FStar_Compiler_List.op_At uu___3 uu___4 in + FStar_Compiler_List.op_At uu___1 uu___2 +and (lidents_of_pattern : + FStar_Parser_AST.pattern -> FStar_Ident.lident Prims.list) = + fun uu___ -> [] +and (lidents_of_binder : + FStar_Parser_AST.binder -> FStar_Ident.lident Prims.list) = fun uu___ -> [] \ No newline at end of file diff --git a/ocaml/fstar-lib/generated/FStar_TypeChecker_Common.ml b/ocaml/fstar-lib/generated/FStar_TypeChecker_Common.ml index 277d1b15ccf..5c2de6825ca 100644 --- a/ocaml/fstar-lib/generated/FStar_TypeChecker_Common.ml +++ b/ocaml/fstar-lib/generated/FStar_TypeChecker_Common.ml @@ -338,7 +338,7 @@ let (id_info__insert : FStar_Compiler_Range.set_def_range range uu___ in let id_ty = match info.identifier with - | FStar_Pervasives.Inr uu___ -> info.identifier_ty + | FStar_Pervasives.Inr uu___ -> ty_map info.identifier_ty | FStar_Pervasives.Inl x -> ty_map info.identifier_ty in let info1 = { diff --git a/src/fstar/FStar.Interactive.Ide.Types.fst b/src/fstar/FStar.Interactive.Ide.Types.fst index 7bf22b125af..b518daa81ca 100755 --- a/src/fstar/FStar.Interactive.Ide.Types.fst +++ b/src/fstar/FStar.Interactive.Ide.Types.fst @@ -100,6 +100,7 @@ type repl_task = type full_buffer_request_kind = | Full : full_buffer_request_kind + | FullBufferWithSymbols : full_buffer_request_kind | Cache : full_buffer_request_kind | ReloadDeps : full_buffer_request_kind | VerifyToPosition of position @@ -241,7 +242,14 @@ let query_to_string q = match q.qq with | Push pq -> "(Push " ^ push_query_to_string pq ^ ")" | VfsAdd _ -> "VfsAdd" | AutoComplete _ -> "AutoComplete" -| Lookup _ -> "Lookup" +| Lookup(s, _lc, pos, features, _sr) -> + BU.format3 "(Lookup %s %s [%s])" + s (match pos with + | None -> "None" + | Some (f, i, j) -> + BU.format3 "(%s, %s, %s)" + f (string_of_int i) (string_of_int j)) + (String.concat "; " features) | Compute _ -> "Compute" | Search _ -> "Search" | GenericError _ -> "GenericError" diff --git a/src/fstar/FStar.Interactive.Ide.fst b/src/fstar/FStar.Interactive.Ide.fst index a59a3dac3bd..dd25dcb06f1 100644 --- a/src/fstar/FStar.Interactive.Ide.fst +++ b/src/fstar/FStar.Interactive.Ide.fst @@ -225,6 +225,7 @@ let unpack_interactive_query json = in let parse_full_buffer_kind = function | "full" -> Full + | "full-with-symbols" -> FullBufferWithSymbols | "cache" -> Cache | "reload-deps" -> ReloadDeps | "verify-to-position" -> VerifyToPosition (read_to_position ()) @@ -696,7 +697,7 @@ let run_push st query = let run_symbol_lookup st symbol pos_opt requested_info (symbol_range_opt:option json) = match QH.symlookup st.repl_env symbol pos_opt requested_info with | None -> Inl "Symbol not found" - | Some result -> + | Some result -> Inr ("symbol", alist_of_symbol_lookup_result result symbol symbol_range_opt) let run_option_lookup opt_name = diff --git a/src/fstar/FStar.Interactive.Incremental.fst b/src/fstar/FStar.Interactive.Incremental.fst index 0a967aa2156..aab3d322b98 100755 --- a/src/fstar/FStar.Interactive.Incremental.fst +++ b/src/fstar/FStar.Interactive.Incremental.fst @@ -86,10 +86,43 @@ let as_query (q:query') qid=qid_prefix ^ "." ^ string_of_int i } +(* This function dumps a symbol table for the decl that has just been checked *) +let dump_symbols_of_term (t:term) +: qst (list query) += let lids = FStar.Parser.AST.Comparison.lidents_of_term t in + let lookup_lid l = + let r = Ident.range_of_lid l in + let start_pos = Range.start_of_range r in + let end_pos = Range.end_of_range r in + let start_line = Range.line_of_pos start_pos in + let start_col = Range.col_of_pos start_pos in + let end_line = Range.line_of_pos end_pos in + let end_col = Range.col_of_pos end_pos in + let position = "", start_line, start_col in + as_query (Lookup(Ident.string_of_lid l, + LKCode, + Some position, + ["type"; "documentation"; "defined-at"], + Some (JsonAssoc [("fname", JsonStr ""); + ("beg", JsonList [JsonInt start_line; JsonInt start_col]); + ("end", JsonList [JsonInt end_line; JsonInt end_col])]))) + in + map lookup_lid lids + +let dump_symbols (d:decl) +: qst (list query) += let open FStar.Parser.AST in + match d.d with + | TopLevelLet (_, defs) -> + let! queries = map (fun (_, t) -> dump_symbols_of_term t) defs in + return (List.concat queries) + | _ -> return [] + (* Push a decl for checking, and before it runs, print a progress message "fragment-started" for the decl that is about to run *) let push_decl (push_kind:push_kind) + (with_symbols:bool) (write_full_buffer_fragment_progress: fragment_progress -> unit) (ds:decl * code_fragment) : qst (list query) @@ -108,13 +141,21 @@ let push_decl (push_kind:push_kind) in let! cb = as_query (Callback progress) in let! push = as_query (Push pq) in - return [cb; push] + if with_symbols + then ( + let! lookups = dump_symbols d in + return ([cb; push] @ lookups) + ) + else ( + return [cb; push] + ) let push_decls (push_kind:push_kind) + (with_symbols:bool) (write_full_buffer_fragment_progress : fragment_progress -> unit) (ds:list (decl & code_fragment)) : qst (list query) - = let! qs = map (push_decl push_kind write_full_buffer_fragment_progress) ds in + = let! qs = map (push_decl push_kind with_symbols write_full_buffer_fragment_progress) ds in return (List.flatten qs) let pop_entries (e:list repl_stack_entry_t) @@ -130,10 +171,11 @@ let repl_task (_, (p, _)) = p let inspect_repl_stack (s:repl_stack_t) (ds:list (decl & code_fragment)) (push_kind : push_kind) + (with_symbols:bool) (write_full_buffer_fragment_progress: fragment_progress -> unit) : qst (list query) = let entries = List.rev s in - let push_decls = push_decls push_kind write_full_buffer_fragment_progress in + let push_decls = push_decls push_kind with_symbols write_full_buffer_fragment_progress in match BU.prefix_until (function (_, (PushFragment _, _)) -> true | _ -> false) entries @@ -239,7 +281,6 @@ let run_full_buffer (st:repl_state) match request_type with | VerifyToPosition (_, line, _col) | LaxToPosition (_, line, _col) -> - BU.print1 "Got to-position: %s" (string_of_int line); List.filter (fun (d, _) -> let start = Range.start_of_range d.drange in @@ -248,13 +289,14 @@ let run_full_buffer (st:repl_state) decls | _ -> decls in + let with_symbols = request_type = FullBufferWithSymbols in let qs = match parse_result with | IncrementalFragment (decls, _, err_opt) -> ( match request_type, decls with | ReloadDeps, d::_ -> run_qst (let! queries = reload_deps (!repl_stack) in - let! push_mod = push_decl FullCheck write_full_buffer_fragment_progress d in + let! push_mod = push_decl FullCheck with_symbols write_full_buffer_fragment_progress d in return (queries @ push_mod)) qid @@ -263,10 +305,12 @@ let run_full_buffer (st:repl_state) let push_kind = match request_type with | LaxToPosition _ -> LaxCheck + | FullBufferWithSymbols -> LaxCheck | _ -> FullCheck in + let queries = - run_qst (inspect_repl_stack (!repl_stack) decls push_kind write_full_buffer_fragment_progress) qid + run_qst (inspect_repl_stack (!repl_stack) decls push_kind with_symbols write_full_buffer_fragment_progress) qid in if request_type = Full then log_syntax_issues err_opt; if Options.debug_any() diff --git a/src/fstar/FStar.Interactive.QueryHelper.fst b/src/fstar/FStar.Interactive.QueryHelper.fst index ce5d202c4f2..022220731fd 100644 --- a/src/fstar/FStar.Interactive.QueryHelper.fst +++ b/src/fstar/FStar.Interactive.QueryHelper.fst @@ -42,7 +42,7 @@ let with_printed_effect_args k = (fun () -> Options.set_option "print_effect_args" (Options.Bool true); k ()) let term_to_string tcenv t = - with_printed_effect_args (fun () -> FStar.TypeChecker.Normalize.term_to_string tcenv t) + with_printed_effect_args (fun () -> FStar.Syntax.Print.term_to_string t) //FStar.TypeChecker.Normalize.term_to_string tcenv t) let sigelt_to_string se = with_printed_effect_args (fun () -> Syntax.Print.sigelt_to_string se) @@ -76,6 +76,9 @@ let symlookup tcenv symbol pos_opt requested_info = match name_or_lid with | Inl name -> name | Inr lid -> Ident.string_of_lid lid in + let str_of_opt = function + | None -> "" + | Some s -> s in let typ_str = if List.mem "type" requested_info then Some (term_to_string tcenv typ) @@ -90,7 +93,6 @@ let symlookup tcenv symbol pos_opt requested_info = | _ -> None in let def_range = if List.mem "defined-at" requested_info then Some rng else None in - Some ({ slr_name = name; slr_def_range = def_range; slr_typ = typ_str; slr_doc = doc_str; slr_def = def_str }) diff --git a/src/fstar/FStar.Main.fst b/src/fstar/FStar.Main.fst index 79bc492b625..3cba3f84d84 100644 --- a/src/fstar/FStar.Main.fst +++ b/src/fstar/FStar.Main.fst @@ -232,7 +232,7 @@ let handle_error e = report_errors [] let main () = - try + // try setup_hooks (); let _, time = Util.record_time go in if FStar.Options.query_stats() @@ -241,6 +241,6 @@ let main () = (String.concat " " (FStar.Getopt.cmdline())); cleanup (); exit 0 - with - | e -> handle_error e; - exit 1 + // with + // | e -> handle_error e; + // exit 1 diff --git a/src/parser/FStar.Parser.AST.Comparison.fst b/src/parser/FStar.Parser.AST.Comparison.fst index b9cb3382cf5..95d88fa33d9 100755 --- a/src/parser/FStar.Parser.AST.Comparison.fst +++ b/src/parser/FStar.Parser.AST.Comparison.fst @@ -547,3 +547,70 @@ and eq_decl (d1 d2:decl) : bool = eq_decl' d1.d d2.d && eq_list eq_qualifier d1.quals d2.quals && eq_list eq_term d1.attrs d2.attrs + +let concat_map = List.collect +let opt_map f (x:option 'a) = match x with | None -> [] | Some x -> f x + +let rec lidents_of_term (t:term) +: list FStar.Ident.lident += lidents_of_term' t.tm +and lidents_of_term' (t:term') +: list FStar.Ident.lident += match t with + | Wild -> [] + | Const _ -> [] + | Op (s, ts) -> concat_map lidents_of_term ts + | Tvar _ -> [] + | Uvar _ -> [] + | Var lid -> [lid] + | Name lid -> [lid] + | Projector (lid, _) -> [lid] + | Construct (lid, ts) -> lid :: concat_map (fun (t, _) -> lidents_of_term t) ts + | Abs (ps, t) -> concat_map lidents_of_pattern ps @ lidents_of_term t + | App (t1, t2, _) -> lidents_of_term t1 @ lidents_of_term t2 + | Let (_, lbs, t) -> concat_map (fun (_, (_, t)) -> lidents_of_term t) lbs @ lidents_of_term t + | LetOperator (lbs, t) -> concat_map (fun (_, _, t) -> lidents_of_term t) lbs @ lidents_of_term t + | LetOpen (lid, t) -> lid :: lidents_of_term t + | LetOpenRecord (t1, t2, t3) -> lidents_of_term t1 @ lidents_of_term t2 @ lidents_of_term t3 + | Seq (t1, t2) -> lidents_of_term t1 @ lidents_of_term t2 + | Bind (_, t1, t2) -> lidents_of_term t1 @ lidents_of_term t2 + | If (t1, _, _, t2, t3) -> lidents_of_term t1 @ lidents_of_term t2 @ lidents_of_term t3 + | Match (t, _, _, bs) -> lidents_of_term t @ concat_map lidents_of_branch bs + | TryWith (t, bs) -> lidents_of_term t @ concat_map lidents_of_branch bs + | Ascribed (t1, t2, _, _) -> lidents_of_term t1 @ lidents_of_term t2 + | Record (t, ts) -> concat_map (fun (_, t) -> lidents_of_term t) ts @ opt_map lidents_of_term t + | Project (t, _) -> lidents_of_term t + | Product (ts, t) -> concat_map lidents_of_binder ts @ lidents_of_term t + | Sum (ts, t) -> concat_map (function Inl b -> lidents_of_binder b | Inr t -> lidents_of_term t) ts @ lidents_of_term t + | QForall (bs, _pats, t) -> lidents_of_term t + | QExists (bs, _pats, t) -> lidents_of_term t + | Refine (b, t) -> lidents_of_term t + | NamedTyp (i, t) -> lidents_of_term t + | Paren t -> lidents_of_term t + | Requires (t, _) -> lidents_of_term t + | Ensures (t, _) -> lidents_of_term t + | LexList ts -> concat_map lidents_of_term ts + | WFOrder (t1, t2) -> lidents_of_term t1 @ lidents_of_term t2 + | Decreases (t, _) -> lidents_of_term t + | Labeled (t, _, _) -> lidents_of_term t + | Discrim lid -> [lid] + | Attributes ts -> concat_map lidents_of_term ts + | Antiquote t -> lidents_of_term t + | Quote (t, _) -> lidents_of_term t + | VQuote t -> lidents_of_term t + | CalcProof (t1, t2, ts) -> lidents_of_term t1 @ lidents_of_term t2 @ concat_map lidents_of_calc_step ts + | IntroForall (bs, t1, t2) -> lidents_of_term t1 @ lidents_of_term t2 + | IntroExists (bs, t1, ts, t2) -> lidents_of_term t1 @ concat_map lidents_of_term ts @ lidents_of_term t2 + | IntroImplies (t1, t2, b, t3) -> lidents_of_term t1 @ lidents_of_term t2 @ lidents_of_term t3 + | IntroOr (b, t1, t2, t3) -> lidents_of_term t1 @ lidents_of_term t2 @ lidents_of_term t3 + | IntroAnd (t1, t2, t3, t4) -> lidents_of_term t1 @ lidents_of_term t2 @ lidents_of_term t3 @ lidents_of_term t4 + | ElimForall (bs, t1, ts) -> lidents_of_term t1 @ concat_map lidents_of_term ts + | ElimExists (bs, t1, t2, b, t3) -> lidents_of_term t1 @ lidents_of_term t2 @ lidents_of_term t3 + | ElimImplies (t1, t2, t3) -> lidents_of_term t1 @ lidents_of_term t2 @ lidents_of_term t3 + | ElimOr (t1, t2, t3, b1, t4, b2, t5) -> lidents_of_term t1 @ lidents_of_term t2 @ lidents_of_term t3 @ lidents_of_term t4 @ lidents_of_term t5 + | ElimAnd (t1, t2, t3, b1, b2, t4) -> lidents_of_term t1 @ lidents_of_term t2 @ lidents_of_term t3 @ lidents_of_term t4 +and lidents_of_branch (p, _, t) = lidents_of_pattern p @ lidents_of_term t +and lidents_of_calc_step = function + | CalcStep (t1, t2, t3) -> lidents_of_term t1 @ lidents_of_term t2 @ lidents_of_term t3 +and lidents_of_pattern _ = [] +and lidents_of_binder _ = [] \ No newline at end of file diff --git a/src/parser/FStar.Parser.AST.Comparison.fsti b/src/parser/FStar.Parser.AST.Comparison.fsti index 71a4e2e91b6..ba2dd14bc26 100755 --- a/src/parser/FStar.Parser.AST.Comparison.fsti +++ b/src/parser/FStar.Parser.AST.Comparison.fsti @@ -24,3 +24,5 @@ open FStar.Parser.AST (* Check if two decls are equal, ignoring range metadata. Used in FStar.Interactive.Incremental *) val eq_decl (d1 d2:decl) : bool + +val lidents_of_term (t:term) : list FStar.Ident.lident \ No newline at end of file From 0344ab7adc1e96e5f6b808e1101df151277abc40 Mon Sep 17 00:00:00 2001 From: Nikhil Swamy Date: Wed, 22 Mar 2023 19:00:00 -0700 Subject: [PATCH 34/48] collecting more names from terms for symbol lookup --- .../generated/FStar_Parser_AST_Comparison.ml | 64 ++++++++++++++++--- src/parser/FStar.Parser.AST.Comparison.fst | 30 +++++++-- 2 files changed, 78 insertions(+), 16 deletions(-) diff --git a/ocaml/fstar-lib/generated/FStar_Parser_AST_Comparison.ml b/ocaml/fstar-lib/generated/FStar_Parser_AST_Comparison.ml index fb250b57829..d754121ba81 100644 --- a/ocaml/fstar-lib/generated/FStar_Parser_AST_Comparison.ml +++ b/ocaml/fstar-lib/generated/FStar_Parser_AST_Comparison.ml @@ -655,15 +655,21 @@ and (lidents_of_term' : (concat_map ()) (fun uu___2 -> match uu___2 with - | (uu___3, (uu___4, t2)) -> lidents_of_term t2) lbs in + | (uu___3, (p, t2)) -> + let uu___4 = lidents_of_pattern p in + let uu___5 = lidents_of_term t2 in + FStar_Compiler_List.op_At uu___4 uu___5) lbs in let uu___2 = lidents_of_term t1 in FStar_Compiler_List.op_At uu___1 uu___2 | FStar_Parser_AST.LetOperator (lbs, t1) -> let uu___ = (concat_map ()) (fun uu___1 -> - match uu___1 with | (uu___2, uu___3, t2) -> lidents_of_term t2) - lbs in + match uu___1 with + | (uu___2, p, t2) -> + let uu___3 = lidents_of_pattern p in + let uu___4 = lidents_of_term t2 in + FStar_Compiler_List.op_At uu___3 uu___4) lbs in let uu___1 = lidents_of_term t1 in FStar_Compiler_List.op_At uu___ uu___1 | FStar_Parser_AST.LetOpen (lid, t1) -> @@ -785,14 +791,20 @@ and (lidents_of_term' : FStar_Compiler_List.op_At uu___2 uu___3 in FStar_Compiler_List.op_At uu___ uu___1 | FStar_Parser_AST.ElimForall (bs, t1, ts) -> - let uu___ = lidents_of_term t1 in - let uu___1 = (concat_map ()) lidents_of_term ts in + let uu___ = (concat_map ()) lidents_of_binder bs in + let uu___1 = + let uu___2 = lidents_of_term t1 in + let uu___3 = (concat_map ()) lidents_of_term ts in + FStar_Compiler_List.op_At uu___2 uu___3 in FStar_Compiler_List.op_At uu___ uu___1 | FStar_Parser_AST.ElimExists (bs, t1, t2, b, t3) -> - let uu___ = lidents_of_term t1 in + let uu___ = (concat_map ()) lidents_of_binder bs in let uu___1 = - let uu___2 = lidents_of_term t2 in - let uu___3 = lidents_of_term t3 in + let uu___2 = lidents_of_term t1 in + let uu___3 = + let uu___4 = lidents_of_term t2 in + let uu___5 = lidents_of_term t3 in + FStar_Compiler_List.op_At uu___4 uu___5 in FStar_Compiler_List.op_At uu___2 uu___3 in FStar_Compiler_List.op_At uu___ uu___1 | FStar_Parser_AST.ElimImplies (t1, t2, t3) -> @@ -849,6 +861,38 @@ and (lidents_of_calc_step : FStar_Compiler_List.op_At uu___1 uu___2 and (lidents_of_pattern : FStar_Parser_AST.pattern -> FStar_Ident.lident Prims.list) = - fun uu___ -> [] + fun p -> + match p.FStar_Parser_AST.pat with + | FStar_Parser_AST.PatWild uu___ -> [] + | FStar_Parser_AST.PatConst uu___ -> [] + | FStar_Parser_AST.PatApp (p1, ps) -> + let uu___ = lidents_of_pattern p1 in + let uu___1 = (concat_map ()) lidents_of_pattern ps in + FStar_Compiler_List.op_At uu___ uu___1 + | FStar_Parser_AST.PatVar (i, uu___, uu___1) -> [] + | FStar_Parser_AST.PatName lid -> [lid] + | FStar_Parser_AST.PatTvar (i, uu___, uu___1) -> [] + | FStar_Parser_AST.PatList ps -> (concat_map ()) lidents_of_pattern ps + | FStar_Parser_AST.PatTuple (ps, uu___) -> + (concat_map ()) lidents_of_pattern ps + | FStar_Parser_AST.PatRecord ps -> + (concat_map ()) + (fun uu___ -> + match uu___ with | (uu___1, p1) -> lidents_of_pattern p1) ps + | FStar_Parser_AST.PatAscribed (p1, (t1, t2)) -> + let uu___ = lidents_of_pattern p1 in + let uu___1 = + let uu___2 = lidents_of_term t1 in + let uu___3 = opt_map lidents_of_term t2 in + FStar_Compiler_List.op_At uu___2 uu___3 in + FStar_Compiler_List.op_At uu___ uu___1 + | FStar_Parser_AST.PatOr ps -> (concat_map ()) lidents_of_pattern ps + | FStar_Parser_AST.PatOp uu___ -> [] + | FStar_Parser_AST.PatVQuote t -> lidents_of_term t and (lidents_of_binder : - FStar_Parser_AST.binder -> FStar_Ident.lident Prims.list) = fun uu___ -> [] \ No newline at end of file + FStar_Parser_AST.binder -> FStar_Ident.lident Prims.list) = + fun b -> + match b.FStar_Parser_AST.b with + | FStar_Parser_AST.Annotated (uu___, t) -> lidents_of_term t + | FStar_Parser_AST.TAnnotated (uu___, t) -> lidents_of_term t + | uu___ -> [] \ No newline at end of file diff --git a/src/parser/FStar.Parser.AST.Comparison.fst b/src/parser/FStar.Parser.AST.Comparison.fst index 95d88fa33d9..f491c163f48 100755 --- a/src/parser/FStar.Parser.AST.Comparison.fst +++ b/src/parser/FStar.Parser.AST.Comparison.fst @@ -568,8 +568,8 @@ and lidents_of_term' (t:term') | Construct (lid, ts) -> lid :: concat_map (fun (t, _) -> lidents_of_term t) ts | Abs (ps, t) -> concat_map lidents_of_pattern ps @ lidents_of_term t | App (t1, t2, _) -> lidents_of_term t1 @ lidents_of_term t2 - | Let (_, lbs, t) -> concat_map (fun (_, (_, t)) -> lidents_of_term t) lbs @ lidents_of_term t - | LetOperator (lbs, t) -> concat_map (fun (_, _, t) -> lidents_of_term t) lbs @ lidents_of_term t + | Let (_, lbs, t) -> concat_map (fun (_, (p, t)) -> lidents_of_pattern p @ lidents_of_term t) lbs @ lidents_of_term t + | LetOperator (lbs, t) -> concat_map (fun (_, p, t) -> lidents_of_pattern p @ lidents_of_term t) lbs @ lidents_of_term t | LetOpen (lid, t) -> lid :: lidents_of_term t | LetOpenRecord (t1, t2, t3) -> lidents_of_term t1 @ lidents_of_term t2 @ lidents_of_term t3 | Seq (t1, t2) -> lidents_of_term t1 @ lidents_of_term t2 @@ -604,13 +604,31 @@ and lidents_of_term' (t:term') | IntroImplies (t1, t2, b, t3) -> lidents_of_term t1 @ lidents_of_term t2 @ lidents_of_term t3 | IntroOr (b, t1, t2, t3) -> lidents_of_term t1 @ lidents_of_term t2 @ lidents_of_term t3 | IntroAnd (t1, t2, t3, t4) -> lidents_of_term t1 @ lidents_of_term t2 @ lidents_of_term t3 @ lidents_of_term t4 - | ElimForall (bs, t1, ts) -> lidents_of_term t1 @ concat_map lidents_of_term ts - | ElimExists (bs, t1, t2, b, t3) -> lidents_of_term t1 @ lidents_of_term t2 @ lidents_of_term t3 + | ElimForall (bs, t1, ts) -> concat_map lidents_of_binder bs @ lidents_of_term t1 @ concat_map lidents_of_term ts + | ElimExists (bs, t1, t2, b, t3) -> concat_map lidents_of_binder bs @ lidents_of_term t1 @ lidents_of_term t2 @ lidents_of_term t3 | ElimImplies (t1, t2, t3) -> lidents_of_term t1 @ lidents_of_term t2 @ lidents_of_term t3 | ElimOr (t1, t2, t3, b1, t4, b2, t5) -> lidents_of_term t1 @ lidents_of_term t2 @ lidents_of_term t3 @ lidents_of_term t4 @ lidents_of_term t5 | ElimAnd (t1, t2, t3, b1, b2, t4) -> lidents_of_term t1 @ lidents_of_term t2 @ lidents_of_term t3 @ lidents_of_term t4 and lidents_of_branch (p, _, t) = lidents_of_pattern p @ lidents_of_term t and lidents_of_calc_step = function | CalcStep (t1, t2, t3) -> lidents_of_term t1 @ lidents_of_term t2 @ lidents_of_term t3 -and lidents_of_pattern _ = [] -and lidents_of_binder _ = [] \ No newline at end of file +and lidents_of_pattern p = + match p.pat with + | PatWild _ -> [] + | PatConst _ -> [] + | PatApp (p, ps) -> lidents_of_pattern p @ concat_map lidents_of_pattern ps + | PatVar (i, _, _) -> [] + | PatName lid -> [lid] + | PatTvar (i, _, _) -> [] + | PatList ps -> concat_map lidents_of_pattern ps + | PatTuple (ps, _) -> concat_map lidents_of_pattern ps + | PatRecord ps -> concat_map (fun (_, p) -> lidents_of_pattern p) ps + | PatAscribed (p, (t1, t2)) -> lidents_of_pattern p @ lidents_of_term t1 @ opt_map lidents_of_term t2 + | PatOr ps -> concat_map lidents_of_pattern ps + | PatOp _ -> [] + | PatVQuote t -> lidents_of_term t +and lidents_of_binder b = + match b.b with + | Annotated (_, t) + | TAnnotated(_, t) -> lidents_of_term t + | _ -> [] \ No newline at end of file From 9671613379fa1ac115ecd22f459b7dcd7438f44a Mon Sep 17 00:00:00 2001 From: Nikhil Swamy Date: Thu, 23 Mar 2023 02:13:02 +0000 Subject: [PATCH 35/48] revert a couple of needless changes --- .../FStar_Interactive_QueryHelper.ml | 2 +- ocaml/fstar-lib/generated/FStar_Main.ml | 39 +++++++++++-------- src/fstar/FStar.Interactive.QueryHelper.fst | 2 +- src/fstar/FStar.Main.fst | 8 ++-- 4 files changed, 29 insertions(+), 22 deletions(-) diff --git a/ocaml/fstar-lib/generated/FStar_Interactive_QueryHelper.ml b/ocaml/fstar-lib/generated/FStar_Interactive_QueryHelper.ml index 97ce93e4c96..306556d0ba0 100644 --- a/ocaml/fstar-lib/generated/FStar_Interactive_QueryHelper.ml +++ b/ocaml/fstar-lib/generated/FStar_Interactive_QueryHelper.ml @@ -44,7 +44,7 @@ let (term_to_string : fun tcenv -> fun t -> with_printed_effect_args - (fun uu___ -> FStar_Syntax_Print.term_to_string t) + (fun uu___ -> FStar_TypeChecker_Normalize.term_to_string tcenv t) let (sigelt_to_string : FStar_Syntax_Syntax.sigelt -> Prims.string) = fun se -> with_printed_effect_args diff --git a/ocaml/fstar-lib/generated/FStar_Main.ml b/ocaml/fstar-lib/generated/FStar_Main.ml index dff0561b53c..8ad6b788389 100644 --- a/ocaml/fstar-lib/generated/FStar_Main.ml +++ b/ocaml/fstar-lib/generated/FStar_Main.ml @@ -320,19 +320,26 @@ let (handle_error : Prims.exn -> unit) = report_errors [] let main : 'uuuuu . unit -> 'uuuuu = fun uu___ -> - setup_hooks (); - (let uu___2 = FStar_Compiler_Util.record_time go in - match uu___2 with - | (uu___3, time) -> - ((let uu___5 = FStar_Options.query_stats () in - if uu___5 - then - let uu___6 = FStar_Compiler_Util.string_of_int time in - let uu___7 = - let uu___8 = FStar_Getopt.cmdline () in - FStar_String.concat " " uu___8 in - FStar_Compiler_Util.print2_error "TOTAL TIME %s ms: %s\n" uu___6 - uu___7 - else ()); - cleanup (); - FStar_Compiler_Effect.exit Prims.int_zero)) \ No newline at end of file + try + (fun uu___1 -> + match () with + | () -> + (setup_hooks (); + (let uu___3 = FStar_Compiler_Util.record_time go in + match uu___3 with + | (uu___4, time) -> + ((let uu___6 = FStar_Options.query_stats () in + if uu___6 + then + let uu___7 = FStar_Compiler_Util.string_of_int time in + let uu___8 = + let uu___9 = FStar_Getopt.cmdline () in + FStar_String.concat " " uu___9 in + FStar_Compiler_Util.print2_error + "TOTAL TIME %s ms: %s\n" uu___7 uu___8 + else ()); + cleanup (); + FStar_Compiler_Effect.exit Prims.int_zero)))) () + with + | uu___1 -> + (handle_error uu___1; FStar_Compiler_Effect.exit Prims.int_one) \ No newline at end of file diff --git a/src/fstar/FStar.Interactive.QueryHelper.fst b/src/fstar/FStar.Interactive.QueryHelper.fst index 022220731fd..48680b3abaf 100644 --- a/src/fstar/FStar.Interactive.QueryHelper.fst +++ b/src/fstar/FStar.Interactive.QueryHelper.fst @@ -42,7 +42,7 @@ let with_printed_effect_args k = (fun () -> Options.set_option "print_effect_args" (Options.Bool true); k ()) let term_to_string tcenv t = - with_printed_effect_args (fun () -> FStar.Syntax.Print.term_to_string t) //FStar.TypeChecker.Normalize.term_to_string tcenv t) + with_printed_effect_args (fun () -> FStar.TypeChecker.Normalize.term_to_string tcenv t) let sigelt_to_string se = with_printed_effect_args (fun () -> Syntax.Print.sigelt_to_string se) diff --git a/src/fstar/FStar.Main.fst b/src/fstar/FStar.Main.fst index 3cba3f84d84..79bc492b625 100644 --- a/src/fstar/FStar.Main.fst +++ b/src/fstar/FStar.Main.fst @@ -232,7 +232,7 @@ let handle_error e = report_errors [] let main () = - // try + try setup_hooks (); let _, time = Util.record_time go in if FStar.Options.query_stats() @@ -241,6 +241,6 @@ let main () = (String.concat " " (FStar.Getopt.cmdline())); cleanup (); exit 0 - // with - // | e -> handle_error e; - // exit 1 + with + | e -> handle_error e; + exit 1 From 81ac27422b7e46a49eb2e54d91eb381277727f71 Mon Sep 17 00:00:00 2001 From: Nikhil Swamy Date: Wed, 22 Mar 2023 22:14:41 -0700 Subject: [PATCH 36/48] symbols for all decls --- .../FStar_Interactive_Incremental.ml | 78 +++++------ ...Comparison.ml => FStar_Parser_AST_Util.ml} | 125 +++++++++++++++++- src/fstar/FStar.Interactive.Incremental.fst | 48 +++---- ...mparison.fst => FStar.Parser.AST.Util.fst} | 70 +++++++++- ...arison.fsti => FStar.Parser.AST.Util.fsti} | 4 +- 5 files changed, 248 insertions(+), 77 deletions(-) rename ocaml/fstar-lib/generated/{FStar_Parser_AST_Comparison.ml => FStar_Parser_AST_Util.ml} (88%) mode change 100644 => 100755 rename src/parser/{FStar.Parser.AST.Comparison.fst => FStar.Parser.AST.Util.fst} (89%) rename src/parser/{FStar.Parser.AST.Comparison.fsti => FStar.Parser.AST.Util.fsti} (90%) diff --git a/ocaml/fstar-lib/generated/FStar_Interactive_Incremental.ml b/ocaml/fstar-lib/generated/FStar_Interactive_Incremental.ml index 5f57da0aeea..8da7aa3a80c 100644 --- a/ocaml/fstar-lib/generated/FStar_Interactive_Incremental.ml +++ b/ocaml/fstar-lib/generated/FStar_Interactive_Incremental.ml @@ -83,54 +83,42 @@ let (as_query : FStar_Interactive_Ide_Types.qid = uu___2 } in return uu___1) -let (dump_symbols_of_term : - FStar_Parser_AST.term -> FStar_Interactive_Ide_Types.query Prims.list qst) - = - fun t -> - let lids = FStar_Parser_AST_Comparison.lidents_of_term t in - let lookup_lid l = - let r = FStar_Ident.range_of_lid l in - let start_pos = FStar_Compiler_Range.start_of_range r in - let end_pos = FStar_Compiler_Range.end_of_range r in - let start_line = FStar_Compiler_Range.line_of_pos start_pos in - let start_col = FStar_Compiler_Range.col_of_pos start_pos in - let end_line = FStar_Compiler_Range.line_of_pos end_pos in - let end_col = FStar_Compiler_Range.col_of_pos end_pos in - let position = ("", start_line, start_col) in - let uu___ = - let uu___1 = - let uu___2 = FStar_Ident.string_of_lid l in - (uu___2, FStar_Interactive_Ide_Types.LKCode, - (FStar_Pervasives_Native.Some position), - ["type"; "documentation"; "defined-at"], - (FStar_Pervasives_Native.Some - (FStar_Compiler_Util.JsonAssoc - [("fname", (FStar_Compiler_Util.JsonStr "")); - ("beg", - (FStar_Compiler_Util.JsonList - [FStar_Compiler_Util.JsonInt start_line; - FStar_Compiler_Util.JsonInt start_col])); - ("end", - (FStar_Compiler_Util.JsonList - [FStar_Compiler_Util.JsonInt end_line; - FStar_Compiler_Util.JsonInt end_col]))]))) in - FStar_Interactive_Ide_Types.Lookup uu___1 in - as_query uu___ in - map lookup_lid lids +let (dump_symbols_for_lid : + FStar_Ident.lident -> FStar_Interactive_Ide_Types.query qst) = + fun l -> + let r = FStar_Ident.range_of_lid l in + let start_pos = FStar_Compiler_Range.start_of_range r in + let end_pos = FStar_Compiler_Range.end_of_range r in + let start_line = FStar_Compiler_Range.line_of_pos start_pos in + let start_col = FStar_Compiler_Range.col_of_pos start_pos in + let end_line = FStar_Compiler_Range.line_of_pos end_pos in + let end_col = FStar_Compiler_Range.col_of_pos end_pos in + let position = ("", start_line, start_col) in + let uu___ = + let uu___1 = + let uu___2 = FStar_Ident.string_of_lid l in + (uu___2, FStar_Interactive_Ide_Types.LKCode, + (FStar_Pervasives_Native.Some position), + ["type"; "documentation"; "defined-at"], + (FStar_Pervasives_Native.Some + (FStar_Compiler_Util.JsonAssoc + [("fname", (FStar_Compiler_Util.JsonStr "")); + ("beg", + (FStar_Compiler_Util.JsonList + [FStar_Compiler_Util.JsonInt start_line; + FStar_Compiler_Util.JsonInt start_col])); + ("end", + (FStar_Compiler_Util.JsonList + [FStar_Compiler_Util.JsonInt end_line; + FStar_Compiler_Util.JsonInt end_col]))]))) in + FStar_Interactive_Ide_Types.Lookup uu___1 in + as_query uu___ let (dump_symbols : FStar_Parser_AST.decl -> FStar_Interactive_Ide_Types.query Prims.list qst) = fun d -> - match d.FStar_Parser_AST.d with - | FStar_Parser_AST.TopLevelLet (uu___, defs) -> - let uu___1 = - map - (fun uu___2 -> - match uu___2 with | (uu___3, t) -> dump_symbols_of_term t) - defs in - op_let_Bang uu___1 - (fun queries -> return (FStar_Compiler_List.concat queries)) - | uu___ -> return [] + let ls = FStar_Parser_AST_Util.lidents_of_decl d in + map dump_symbols_for_lid ls let (push_decl : FStar_Interactive_Ide_Types.push_kind -> Prims.bool -> @@ -263,7 +251,7 @@ let (inspect_repl_stack : | FStar_Interactive_Ide_Types.PushFragment (FStar_Pervasives.Inr d', pk) -> let uu___1 = - FStar_Parser_AST_Comparison.eq_decl + FStar_Parser_AST_Util.eq_decl (FStar_Pervasives_Native.fst d) d' in if uu___1 then diff --git a/ocaml/fstar-lib/generated/FStar_Parser_AST_Comparison.ml b/ocaml/fstar-lib/generated/FStar_Parser_AST_Util.ml old mode 100644 new mode 100755 similarity index 88% rename from ocaml/fstar-lib/generated/FStar_Parser_AST_Comparison.ml rename to ocaml/fstar-lib/generated/FStar_Parser_AST_Util.ml index d754121ba81..95fb06e2a0c --- a/ocaml/fstar-lib/generated/FStar_Parser_AST_Comparison.ml +++ b/ocaml/fstar-lib/generated/FStar_Parser_AST_Util.ml @@ -895,4 +895,127 @@ and (lidents_of_binder : match b.FStar_Parser_AST.b with | FStar_Parser_AST.Annotated (uu___, t) -> lidents_of_term t | FStar_Parser_AST.TAnnotated (uu___, t) -> lidents_of_term t - | uu___ -> [] \ No newline at end of file + | uu___ -> [] +let lidents_of_tycon_record : + 'uuuuu 'uuuuu1 'uuuuu2 . + ('uuuuu * 'uuuuu1 * 'uuuuu2 * FStar_Parser_AST.term) -> + FStar_Ident.lident Prims.list + = + fun uu___ -> + match uu___ with | (uu___1, uu___2, uu___3, t) -> lidents_of_term t +let (lidents_of_constructor_payload : + FStar_Parser_AST.constructor_payload -> FStar_Ident.lident Prims.list) = + fun t -> + match t with + | FStar_Parser_AST.VpOfNotation t1 -> lidents_of_term t1 + | FStar_Parser_AST.VpArbitrary t1 -> lidents_of_term t1 + | FStar_Parser_AST.VpRecord (tc, FStar_Pervasives_Native.None) -> + (concat_map ()) lidents_of_tycon_record tc + | FStar_Parser_AST.VpRecord (tc, FStar_Pervasives_Native.Some t1) -> + let uu___ = (concat_map ()) lidents_of_tycon_record tc in + let uu___1 = lidents_of_term t1 in + FStar_Compiler_List.op_At uu___ uu___1 +let (lidents_of_tycon_variant : + (FStar_Ident.ident * FStar_Parser_AST.constructor_payload + FStar_Pervasives_Native.option * FStar_Parser_AST.attributes_) -> + FStar_Ident.lident Prims.list) + = + fun tc -> + match tc with + | (uu___, FStar_Pervasives_Native.None, uu___1) -> [] + | (uu___, FStar_Pervasives_Native.Some t, uu___1) -> + lidents_of_constructor_payload t +let (lidents_of_tycon : + FStar_Parser_AST.tycon -> FStar_Ident.lident Prims.list) = + fun tc -> + match tc with + | FStar_Parser_AST.TyconAbstract (uu___, bs, k) -> + let uu___1 = (concat_map ()) lidents_of_binder bs in + let uu___2 = opt_map lidents_of_term k in + FStar_Compiler_List.op_At uu___1 uu___2 + | FStar_Parser_AST.TyconAbbrev (uu___, bs, k, t) -> + let uu___1 = (concat_map ()) lidents_of_binder bs in + let uu___2 = + let uu___3 = opt_map lidents_of_term k in + let uu___4 = lidents_of_term t in + FStar_Compiler_List.op_At uu___3 uu___4 in + FStar_Compiler_List.op_At uu___1 uu___2 + | FStar_Parser_AST.TyconRecord (uu___, bs, k, uu___1, tcs) -> + let uu___2 = (concat_map ()) lidents_of_binder bs in + let uu___3 = + let uu___4 = opt_map lidents_of_term k in + let uu___5 = (concat_map ()) lidents_of_tycon_record tcs in + FStar_Compiler_List.op_At uu___4 uu___5 in + FStar_Compiler_List.op_At uu___2 uu___3 + | FStar_Parser_AST.TyconVariant (uu___, bs, k, tcs) -> + let uu___1 = (concat_map ()) lidents_of_binder bs in + let uu___2 = + let uu___3 = opt_map lidents_of_term k in + let uu___4 = (concat_map ()) lidents_of_tycon_variant tcs in + FStar_Compiler_List.op_At uu___3 uu___4 in + FStar_Compiler_List.op_At uu___1 uu___2 +let (lidents_of_lift : + FStar_Parser_AST.lift -> FStar_Ident.lident Prims.list) = + fun l -> + let uu___ = + match l.FStar_Parser_AST.lift_op with + | FStar_Parser_AST.NonReifiableLift t -> lidents_of_term t + | FStar_Parser_AST.ReifiableLift (t1, t2) -> + let uu___1 = lidents_of_term t1 in + let uu___2 = lidents_of_term t2 in + FStar_Compiler_List.op_At uu___1 uu___2 + | FStar_Parser_AST.LiftForFree t -> lidents_of_term t in + FStar_Compiler_List.op_At + [l.FStar_Parser_AST.msource; l.FStar_Parser_AST.mdest] uu___ +let rec (lidents_of_decl : + FStar_Parser_AST.decl -> FStar_Ident.lident Prims.list) = + fun d -> + match d.FStar_Parser_AST.d with + | FStar_Parser_AST.TopLevelModule uu___ -> [] + | FStar_Parser_AST.Open l -> [l] + | FStar_Parser_AST.Friend l -> [l] + | FStar_Parser_AST.Include l -> [l] + | FStar_Parser_AST.ModuleAbbrev (uu___, l) -> [l] + | FStar_Parser_AST.TopLevelLet (_q, lbs) -> + (concat_map ()) + (fun uu___ -> + match uu___ with + | (p, t) -> + let uu___1 = lidents_of_pattern p in + let uu___2 = lidents_of_term t in + FStar_Compiler_List.op_At uu___1 uu___2) lbs + | FStar_Parser_AST.Tycon (uu___, uu___1, tcs) -> + (concat_map ()) lidents_of_tycon tcs + | FStar_Parser_AST.Val (uu___, t) -> lidents_of_term t + | FStar_Parser_AST.Exception (uu___, FStar_Pervasives_Native.None) -> [] + | FStar_Parser_AST.Exception (uu___, FStar_Pervasives_Native.Some t) -> + lidents_of_term t + | FStar_Parser_AST.NewEffect ed -> lidents_of_effect_decl ed + | FStar_Parser_AST.LayeredEffect ed -> lidents_of_effect_decl ed + | FStar_Parser_AST.SubEffect lift -> lidents_of_lift lift + | FStar_Parser_AST.Polymonadic_bind (l0, l1, l2, t) -> + let uu___ = + let uu___1 = let uu___2 = lidents_of_term t in l2 :: uu___2 in l1 + :: uu___1 in + l0 :: uu___ + | FStar_Parser_AST.Polymonadic_subcomp (l0, l1, t) -> + let uu___ = let uu___1 = lidents_of_term t in l1 :: uu___1 in l0 :: + uu___ + | FStar_Parser_AST.Pragma uu___ -> [] + | FStar_Parser_AST.Assume (uu___, t) -> lidents_of_term t + | FStar_Parser_AST.Splice (uu___, t) -> lidents_of_term t +and (lidents_of_effect_decl : + FStar_Parser_AST.effect_decl -> FStar_Ident.lident Prims.list) = + fun ed -> + match ed with + | FStar_Parser_AST.DefineEffect (uu___, bs, t, ds) -> + let uu___1 = (concat_map ()) lidents_of_binder bs in + let uu___2 = + let uu___3 = lidents_of_term t in + let uu___4 = (concat_map ()) lidents_of_decl ds in + FStar_Compiler_List.op_At uu___3 uu___4 in + FStar_Compiler_List.op_At uu___1 uu___2 + | FStar_Parser_AST.RedefineEffect (uu___, bs, t) -> + let uu___1 = (concat_map ()) lidents_of_binder bs in + let uu___2 = lidents_of_term t in + FStar_Compiler_List.op_At uu___1 uu___2 \ No newline at end of file diff --git a/src/fstar/FStar.Interactive.Incremental.fst b/src/fstar/FStar.Interactive.Incremental.fst index aab3d322b98..4c71eec518a 100755 --- a/src/fstar/FStar.Interactive.Incremental.fst +++ b/src/fstar/FStar.Interactive.Incremental.fst @@ -42,7 +42,7 @@ open FStar.Interactive.Ide.Types module P = FStar.Parser.ParseIt module BU = FStar.Compiler.Util open FStar.Parser.AST -open FStar.Parser.AST.Comparison +open FStar.Parser.AST.Util let qid = string & int let qst a = qid -> a & qid @@ -87,36 +87,30 @@ let as_query (q:query') } (* This function dumps a symbol table for the decl that has just been checked *) -let dump_symbols_of_term (t:term) -: qst (list query) -= let lids = FStar.Parser.AST.Comparison.lidents_of_term t in - let lookup_lid l = - let r = Ident.range_of_lid l in - let start_pos = Range.start_of_range r in - let end_pos = Range.end_of_range r in - let start_line = Range.line_of_pos start_pos in - let start_col = Range.col_of_pos start_pos in - let end_line = Range.line_of_pos end_pos in - let end_col = Range.col_of_pos end_pos in - let position = "", start_line, start_col in - as_query (Lookup(Ident.string_of_lid l, - LKCode, - Some position, - ["type"; "documentation"; "defined-at"], - Some (JsonAssoc [("fname", JsonStr ""); - ("beg", JsonList [JsonInt start_line; JsonInt start_col]); - ("end", JsonList [JsonInt end_line; JsonInt end_col])]))) - in - map lookup_lid lids +let dump_symbols_for_lid (l:lident) +: qst query += let r = Ident.range_of_lid l in + let start_pos = Range.start_of_range r in + let end_pos = Range.end_of_range r in + let start_line = Range.line_of_pos start_pos in + let start_col = Range.col_of_pos start_pos in + let end_line = Range.line_of_pos end_pos in + let end_col = Range.col_of_pos end_pos in + let position = "", start_line, start_col in + as_query (Lookup(Ident.string_of_lid l, + LKCode, + Some position, + ["type"; "documentation"; "defined-at"], + Some (JsonAssoc [("fname", JsonStr ""); + ("beg", JsonList [JsonInt start_line; JsonInt start_col]); + ("end", JsonList [JsonInt end_line; JsonInt end_col])]))) let dump_symbols (d:decl) : qst (list query) = let open FStar.Parser.AST in - match d.d with - | TopLevelLet (_, defs) -> - let! queries = map (fun (_, t) -> dump_symbols_of_term t) defs in - return (List.concat queries) - | _ -> return [] + let ls = lidents_of_decl d in + map dump_symbols_for_lid ls + (* Push a decl for checking, and before it runs, print a progress message "fragment-started" diff --git a/src/parser/FStar.Parser.AST.Comparison.fst b/src/parser/FStar.Parser.AST.Util.fst similarity index 89% rename from src/parser/FStar.Parser.AST.Comparison.fst rename to src/parser/FStar.Parser.AST.Util.fst index f491c163f48..6da4e324eb1 100755 --- a/src/parser/FStar.Parser.AST.Comparison.fst +++ b/src/parser/FStar.Parser.AST.Util.fst @@ -15,7 +15,7 @@ Authors: N. Swamy and Copilot *) -module FStar.Parser.AST.Comparison +module FStar.Parser.AST.Util open FStar.Pervasives open FStar.Compiler.Effect open FStar.Compiler.List @@ -631,4 +631,70 @@ and lidents_of_binder b = match b.b with | Annotated (_, t) | TAnnotated(_, t) -> lidents_of_term t - | _ -> [] \ No newline at end of file + | _ -> [] + +let lidents_of_tycon_record (_, _, _, t) = + lidents_of_term t + +let lidents_of_constructor_payload (t:constructor_payload) = + match t with + | VpOfNotation t -> lidents_of_term t + | VpArbitrary t -> lidents_of_term t + | VpRecord (tc, None) -> concat_map lidents_of_tycon_record tc + | VpRecord (tc, Some t) -> concat_map lidents_of_tycon_record tc @ lidents_of_term t + +let lidents_of_tycon_variant (tc:(ident * option constructor_payload * attributes_)) = + match tc with + | _, None, _ -> [] + | _, Some t, _ -> lidents_of_constructor_payload t + +let lidents_of_tycon (tc:tycon) = + match tc with + | TyconAbstract (_, bs, k) -> concat_map lidents_of_binder bs @ opt_map lidents_of_term k + | TyconAbbrev (_, bs, k, t) -> concat_map lidents_of_binder bs @ opt_map lidents_of_term k @ lidents_of_term t + | TyconRecord (_, bs, k, _, tcs) -> + concat_map lidents_of_binder bs @ + opt_map lidents_of_term k @ + concat_map lidents_of_tycon_record tcs + | TyconVariant (_, bs, k, tcs) -> + concat_map lidents_of_binder bs @ + opt_map lidents_of_term k @ + concat_map lidents_of_tycon_variant tcs + +let lidents_of_lift (l:lift) = + [l.msource; l.mdest]@ + (match l.lift_op with + | NonReifiableLift t -> lidents_of_term t + | ReifiableLift (t1, t2) -> lidents_of_term t1 @ lidents_of_term t2 + | LiftForFree t -> lidents_of_term t) + +let rec lidents_of_decl (d:decl) = + match d.d with + | TopLevelModule _ -> [] + | Open l + | Friend l + | Include l + | ModuleAbbrev (_, l) -> [l] + | TopLevelLet (_q, lbs) -> concat_map (fun (p, t) -> lidents_of_pattern p @ lidents_of_term t) lbs + | Tycon (_, _, tcs) -> concat_map lidents_of_tycon tcs + | Val (_, t) -> lidents_of_term t + | Exception (_, None) -> [] + | Exception (_, Some t) -> lidents_of_term t + | NewEffect ed + | LayeredEffect ed -> lidents_of_effect_decl ed + | SubEffect lift -> lidents_of_lift lift + | Polymonadic_bind(l0, l1, l2, t) -> l0::l1::l2::lidents_of_term t + | Polymonadic_subcomp(l0, l1, t) -> l0::l1::lidents_of_term t + | Pragma _ -> [] + | Assume (_, t) -> lidents_of_term t + | Splice (_, t) -> lidents_of_term t + +and lidents_of_effect_decl (ed:effect_decl) = + match ed with + | DefineEffect (_, bs, t, ds) -> + concat_map lidents_of_binder bs @ + lidents_of_term t @ + concat_map lidents_of_decl ds + | RedefineEffect (_, bs, t) -> + concat_map lidents_of_binder bs @ + lidents_of_term t \ No newline at end of file diff --git a/src/parser/FStar.Parser.AST.Comparison.fsti b/src/parser/FStar.Parser.AST.Util.fsti similarity index 90% rename from src/parser/FStar.Parser.AST.Comparison.fsti rename to src/parser/FStar.Parser.AST.Util.fsti index ba2dd14bc26..88098d60e8c 100755 --- a/src/parser/FStar.Parser.AST.Comparison.fsti +++ b/src/parser/FStar.Parser.AST.Util.fsti @@ -15,7 +15,7 @@ Authors: N. Swamy and Copilot *) -module FStar.Parser.AST.Comparison +module FStar.Parser.AST.Util open FStar.Pervasives open FStar.Compiler.Effect open FStar.Compiler.List @@ -25,4 +25,4 @@ open FStar.Parser.AST Used in FStar.Interactive.Incremental *) val eq_decl (d1 d2:decl) : bool -val lidents_of_term (t:term) : list FStar.Ident.lident \ No newline at end of file +val lidents_of_decl (t:decl) : list FStar.Ident.lident From 3e036df012504e49a4cd4a3e71f1e9b9e1fc9d76 Mon Sep 17 00:00:00 2001 From: Nikhil Swamy Date: Thu, 23 Mar 2023 17:05:24 -0700 Subject: [PATCH 37/48] some more fine-tuning & simplifcation of raw contents of a buffer in incremental parsing --- ocaml/fstar-lib/FStar_Parser_ParseIt.ml | 123 ++++++++++++++++++------ ocaml/fstar-lib/FStar_Sedlexing.ml | 2 +- src/basic/FStar.Ident.fsti | 10 +- src/parser/FStar.Parser.AST.Util.fst | 5 +- 4 files changed, 103 insertions(+), 37 deletions(-) diff --git a/ocaml/fstar-lib/FStar_Parser_ParseIt.ml b/ocaml/fstar-lib/FStar_Parser_ParseIt.ml index 0b65ff74ed6..f3692a4dc05 100644 --- a/ocaml/fstar-lib/FStar_Parser_ParseIt.ml +++ b/ocaml/fstar-lib/FStar_Parser_ParseIt.ml @@ -111,6 +111,9 @@ type parse_result = | Term of FStar_Parser_AST.term | ParseError of parse_error +module BU = FStar_Compiler_Util +module Range = FStar_Compiler_Range + let parse fn = FStar_Parser_Util.warningHandler := (function | e -> Printf.printf "There was some warning (TODO)\n"); @@ -146,41 +149,86 @@ let parse fn = let parse_one_decl = MenhirLib.Convert.Simplified.traditional2revised FStar_Parser_Parse.oneDeclOrEOF in let contents_at = let lines = U.splitlines contents in + let split_line_at_col line col = + if col > 0 + then ( + (* Don't index directly into the string, since this is a UTF-8 string. + Convert first to a list of charaters, index into that, and then convert + back to a string *) + let chars = FStar_String.list_of_string line in + if col <= List.length chars + then ( + let prefix, suffix = FStar_Compiler_Util.first_N (Z.of_int col) chars in + Some (FStar_String.string_of_list prefix, + FStar_String.string_of_list suffix) + ) + else ( + None + ) + ) + else None + in + let line_from_col line pos = + match split_line_at_col line pos with + | None -> None + | Some (_, p) -> Some p + in + let line_to_col line pos = + match split_line_at_col line pos with + | None -> None + | Some (p, _) -> Some p + in (* Find the raw content of the input from the line of the start_pos to the end_pos. This is used by Interactive.Incremental to record exactly the raw content of the fragment that was checked *) - fun (start_pos:Lexing.position) (end_pos:Lexing.position) -> + fun (range:Range.range) -> (* discard all lines until the start line *) - let suffix = FStar_Compiler_Util.nth_tail (Z.of_int (if start_pos.pos_lnum > 0 then start_pos.pos_lnum - 1 else 0)) lines in + let start_pos = Range.start_of_range range in + let end_pos = Range.end_of_range range in + let start_line = Z.to_int (Range.line_of_pos start_pos) in + let start_col = Z.to_int (Range.col_of_pos start_pos) in + let end_line = Z.to_int (Range.line_of_pos end_pos) in + let end_col = Z.to_int (Range.col_of_pos end_pos) in + let suffix = + FStar_Compiler_Util.nth_tail + (Z.of_int (if start_line > 0 then start_line - 1 else 0)) + lines + in (* Take all the lines between the start and end lines *) - let text, rest = FStar_Compiler_Util.first_N (Z.of_int (end_pos.pos_lnum - start_pos.pos_lnum)) suffix in + let text, rest = + FStar_Compiler_Util.first_N + (Z.of_int (end_line - start_line)) + suffix + in + let text = + match text with + | first_line::rest -> ( + match line_from_col first_line start_col with + | Some s -> s :: rest + | _ -> text + ) + | _ -> text + in let text = (* For the last line itself, take the prefix of it up to the character of the end_pos *) match rest with - | last::_ -> - let col = end_pos.pos_cnum - end_pos.pos_bol in - if col > 0 - then ( - (* Don't index directly into the string, since this is a UTF-8 string. - Convert first to a list of charaters, index into that, and then convert - back to a string *) - let chars = FStar_String.list_of_string last in - if col <= List.length chars - then ( - let chars, _ = FStar_Compiler_Util.first_N (Z.of_int col) chars in - let last = FStar_String.string_of_list chars in - text@[last] + | last::_ -> ( + match line_to_col last end_col with + | None -> text + | Some last -> + (* The last line is also the first line *) + match text with + | [] -> ( + match line_from_col last start_col with + | None -> [last] + | Some l -> [l] ) - else ( - text - ) - ) - else text + | _ -> text @ [last] + ) | _ -> text in - let range = range_of_positions start_pos end_pos in - { range; - code = FStar_String.concat "\n" text } + { range; + code = FStar_String.concat "\n" text } in let open FStar_Pervasives in let rec parse decls = @@ -208,10 +256,25 @@ let parse fn = Rollback the lexer one token for declarations whose syntax requires such lookahead to complete a production. *) - if not (FStar_Parser_AST.decl_syntax_is_delimited d) - then rollback lexbuf; - let end_pos = current_pos lexbuf in - let raw_contents = contents_at start_pos end_pos in + let end_pos = + if not (FStar_Parser_AST.decl_syntax_is_delimited d) + then ( + rollback lexbuf; + current_pos lexbuf + ) + else ( + current_pos lexbuf + ) + in + let raw_contents = contents_at d.drange in + if FStar_Options.debug_any() + then ( + FStar_Compiler_Util.print4 "Parsed decl@%s=%s\nRaw contents@%s=%s\n" + (FStar_Compiler_Range.string_of_def_range d.drange) + (FStar_Parser_AST.decl_to_string d) + (FStar_Compiler_Range.string_of_def_range raw_contents.range) + raw_contents.code + ); parse ((d, raw_contents)::decls) | Inr err -> List.rev decls, Some err in @@ -231,7 +294,9 @@ let parse fn = | Filename _ | Toplevel _ -> begin let fileOrFragment = - MenhirLib.Convert.Simplified.traditional2revised FStar_Parser_Parse.inputFragment lexer + if FStar_Options.debug_any() + then parse_incremental_fragment() + else MenhirLib.Convert.Simplified.traditional2revised FStar_Parser_Parse.inputFragment lexer in let frags = match fileOrFragment with | FStar_Pervasives.Inl modul -> diff --git a/ocaml/fstar-lib/FStar_Sedlexing.ml b/ocaml/fstar-lib/FStar_Sedlexing.ml index 1c4fe22f2f9..b0434303dc7 100644 --- a/ocaml/fstar-lib/FStar_Sedlexing.ml +++ b/ocaml/fstar-lib/FStar_Sedlexing.ml @@ -53,7 +53,7 @@ let create (s:string) fn loffset coffset = mark_val = 0; } -let current_pos b = b.start_p +let current_pos b = b.cur_p let start b = b.mark <- b.cur; diff --git a/src/basic/FStar.Ident.fsti b/src/basic/FStar.Ident.fsti index 24113c6782f..fb69e389ccf 100644 --- a/src/basic/FStar.Ident.fsti +++ b/src/basic/FStar.Ident.fsti @@ -19,17 +19,17 @@ open FStar.Compiler.Range (** A (short) identifier for a local name. * e.g. x in `fun x -> ...` *) -// IN F*: [@@ PpxDerivingYoJson; PpxDerivingShow ] +[@@ PpxDerivingYoJson; PpxDerivingShow ] new val ident : Type0 // type ident (** A module path *) -// IN F*: [@@ PpxDerivingYoJson; PpxDerivingShow ] +[@@ PpxDerivingYoJson; PpxDerivingShow ] type path = list string (** A module path, as idents *) -// IN F*: [@@ PpxDerivingYoJson; PpxDerivingShow ] +[@@ PpxDerivingYoJson; PpxDerivingShow ] type ipath = list ident (** Create an ident *) @@ -81,10 +81,10 @@ val path_of_ns : ipath -> path e.g. Prims.string. Essentially a list of idents where the last one denotes a name, and all the others denote a module path that qualifies the name. *) -// IN F*: [@@ PpxDerivingYoJson; PpxDerivingShow ] +[@@ PpxDerivingYoJson; PpxDerivingShow ] new val lident : Type0 -// IN F*: [@@ PpxDerivingYoJson; PpxDerivingShow ] +[@@ PpxDerivingYoJson; PpxDerivingShow ] type lid = lident (** Obtain the range of an lid *) diff --git a/src/parser/FStar.Parser.AST.Util.fst b/src/parser/FStar.Parser.AST.Util.fst index 6da4e324eb1..5e3541fb39a 100755 --- a/src/parser/FStar.Parser.AST.Util.fst +++ b/src/parser/FStar.Parser.AST.Util.fst @@ -630,7 +630,8 @@ and lidents_of_pattern p = and lidents_of_binder b = match b.b with | Annotated (_, t) - | TAnnotated(_, t) -> lidents_of_term t + | TAnnotated(_, t) + | NoName t -> lidents_of_term t | _ -> [] let lidents_of_tycon_record (_, _, _, t) = @@ -697,4 +698,4 @@ and lidents_of_effect_decl (ed:effect_decl) = concat_map lidents_of_decl ds | RedefineEffect (_, bs, t) -> concat_map lidents_of_binder bs @ - lidents_of_term t \ No newline at end of file + lidents_of_term t From 57b5b20ed7ed60e1bd353e2e07c2a9be09dd4ba0 Mon Sep 17 00:00:00 2001 From: Nikhil Swamy Date: Fri, 24 Mar 2023 04:55:34 +0000 Subject: [PATCH 38/48] snap --- ocaml/fstar-lib/generated/FStar_Parser_AST_Util.ml | 1 + 1 file changed, 1 insertion(+) diff --git a/ocaml/fstar-lib/generated/FStar_Parser_AST_Util.ml b/ocaml/fstar-lib/generated/FStar_Parser_AST_Util.ml index 95fb06e2a0c..88df1ed74cf 100755 --- a/ocaml/fstar-lib/generated/FStar_Parser_AST_Util.ml +++ b/ocaml/fstar-lib/generated/FStar_Parser_AST_Util.ml @@ -895,6 +895,7 @@ and (lidents_of_binder : match b.FStar_Parser_AST.b with | FStar_Parser_AST.Annotated (uu___, t) -> lidents_of_term t | FStar_Parser_AST.TAnnotated (uu___, t) -> lidents_of_term t + | FStar_Parser_AST.NoName t -> lidents_of_term t | uu___ -> [] let lidents_of_tycon_record : 'uuuuu 'uuuuu1 'uuuuu2 . From 754a6d8f1d2d5c6a0816bf24499a12896e0eb369 Mon Sep 17 00:00:00 2001 From: Nikhil Swamy Date: Fri, 24 Mar 2023 11:19:37 -0700 Subject: [PATCH 39/48] classes have delimited syntax (thanks guido); do not fail when lookup sym not found --- ocaml/fstar-lib/generated/FStar_Interactive_Ide.ml | 2 +- ocaml/fstar-lib/generated/FStar_Parser_AST.ml | 1 + src/fstar/FStar.Interactive.Ide.fst | 3 ++- src/parser/FStar.Parser.AST.fst | 1 + 4 files changed, 5 insertions(+), 2 deletions(-) diff --git a/ocaml/fstar-lib/generated/FStar_Interactive_Ide.ml b/ocaml/fstar-lib/generated/FStar_Interactive_Ide.ml index 6031d184712..16067aca6d0 100644 --- a/ocaml/fstar-lib/generated/FStar_Interactive_Ide.ml +++ b/ocaml/fstar-lib/generated/FStar_Interactive_Ide.ml @@ -1923,7 +1923,7 @@ let run_lookup : run_lookup' st symbol context pos_opt requested_info symrange in match uu___ with | FStar_Pervasives.Inl err_msg -> - ((FStar_Interactive_Ide_Types.QueryNOK, + ((FStar_Interactive_Ide_Types.QueryOK, (FStar_Compiler_Util.JsonStr err_msg)), (FStar_Pervasives.Inl st)) | FStar_Pervasives.Inr (kind, info) -> diff --git a/ocaml/fstar-lib/generated/FStar_Parser_AST.ml b/ocaml/fstar-lib/generated/FStar_Parser_AST.ml index 8e01c7900f2..5b0e6267bf5 100644 --- a/ocaml/fstar-lib/generated/FStar_Parser_AST.ml +++ b/ocaml/fstar-lib/generated/FStar_Parser_AST.ml @@ -2438,4 +2438,5 @@ let (decl_syntax_is_delimited : decl -> Prims.bool) = | SubEffect { msource = uu___; mdest = uu___1; lift_op = uu___2; braced = true;_} -> true + | Tycon (uu___, b, uu___1) -> b | uu___ -> false \ No newline at end of file diff --git a/src/fstar/FStar.Interactive.Ide.fst b/src/fstar/FStar.Interactive.Ide.fst index 7850fda6850..6ffc9f61688 100644 --- a/src/fstar/FStar.Interactive.Ide.fst +++ b/src/fstar/FStar.Interactive.Ide.fst @@ -733,7 +733,8 @@ let run_lookup' st symbol context pos_opt requested_info symrange = let run_lookup st symbol context pos_opt requested_info symrange = match run_lookup' st symbol context pos_opt requested_info symrange with | Inl err_msg -> - ((QueryNOK, JsonStr err_msg), Inl st) + // No result found, but don't fail the query + ((QueryOK, JsonStr err_msg), Inl st) | Inr (kind, info) -> ((QueryOK, JsonAssoc (("kind", JsonStr kind) :: info)), Inl st) diff --git a/src/parser/FStar.Parser.AST.fst b/src/parser/FStar.Parser.AST.fst index 79b8ecb3d3d..f990571b82b 100644 --- a/src/parser/FStar.Parser.AST.fst +++ b/src/parser/FStar.Parser.AST.fst @@ -995,5 +995,6 @@ let decl_syntax_is_delimited (d:decl) = | NewEffect (DefineEffect _) | LayeredEffect (DefineEffect _) | SubEffect {braced=true} -> true + | Tycon(_, b, _) -> b | _ -> false From b767e7f8527becacfea604bfe7a7e3eaea672e32 Mon Sep 17 00:00:00 2001 From: Nikhil Swamy Date: Fri, 24 Mar 2023 16:00:15 -0700 Subject: [PATCH 40/48] run the incremental IDE protocol on all files in ulib from test-incremental.py --- .../FStar_Interactive_Incremental.ml | 118 +++++------ src/fstar/FStar.Interactive.Incremental.fst | 1 + tests/ide/Makefile | 1 + tests/ide/test-incremental.py | 189 ++++++++++++++++++ 4 files changed, 253 insertions(+), 56 deletions(-) create mode 100644 tests/ide/test-incremental.py diff --git a/ocaml/fstar-lib/generated/FStar_Interactive_Incremental.ml b/ocaml/fstar-lib/generated/FStar_Interactive_Incremental.ml index 8da7aa3a80c..45b85464893 100644 --- a/ocaml/fstar-lib/generated/FStar_Interactive_Incremental.ml +++ b/ocaml/fstar-lib/generated/FStar_Interactive_Incremental.ml @@ -402,62 +402,68 @@ let (run_full_buffer : match parse_result with | FStar_Parser_ParseIt.IncrementalFragment (decls, uu___, err_opt) -> - (match (request_type, decls) with - | (FStar_Interactive_Ide_Types.ReloadDeps, d::uu___1) -> - let uu___2 = - let uu___3 = - let uu___4 = - FStar_Compiler_Effect.op_Bang - FStar_Interactive_PushHelper.repl_stack in - reload_deps uu___4 in - op_let_Bang uu___3 - (fun queries -> - let uu___4 = - push_decl - FStar_Interactive_Ide_Types.FullCheck - with_symbols - write_full_buffer_fragment_progress d in - op_let_Bang uu___4 - (fun push_mod -> - return - (FStar_Compiler_List.op_At queries - push_mod))) in - run_qst uu___2 qid1 - | uu___1 -> - let decls1 = filter_decls decls in - let push_kind = - match request_type with - | FStar_Interactive_Ide_Types.LaxToPosition uu___2 - -> FStar_Interactive_Ide_Types.LaxCheck - | FStar_Interactive_Ide_Types.FullBufferWithSymbols - -> FStar_Interactive_Ide_Types.LaxCheck - | uu___2 -> FStar_Interactive_Ide_Types.FullCheck in - let queries = - let uu___2 = - let uu___3 = - FStar_Compiler_Effect.op_Bang - FStar_Interactive_PushHelper.repl_stack in - inspect_repl_stack uu___3 decls1 push_kind - with_symbols write_full_buffer_fragment_progress in - run_qst uu___2 qid1 in - (if request_type = FStar_Interactive_Ide_Types.Full - then log_syntax_issues err_opt - else (); - (let uu___4 = FStar_Options.debug_any () in - if uu___4 - then - let uu___5 = - let uu___6 = - FStar_Compiler_List.map - FStar_Interactive_Ide_Types.query_to_string - queries in - FStar_String.concat "\n" uu___6 in - FStar_Compiler_Util.print1 - "Generating queries\n%s\n" uu___5 - else ()); - if request_type <> FStar_Interactive_Ide_Types.Cache - then queries - else [])) + ((let uu___2 = + FStar_Compiler_Util.string_of_int + (FStar_Compiler_List.length decls) in + FStar_Compiler_Util.print1 "Parsed %s declarations\n" + uu___2); + (match (request_type, decls) with + | (FStar_Interactive_Ide_Types.ReloadDeps, d::uu___2) -> + let uu___3 = + let uu___4 = + let uu___5 = + FStar_Compiler_Effect.op_Bang + FStar_Interactive_PushHelper.repl_stack in + reload_deps uu___5 in + op_let_Bang uu___4 + (fun queries -> + let uu___5 = + push_decl + FStar_Interactive_Ide_Types.FullCheck + with_symbols + write_full_buffer_fragment_progress d in + op_let_Bang uu___5 + (fun push_mod -> + return + (FStar_Compiler_List.op_At queries + push_mod))) in + run_qst uu___3 qid1 + | uu___2 -> + let decls1 = filter_decls decls in + let push_kind = + match request_type with + | FStar_Interactive_Ide_Types.LaxToPosition uu___3 + -> FStar_Interactive_Ide_Types.LaxCheck + | FStar_Interactive_Ide_Types.FullBufferWithSymbols + -> FStar_Interactive_Ide_Types.LaxCheck + | uu___3 -> FStar_Interactive_Ide_Types.FullCheck in + let queries = + let uu___3 = + let uu___4 = + FStar_Compiler_Effect.op_Bang + FStar_Interactive_PushHelper.repl_stack in + inspect_repl_stack uu___4 decls1 push_kind + with_symbols + write_full_buffer_fragment_progress in + run_qst uu___3 qid1 in + (if request_type = FStar_Interactive_Ide_Types.Full + then log_syntax_issues err_opt + else (); + (let uu___5 = FStar_Options.debug_any () in + if uu___5 + then + let uu___6 = + let uu___7 = + FStar_Compiler_List.map + FStar_Interactive_Ide_Types.query_to_string + queries in + FStar_String.concat "\n" uu___7 in + FStar_Compiler_Util.print1 + "Generating queries\n%s\n" uu___6 + else ()); + if request_type <> FStar_Interactive_Ide_Types.Cache + then queries + else []))) | FStar_Parser_ParseIt.ParseError err -> (if request_type = FStar_Interactive_Ide_Types.Full then log_syntax_issues (FStar_Pervasives_Native.Some err) diff --git a/src/fstar/FStar.Interactive.Incremental.fst b/src/fstar/FStar.Interactive.Incremental.fst index 4c71eec518a..e4ae4433772 100755 --- a/src/fstar/FStar.Interactive.Incremental.fst +++ b/src/fstar/FStar.Interactive.Incremental.fst @@ -287,6 +287,7 @@ let run_full_buffer (st:repl_state) let qs = match parse_result with | IncrementalFragment (decls, _, err_opt) -> ( + BU.print1 "Parsed %s declarations\n" (string_of_int (List.length decls)); match request_type, decls with | ReloadDeps, d::_ -> run_qst (let! queries = reload_deps (!repl_stack) in diff --git a/tests/ide/Makefile b/tests/ide/Makefile index feecbb62ebf..a18d2ed5b6b 100644 --- a/tests/ide/Makefile +++ b/tests/ide/Makefile @@ -10,6 +10,7 @@ clean: $(addsuffix .clean, $(ALL_TEST_DIRS)) %.all: % +$(MAKE) -C $^ all + python test-incremental.py %.clean: % +$(MAKE) -C $^ clean diff --git a/tests/ide/test-incremental.py b/tests/ide/test-incremental.py new file mode 100644 index 00000000000..863c8d14eb5 --- /dev/null +++ b/tests/ide/test-incremental.py @@ -0,0 +1,189 @@ +# A script to test F* interactive mode's incremental capabilities + +import os +import sys +import subprocess +import json +import re + + +# The path to the F* executable +fstar = sys.argv[1] + +# A function to validate the response from F* interactive mode +def validate_response(response): + # parse the each line of the response into a JSON object + # if the line is not valid JSON, print an error message and exit + # store the JSON objects in a list + json_objects = [] + for line in response.splitlines(): + try: + resp = json.loads(line) + json_objects.append(resp) + except json.JSONDecodeError: + print(f"Invalid JSON: {line}") + return False + # Check that the first line is a protocol-info message + if json_objects[0]["kind"] != "protocol-info": + print("First line is not a protocol-info message") + return False + # Second line has kind "response" and query-id "1" and status "success" + if json_objects[1]["kind"] != "response": + print("Second line is not a response message") + return False + if json_objects[1]["query-id"] != "1": + print("Second line does not have query-id 1") + return False + if json_objects[1]["status"] != "success": + print("Second line does not have status success") + return False + + # Third line has the form {"kind":"message","query-id":"2","level":"info","contents":"Parsed 138 declarations\n"} + if json_objects[2]["kind"] != "message": + print(f"Second line is not a message {json_objects[2]}") + return False + if json_objects[2]["query-id"] != "2": + print("Second line does not have query-id 2") + return False + if json_objects[2]["level"] != "info": + print("Second line does not have level info") + return False + # the contents should match a regular expression of the form "Parsed \d+ declarations" + # store the number of declarations in a variable + if not re.match(r"Parsed \d+ declarations", json_objects[2]["contents"]): + print("Second line does not have the correct contents") + return False + # Check that the number of declarations is 138 + num_decls = int(re.search(r"\d+", json_objects[2]["contents"]).group()) + + # Fourth line has kind "message" and level "progress" and contents.stage = "full-buffer-fragment-started" + if json_objects[3]["kind"] != "message": + print("Third line is not a message") + return False + if json_objects[3]["level"] != "progress": + print("Third line does not have level progress") + return False + if json_objects[3]["contents"]["stage"] != "full-buffer-fragment-started": + print("Third line does not have stage full-buffer-fragment-started") + return False + + # Next several messages are progress messages with contents.stage = "loading-dependency" + # Check all of these messages for the correct kind, level, and stage and stop + # when the first message with a different kind or level or stage is found + for i in range(4, len(json_objects) - 1): + if json_objects[i]["kind"] != "message": + break + if json_objects[i]["level"] != "progress": + break + if json_objects[i]["contents"]["stage"] != "loading-dependency": + break + # the message and index i has contents.stage = None + if json_objects[i]["contents"]["stage"] != None: + print(f"Message {i} has contents {json_objects[i]} does not have stage None") + return False + # the next message has conents.stage = "full-buffer-fragment-lax-ok" + if json_objects[i + 1]["contents"]["stage"] != "full-buffer-fragment-lax-ok": + print("Message does not have stage full-buffer-fragment-lax-ok") + return False + # Then, we have a sequence of pairs of messages + # where the first message in the pair has contents.stage = "full-buffer-fragment-started" + # and the second message in the pair has contents.stage = "full-buffer-fragment-lax-ok" + # Check all of these messages for the correct kind, level, and stage and stop + # when the first message with a different kind or level or stage is found + # The first message in the pair has contents.stage = "full-buffer-fragment-started" + # The second message in the pair has contents.stage = "full-buffer-fragment-lax-ok" + num_successes = 1 + for j in range(i + 2, len(json_objects) - 1, 2): + if json_objects[j]["kind"] != "message": + break + if json_objects[j]["level"] != "progress": + break + if json_objects[j]["contents"]["stage"] != "full-buffer-fragment-started": + break + if json_objects[j + 1]["contents"]["stage"] != "full-buffer-fragment-lax-ok": + break + num_successes = num_successes + 1 + + if num_successes != num_decls: + print(f"Number of successes {num_successes} does not match number of declarations {num_decls}") + return False + + # The next sequence of messages have status success and query-id "2" + # Check all of these messages for the correct kind and status and stop + # when the first message with a different kind or status is found + for k in range(j + 1, len(json_objects)): + if json_objects[k]["kind"] != "response": + break + if json_objects[k]["status"] != "success": + break + if json_objects[k]["query-id"] != "2": + break + + # check that k is the index of the last message + if k != len(json_objects) - 1: + print(f"Unexpected last message at index {k}, contents = {json_objects[k]}") + return False + + # The last message has query-id "3" and status "success" + if json_objects[k]["query-id"] != "3": + print("Last message does not have query-id 3") + return False + + return True + +# A function to test fstar on a single file +def test_file(file): + # prepend the path ../../ulib to the file + filepath = os.path.join("../../ulib", file) + print(f"Testing {filepath}") + # flush stdout + sys.stdout.flush() + with open(filepath, "r") as f: + contents = f.read() + # Escape the contents of the file for JSON + contents = json.dumps(contents) + # print(contents) + # Format the contents of the file into a request for the F* ide + # The first line is a JSON object initializing the ide + # The second line is a JSON object containing the contents of the file + # The third line is an exit command + request = f'{{"query-id":"1", "query": "vfs-add", "args":{{"filename":null, "contents": {contents}}}}}\n{{"query-id":"2", "query": "full-buffer", "args":{{"kind": "full-with-symbols", "code": {contents}, "line":1, "column":0}}}}\n{{"query-id":"3", "query": "exit", "args":{{}}}}\n' + # print the request to the console for debugging + #print(request) + # Run fstar on the file with the request as stdin + p = subprocess.run([fstar, "--admit_smt_queries", "true", "--ide", file], input=request, encoding="utf-8", stdout=subprocess.PIPE, stderr=subprocess.PIPE) + # Read the response from stdout + response = p.stdout + # Print the response to the console for debugging + # print(response) + # Check that fstar exited with code 0 + if p.returncode != 0: + print("F* returned non-zero exit code") + print(p.stderr.read()) + print(p.stdout.read()) + return False + # Parse the response into a list of lines + lines = response.splitlines() + # Print the number of lines in the response + #print(f"Response has {len(lines)} lines") + # Validate the response + return validate_response(response) + +# List all files in ../../ulib +files = os.listdir("../../ulib") +# Filter the list to only include .fst and .fsti files +files = [f for f in files if f.endswith(".fst") or f.endswith(".fsti")] +succeeded = True +# For each file files, test fstar on the file +for file in files: + # If the test fails, exit with code 1 + if not test_file(file): + print(f"Failed test on {file}") + succeeded = False + +if succeeded: + print("All tests passed") + sys.exit(0) +else: + print("Some tests failed") + sys.exit(1) \ No newline at end of file From 591c35da566f5b86c3da5315141b6f20a55a8cdc Mon Sep 17 00:00:00 2001 From: Nikhil Swamy Date: Fri, 24 Mar 2023 17:19:10 -0700 Subject: [PATCH 41/48] update test case --- tests/ide/Makefile | 3 ++- tests/ide/emacs/FullBuffer.fst | 3 --- tests/ide/emacs/FullBuffer.full-buffer.in | 2 -- 3 files changed, 2 insertions(+), 6 deletions(-) delete mode 100755 tests/ide/emacs/FullBuffer.fst delete mode 100755 tests/ide/emacs/FullBuffer.full-buffer.in diff --git a/tests/ide/Makefile b/tests/ide/Makefile index a18d2ed5b6b..036bdd97ac9 100644 --- a/tests/ide/Makefile +++ b/tests/ide/Makefile @@ -6,11 +6,12 @@ lsp \ emacs all: $(addsuffix .all, $(ALL_TEST_DIRS)) + python test-incremental.py $(FSTAR_HOME)/bin/fstar.exe + clean: $(addsuffix .clean, $(ALL_TEST_DIRS)) %.all: % +$(MAKE) -C $^ all - python test-incremental.py %.clean: % +$(MAKE) -C $^ clean diff --git a/tests/ide/emacs/FullBuffer.fst b/tests/ide/emacs/FullBuffer.fst deleted file mode 100755 index eaa42fcb598..00000000000 --- a/tests/ide/emacs/FullBuffer.fst +++ /dev/null @@ -1,3 +0,0 @@ -module FullBuffer -let f x = match x with | Some x -> true | None -> false -let test y = if Some? y then f y else true diff --git a/tests/ide/emacs/FullBuffer.full-buffer.in b/tests/ide/emacs/FullBuffer.full-buffer.in deleted file mode 100755 index 497765a9d64..00000000000 --- a/tests/ide/emacs/FullBuffer.full-buffer.in +++ /dev/null @@ -1,2 +0,0 @@ -{"query-id":"1","query":"vfs-add","args":{"filename":null,"contents":"module FullBuffer\nlet f x = match x with | Some x -> true | None -> false\nlet test y = if Some? y then f y else true\n"}} -{"query-id":"2","query":"full-buffer","args":{"kind":"full","code":"module FullBuffer\nlet f x = match x with | Some x -> true | None -> false\nlet test y = if Some? y then f y else true\nlet test2 : bool = 0\n","line":1,"column":0}} From 55bf6770987f856a781f7f152ed9db4d920bd259 Mon Sep 17 00:00:00 2001 From: Nikhil Swamy Date: Sun, 26 Mar 2023 13:18:45 -0700 Subject: [PATCH 42/48] Record any warnings raised on a checked fragment so it can be replayed when cached for the IDE; fix incremental parsing of pragmas; strengthen test-interactive.py --- ocaml/fstar-lib/FStar_Parser_ParseIt.ml | 6 +- .../generated/FStar_Interactive_Ide.ml | 78 +++++++------ .../generated/FStar_Interactive_Ide_Types.ml | 22 ++-- .../FStar_Interactive_Incremental.ml | 109 ++++++++++-------- .../generated/FStar_Interactive_PushHelper.ml | 70 +++++++---- ocaml/fstar-lib/generated/FStar_Parser_AST.ml | 2 + .../generated/FStar_Parser_AST_Util.ml | 3 +- src/fstar/FStar.Interactive.Ide.Types.fst | 13 ++- src/fstar/FStar.Interactive.Ide.fst | 17 ++- src/fstar/FStar.Interactive.Incremental.fst | 48 ++++---- src/fstar/FStar.Interactive.Incremental.fsti | 6 +- src/fstar/FStar.Interactive.PushHelper.fst | 22 +++- src/fstar/FStar.Interactive.PushHelper.fsti | 5 +- src/parser/FStar.Parser.AST.Util.fst | 2 +- src/parser/FStar.Parser.AST.fst | 2 + tests/ide/test-incremental.py | 68 +++++++++-- 16 files changed, 300 insertions(+), 173 deletions(-) diff --git a/ocaml/fstar-lib/FStar_Parser_ParseIt.ml b/ocaml/fstar-lib/FStar_Parser_ParseIt.ml index f3692a4dc05..ae2c411c67c 100644 --- a/ocaml/fstar-lib/FStar_Parser_ParseIt.ml +++ b/ocaml/fstar-lib/FStar_Parser_ParseIt.ml @@ -267,6 +267,7 @@ let parse fn = ) in let raw_contents = contents_at d.drange in + (* if FStar_Options.debug_any() then ( FStar_Compiler_Util.print4 "Parsed decl@%s=%s\nRaw contents@%s=%s\n" @@ -275,6 +276,7 @@ let parse fn = (FStar_Compiler_Range.string_of_def_range raw_contents.range) raw_contents.code ); + *) parse ((d, raw_contents)::decls) | Inr err -> List.rev decls, Some err in @@ -294,9 +296,7 @@ let parse fn = | Filename _ | Toplevel _ -> begin let fileOrFragment = - if FStar_Options.debug_any() - then parse_incremental_fragment() - else MenhirLib.Convert.Simplified.traditional2revised FStar_Parser_Parse.inputFragment lexer + MenhirLib.Convert.Simplified.traditional2revised FStar_Parser_Parse.inputFragment lexer in let frags = match fileOrFragment with | FStar_Pervasives.Inl modul -> diff --git a/ocaml/fstar-lib/generated/FStar_Interactive_Ide.ml b/ocaml/fstar-lib/generated/FStar_Interactive_Ide.ml index 16067aca6d0..bf4e5598fbd 100644 --- a/ocaml/fstar-lib/generated/FStar_Interactive_Ide.ml +++ b/ocaml/fstar-lib/generated/FStar_Interactive_Ide.ml @@ -72,7 +72,7 @@ let (nothing_left_to_pop : st.FStar_Interactive_Ide_Types.repl_deps_stack) let (run_repl_transaction : FStar_Interactive_Ide_Types.repl_state -> - FStar_Interactive_Ide_Types.push_kind -> + FStar_Interactive_Ide_Types.push_kind FStar_Pervasives_Native.option -> Prims.bool -> FStar_Interactive_Ide_Types.repl_task -> (Prims.bool * FStar_Interactive_Ide_Types.repl_state)) @@ -226,7 +226,9 @@ let (run_repl_ld_transactions : then FStar_Interactive_Ide_Types.LaxCheck else FStar_Interactive_Ide_Types.FullCheck in let uu___3 = - run_repl_transaction st1 push_kind false timestamped_task in + run_repl_transaction st1 + (FStar_Pervasives_Native.Some push_kind) false + timestamped_task in match uu___3 with | (success, st2) -> if success @@ -358,8 +360,8 @@ let (unpack_interactive_query : let parse_full_buffer_kind uu___1 = match uu___1 with | "full" -> FStar_Interactive_Ide_Types.Full - | "full-with-symbols" -> - FStar_Interactive_Ide_Types.FullBufferWithSymbols + | "lax-with-symbols" -> + FStar_Interactive_Ide_Types.LaxWithSymbols | "cache" -> FStar_Interactive_Ide_Types.Cache | "reload-deps" -> FStar_Interactive_Ide_Types.ReloadDeps | "verify-to-position" -> @@ -1633,8 +1635,10 @@ let (run_push_without_deps : FStar_Pervasives.Inr decl in let st1 = set_nosynth_flag st peek_only in let uu___2 = - run_repl_transaction st1 push_kind peek_only - (FStar_Interactive_Ide_Types.PushFragment (frag, push_kind)) in + run_repl_transaction st1 + (FStar_Pervasives_Native.Some push_kind) peek_only + (FStar_Interactive_Ide_Types.PushFragment + (frag, push_kind, [])) in match uu___2 with | (success, st2) -> let st3 = set_nosynth_flag st2 false in @@ -1667,30 +1671,35 @@ let (run_push_without_deps : (FStar_Compiler_List.map FStar_Interactive_Ide_Types.json_of_issue) in FStar_Compiler_Util.JsonList uu___4 in - let st4 = - if success - then - { - FStar_Interactive_Ide_Types.repl_line = line; - FStar_Interactive_Ide_Types.repl_column = column; - FStar_Interactive_Ide_Types.repl_fname = - (st3.FStar_Interactive_Ide_Types.repl_fname); - FStar_Interactive_Ide_Types.repl_deps_stack = - (st3.FStar_Interactive_Ide_Types.repl_deps_stack); - FStar_Interactive_Ide_Types.repl_curmod = - (st3.FStar_Interactive_Ide_Types.repl_curmod); - FStar_Interactive_Ide_Types.repl_env = - (st3.FStar_Interactive_Ide_Types.repl_env); - FStar_Interactive_Ide_Types.repl_stdin = - (st3.FStar_Interactive_Ide_Types.repl_stdin); - FStar_Interactive_Ide_Types.repl_names = - (st3.FStar_Interactive_Ide_Types.repl_names); - FStar_Interactive_Ide_Types.repl_buffered_input_queries - = - (st3.FStar_Interactive_Ide_Types.repl_buffered_input_queries) - } - else st3 in - ((status, json_errors), (FStar_Pervasives.Inl st4)))))) + (match (errs, status) with + | (uu___5::uu___6, FStar_Interactive_Ide_Types.QueryOK) -> + FStar_Interactive_PushHelper.add_issues_to_push_fragment + [json_errors] + | uu___5 -> ()); + (let st4 = + if success + then + { + FStar_Interactive_Ide_Types.repl_line = line; + FStar_Interactive_Ide_Types.repl_column = column; + FStar_Interactive_Ide_Types.repl_fname = + (st3.FStar_Interactive_Ide_Types.repl_fname); + FStar_Interactive_Ide_Types.repl_deps_stack = + (st3.FStar_Interactive_Ide_Types.repl_deps_stack); + FStar_Interactive_Ide_Types.repl_curmod = + (st3.FStar_Interactive_Ide_Types.repl_curmod); + FStar_Interactive_Ide_Types.repl_env = + (st3.FStar_Interactive_Ide_Types.repl_env); + FStar_Interactive_Ide_Types.repl_stdin = + (st3.FStar_Interactive_Ide_Types.repl_stdin); + FStar_Interactive_Ide_Types.repl_names = + (st3.FStar_Interactive_Ide_Types.repl_names); + FStar_Interactive_Ide_Types.repl_buffered_input_queries + = + (st3.FStar_Interactive_Ide_Types.repl_buffered_input_queries) + } + else st3 in + ((status, json_errors), (FStar_Pervasives.Inl st4))))))) let (run_push_with_deps : FStar_Interactive_Ide_Types.repl_state -> FStar_Interactive_Ide_Types.push_query -> @@ -2091,7 +2100,8 @@ let run_and_rewind : fun task -> let st1 = FStar_Interactive_PushHelper.push_repl "run_and_rewind" - FStar_Interactive_Ide_Types.FullCheck + (FStar_Pervasives_Native.Some + FStar_Interactive_Ide_Types.FullCheck) FStar_Interactive_Ide_Types.Noop st in let results = try @@ -2712,11 +2722,13 @@ let rec (run_query : | FStar_Interactive_Ide_Types.Pop -> let uu___ = run_pop st in as_json_list uu___ | FStar_Interactive_Ide_Types.FullBuffer (code, full) -> - let queries = + let uu___ = FStar_Interactive_Incremental.run_full_buffer st q.FStar_Interactive_Ide_Types.qid code full write_full_buffer_fragment_progress in - fold_query validate_and_run_query queries st [] + (match uu___ with + | (queries, issues) -> + fold_query validate_and_run_query queries st issues) | FStar_Interactive_Ide_Types.AutoComplete (search_term1, context) -> let uu___ = run_autocomplete st search_term1 context in as_json_list uu___ diff --git a/ocaml/fstar-lib/generated/FStar_Interactive_Ide_Types.ml b/ocaml/fstar-lib/generated/FStar_Interactive_Ide_Types.ml index 2ac131e733b..0e2fbf3e670 100644 --- a/ocaml/fstar-lib/generated/FStar_Interactive_Ide_Types.ml +++ b/ocaml/fstar-lib/generated/FStar_Interactive_Ide_Types.ml @@ -117,7 +117,8 @@ type repl_task = | LDSingle of timed_fname | LDInterfaceOfCurrentFile of timed_fname | PushFragment of ((FStar_Parser_ParseIt.input_frag, FStar_Parser_AST.decl) - FStar_Pervasives.either * push_kind) + FStar_Pervasives.either * push_kind * FStar_Compiler_Util.json Prims.list) + | Noop let (uu___is_LDInterleaved : repl_task -> Prims.bool) = fun projectee -> @@ -143,23 +144,23 @@ let (uu___is_PushFragment : repl_task -> Prims.bool) = let (__proj__PushFragment__item___0 : repl_task -> ((FStar_Parser_ParseIt.input_frag, FStar_Parser_AST.decl) - FStar_Pervasives.either * push_kind)) + FStar_Pervasives.either * push_kind * FStar_Compiler_Util.json + Prims.list)) = fun projectee -> match projectee with | PushFragment _0 -> _0 let (uu___is_Noop : repl_task -> Prims.bool) = fun projectee -> match projectee with | Noop -> true | uu___ -> false type full_buffer_request_kind = | Full - | FullBufferWithSymbols + | LaxWithSymbols | Cache | ReloadDeps | VerifyToPosition of position | LaxToPosition of position let (uu___is_Full : full_buffer_request_kind -> Prims.bool) = fun projectee -> match projectee with | Full -> true | uu___ -> false -let (uu___is_FullBufferWithSymbols : full_buffer_request_kind -> Prims.bool) - = +let (uu___is_LaxWithSymbols : full_buffer_request_kind -> Prims.bool) = fun projectee -> - match projectee with | FullBufferWithSymbols -> true | uu___ -> false + match projectee with | LaxWithSymbols -> true | uu___ -> false let (uu___is_Cache : full_buffer_request_kind -> Prims.bool) = fun projectee -> match projectee with | Cache -> true | uu___ -> false let (uu___is_ReloadDeps : full_buffer_request_kind -> Prims.bool) = @@ -402,12 +403,12 @@ let (string_of_repl_task : repl_task -> Prims.string) = | LDInterfaceOfCurrentFile intf -> let uu___1 = string_of_timed_fname intf in FStar_Compiler_Util.format1 "LDInterfaceOfCurrentFile %s" uu___1 - | PushFragment (FStar_Pervasives.Inl frag, uu___1) -> + | PushFragment (FStar_Pervasives.Inl frag, uu___1, uu___2) -> FStar_Compiler_Util.format1 "PushFragment { code = %s }" frag.FStar_Parser_ParseIt.frag_text - | PushFragment (FStar_Pervasives.Inr d, uu___1) -> - let uu___2 = FStar_Parser_AST.decl_to_string d in - FStar_Compiler_Util.format1 "PushFragment { decl = %s }" uu___2 + | PushFragment (FStar_Pervasives.Inr d, uu___1, uu___2) -> + let uu___3 = FStar_Parser_AST.decl_to_string d in + FStar_Compiler_Util.format1 "PushFragment { decl = %s }" uu___3 | Noop -> "Noop {}" let (string_of_repl_stack_entry : repl_stack_entry_t -> Prims.string) = fun uu___ -> @@ -519,6 +520,7 @@ let (query_needs_current_module : query' -> Prims.bool) = | FullBuffer uu___1 -> false | Callback uu___1 -> false | Format uu___1 -> false + | Cancel uu___1 -> false | Push uu___1 -> true | AutoComplete uu___1 -> true | Lookup uu___1 -> true diff --git a/ocaml/fstar-lib/generated/FStar_Interactive_Incremental.ml b/ocaml/fstar-lib/generated/FStar_Interactive_Incremental.ml index 45b85464893..f0f0ca5589d 100644 --- a/ocaml/fstar-lib/generated/FStar_Interactive_Incremental.ml +++ b/ocaml/fstar-lib/generated/FStar_Interactive_Incremental.ml @@ -204,7 +204,8 @@ let (inspect_repl_stack : FStar_Interactive_Ide_Types.push_kind -> Prims.bool -> (fragment_progress -> unit) -> - FStar_Interactive_Ide_Types.query Prims.list qst) + (FStar_Interactive_Ide_Types.query Prims.list * + FStar_Compiler_Util.json Prims.list) qst) = fun s -> fun ds -> @@ -226,30 +227,31 @@ let (inspect_repl_stack : match uu___ with | FStar_Pervasives_Native.None -> let uu___1 = push_decls1 ds in - op_let_Bang uu___1 (fun ds1 -> return ds1) + op_let_Bang uu___1 (fun ds1 -> return (ds1, [])) | FStar_Pervasives_Native.Some (prefix, first_push, rest) -> let entries1 = first_push :: rest in let repl_task1 uu___1 = match uu___1 with | (uu___2, (p, uu___3)) -> p in - let rec matching_prefix entries2 ds1 = + let rec matching_prefix accum entries2 ds1 = match (entries2, ds1) with - | ([], []) -> return [] + | ([], []) -> return ([], accum) | (e::entries3, d::ds2) -> (match repl_task1 e with | FStar_Interactive_Ide_Types.Noop -> - matching_prefix entries3 (d :: ds2) + matching_prefix accum entries3 (d :: ds2) | FStar_Interactive_Ide_Types.PushFragment - (FStar_Pervasives.Inl frag, uu___1) -> - let uu___2 = pop_entries (e :: entries3) in - op_let_Bang uu___2 + (FStar_Pervasives.Inl frag, uu___1, uu___2) -> + let uu___3 = pop_entries (e :: entries3) in + op_let_Bang uu___3 (fun pops -> - let uu___3 = push_decls1 (d :: ds2) in - op_let_Bang uu___3 + let uu___4 = push_decls1 (d :: ds2) in + op_let_Bang uu___4 (fun pushes -> return - (FStar_Compiler_List.op_At pops pushes))) + ((FStar_Compiler_List.op_At pops + pushes), accum))) | FStar_Interactive_Ide_Types.PushFragment - (FStar_Pervasives.Inr d', pk) -> + (FStar_Pervasives.Inr d', pk, issues) -> let uu___1 = FStar_Parser_AST_Util.eq_decl (FStar_Pervasives_Native.fst d) d' in @@ -260,7 +262,9 @@ let (inspect_repl_stack : | (d1, s1) -> (write_full_buffer_fragment_progress (FragmentSuccess (d1, s1, pk)); - matching_prefix entries3 ds2)) + matching_prefix + (FStar_Compiler_List.op_At issues accum) + entries3 ds2)) else (let uu___3 = pop_entries (e :: entries3) in op_let_Bang uu___3 @@ -269,15 +273,16 @@ let (inspect_repl_stack : op_let_Bang uu___4 (fun pushes -> return - (FStar_Compiler_List.op_At pops - pushes))))) + ((FStar_Compiler_List.op_At pops + pushes), accum))))) | ([], ds2) -> let uu___1 = push_decls1 ds2 in - op_let_Bang uu___1 (fun pushes -> return pushes) + op_let_Bang uu___1 + (fun pushes -> return (pushes, accum)) | (es, []) -> let uu___1 = pop_entries es in - op_let_Bang uu___1 (fun pops -> return pops) in - matching_prefix entries1 ds + op_let_Bang uu___1 (fun pops -> return (pops, accum)) in + matching_prefix [] entries1 ds let reload_deps : 'uuuuu 'uuuuu1 . ('uuuuu * (FStar_Interactive_Ide_Types.repl_task * 'uuuuu1)) Prims.list @@ -354,7 +359,8 @@ let (run_full_buffer : Prims.string -> FStar_Interactive_Ide_Types.full_buffer_request_kind -> (fragment_progress -> unit) -> - FStar_Interactive_Ide_Types.query Prims.list) + (FStar_Interactive_Ide_Types.query Prims.list * + FStar_Compiler_Util.json Prims.list)) = fun st -> fun qid1 -> @@ -396,8 +402,7 @@ let (run_full_buffer : start_line <= line) decls | uu___ -> decls in let with_symbols = - request_type = - FStar_Interactive_Ide_Types.FullBufferWithSymbols in + request_type = FStar_Interactive_Ide_Types.LaxWithSymbols in let qs = match parse_result with | FStar_Parser_ParseIt.IncrementalFragment @@ -425,8 +430,8 @@ let (run_full_buffer : op_let_Bang uu___5 (fun push_mod -> return - (FStar_Compiler_List.op_At queries - push_mod))) in + ((FStar_Compiler_List.op_At queries + push_mod), []))) in run_qst uu___3 qid1 | uu___2 -> let decls1 = filter_decls decls in @@ -434,41 +439,47 @@ let (run_full_buffer : match request_type with | FStar_Interactive_Ide_Types.LaxToPosition uu___3 -> FStar_Interactive_Ide_Types.LaxCheck - | FStar_Interactive_Ide_Types.FullBufferWithSymbols - -> FStar_Interactive_Ide_Types.LaxCheck + | FStar_Interactive_Ide_Types.LaxWithSymbols -> + FStar_Interactive_Ide_Types.LaxCheck | uu___3 -> FStar_Interactive_Ide_Types.FullCheck in - let queries = - let uu___3 = - let uu___4 = + let uu___3 = + let uu___4 = + let uu___5 = FStar_Compiler_Effect.op_Bang FStar_Interactive_PushHelper.repl_stack in - inspect_repl_stack uu___4 decls1 push_kind + inspect_repl_stack uu___5 decls1 push_kind with_symbols write_full_buffer_fragment_progress in - run_qst uu___3 qid1 in - (if request_type = FStar_Interactive_Ide_Types.Full - then log_syntax_issues err_opt - else (); - (let uu___5 = FStar_Options.debug_any () in - if uu___5 - then - let uu___6 = - let uu___7 = - FStar_Compiler_List.map - FStar_Interactive_Ide_Types.query_to_string - queries in - FStar_String.concat "\n" uu___7 in - FStar_Compiler_Util.print1 - "Generating queries\n%s\n" uu___6 - else ()); - if request_type <> FStar_Interactive_Ide_Types.Cache - then queries - else []))) + run_qst uu___4 qid1 in + (match uu___3 with + | (queries, issues) -> + (if + request_type <> + FStar_Interactive_Ide_Types.Cache + then log_syntax_issues err_opt + else (); + (let uu___6 = FStar_Options.debug_any () in + if uu___6 + then + let uu___7 = + let uu___8 = + FStar_Compiler_List.map + FStar_Interactive_Ide_Types.query_to_string + queries in + FStar_String.concat "\n" uu___8 in + FStar_Compiler_Util.print1 + "Generating queries\n%s\n" uu___7 + else ()); + if + request_type <> + FStar_Interactive_Ide_Types.Cache + then (queries, issues) + else ([], issues))))) | FStar_Parser_ParseIt.ParseError err -> (if request_type = FStar_Interactive_Ide_Types.Full then log_syntax_issues (FStar_Pervasives_Native.Some err) else (); - []) + ([], [])) | uu___ -> failwith "Unexpected parse result" in qs let (format_code : diff --git a/ocaml/fstar-lib/generated/FStar_Interactive_PushHelper.ml b/ocaml/fstar-lib/generated/FStar_Interactive_PushHelper.ml index 045eaf632b0..7646f2c00ce 100644 --- a/ocaml/fstar-lib/generated/FStar_Interactive_PushHelper.ml +++ b/ocaml/fstar-lib/generated/FStar_Interactive_PushHelper.ml @@ -251,13 +251,13 @@ let (snapshot_env : | (opt_depth, ()) -> ((ctx_depth, opt_depth), env1)) let (push_repl : Prims.string -> - FStar_Interactive_Ide_Types.push_kind -> + FStar_Interactive_Ide_Types.push_kind FStar_Pervasives_Native.option -> FStar_Interactive_Ide_Types.repl_task -> FStar_Interactive_Ide_Types.repl_state -> FStar_Interactive_Ide_Types.repl_state) = fun msg -> - fun push_kind -> + fun push_kind_opt -> fun task -> fun st -> let uu___ = @@ -268,26 +268,44 @@ let (push_repl : let uu___3 = FStar_Compiler_Effect.op_Bang repl_stack in (depth, (task, st)) :: uu___3 in FStar_Compiler_Effect.op_Colon_Equals repl_stack uu___2); - (let uu___2 = set_check_kind env push_kind in - { - FStar_Interactive_Ide_Types.repl_line = - (st.FStar_Interactive_Ide_Types.repl_line); - FStar_Interactive_Ide_Types.repl_column = - (st.FStar_Interactive_Ide_Types.repl_column); - FStar_Interactive_Ide_Types.repl_fname = - (st.FStar_Interactive_Ide_Types.repl_fname); - FStar_Interactive_Ide_Types.repl_deps_stack = - (st.FStar_Interactive_Ide_Types.repl_deps_stack); - FStar_Interactive_Ide_Types.repl_curmod = - (st.FStar_Interactive_Ide_Types.repl_curmod); - FStar_Interactive_Ide_Types.repl_env = uu___2; - FStar_Interactive_Ide_Types.repl_stdin = - (st.FStar_Interactive_Ide_Types.repl_stdin); - FStar_Interactive_Ide_Types.repl_names = - (st.FStar_Interactive_Ide_Types.repl_names); - FStar_Interactive_Ide_Types.repl_buffered_input_queries = - (st.FStar_Interactive_Ide_Types.repl_buffered_input_queries) - })) + (match push_kind_opt with + | FStar_Pervasives_Native.None -> st + | FStar_Pervasives_Native.Some push_kind -> + let uu___2 = set_check_kind env push_kind in + { + FStar_Interactive_Ide_Types.repl_line = + (st.FStar_Interactive_Ide_Types.repl_line); + FStar_Interactive_Ide_Types.repl_column = + (st.FStar_Interactive_Ide_Types.repl_column); + FStar_Interactive_Ide_Types.repl_fname = + (st.FStar_Interactive_Ide_Types.repl_fname); + FStar_Interactive_Ide_Types.repl_deps_stack = + (st.FStar_Interactive_Ide_Types.repl_deps_stack); + FStar_Interactive_Ide_Types.repl_curmod = + (st.FStar_Interactive_Ide_Types.repl_curmod); + FStar_Interactive_Ide_Types.repl_env = uu___2; + FStar_Interactive_Ide_Types.repl_stdin = + (st.FStar_Interactive_Ide_Types.repl_stdin); + FStar_Interactive_Ide_Types.repl_names = + (st.FStar_Interactive_Ide_Types.repl_names); + FStar_Interactive_Ide_Types.repl_buffered_input_queries + = + (st.FStar_Interactive_Ide_Types.repl_buffered_input_queries) + })) +let (add_issues_to_push_fragment : + FStar_Compiler_Util.json Prims.list -> unit) = + fun issues -> + let uu___ = FStar_Compiler_Effect.op_Bang repl_stack in + match uu___ with + | (depth, + (FStar_Interactive_Ide_Types.PushFragment (frag, push_kind, i), st))::rest + -> + let pf = + FStar_Interactive_Ide_Types.PushFragment + (frag, push_kind, (FStar_Compiler_List.op_At issues i)) in + FStar_Compiler_Effect.op_Colon_Equals repl_stack ((depth, (pf, st)) + :: rest) + | uu___1 -> () let (rollback_env : FStar_TypeChecker_Env.solver_t -> Prims.string -> @@ -368,7 +386,7 @@ let (run_repl_task : FStar_Universal.load_interface_decls env intf.FStar_Interactive_Ide_Types.tf_fname in (curmod, uu___) - | FStar_Interactive_Ide_Types.PushFragment (frag, uu___) -> + | FStar_Interactive_Ide_Types.PushFragment (frag, uu___, uu___1) -> FStar_Universal.tc_one_fragment curmod env frag | FStar_Interactive_Ide_Types.Noop -> (curmod, env) let (query_of_ids : @@ -575,7 +593,9 @@ let (repl_tx : (fun uu___ -> match () with | () -> - let st1 = push_repl "repl_tx" push_kind task st in + let st1 = + push_repl "repl_tx" + (FStar_Pervasives_Native.Some push_kind) task st in let uu___1 = track_name_changes st1.FStar_Interactive_Ide_Types.repl_env in @@ -904,5 +924,5 @@ let (full_lax : } FStar_Interactive_Ide_Types.LaxCheck (FStar_Interactive_Ide_Types.PushFragment ((FStar_Pervasives.Inl frag), - FStar_Interactive_Ide_Types.LaxCheck)) + FStar_Interactive_Ide_Types.LaxCheck, [])) | FStar_Pervasives.Inr st1 -> (FStar_Pervasives_Native.None, st1)) \ No newline at end of file diff --git a/ocaml/fstar-lib/generated/FStar_Parser_AST.ml b/ocaml/fstar-lib/generated/FStar_Parser_AST.ml index 5b0e6267bf5..e648e23afc2 100644 --- a/ocaml/fstar-lib/generated/FStar_Parser_AST.ml +++ b/ocaml/fstar-lib/generated/FStar_Parser_AST.ml @@ -2432,6 +2432,8 @@ let (idents_of_binders : let (decl_syntax_is_delimited : decl -> Prims.bool) = fun d -> match d.d with + | Pragma (ResetOptions (FStar_Pervasives_Native.None)) -> false + | Pragma (PushOptions (FStar_Pervasives_Native.None)) -> false | Pragma uu___ -> true | NewEffect (DefineEffect uu___) -> true | LayeredEffect (DefineEffect uu___) -> true diff --git a/ocaml/fstar-lib/generated/FStar_Parser_AST_Util.ml b/ocaml/fstar-lib/generated/FStar_Parser_AST_Util.ml index 88df1ed74cf..9cc75932ae5 100755 --- a/ocaml/fstar-lib/generated/FStar_Parser_AST_Util.ml +++ b/ocaml/fstar-lib/generated/FStar_Parser_AST_Util.ml @@ -869,7 +869,8 @@ and (lidents_of_pattern : let uu___ = lidents_of_pattern p1 in let uu___1 = (concat_map ()) lidents_of_pattern ps in FStar_Compiler_List.op_At uu___ uu___1 - | FStar_Parser_AST.PatVar (i, uu___, uu___1) -> [] + | FStar_Parser_AST.PatVar (i, uu___, uu___1) -> + let uu___2 = FStar_Ident.lid_of_ids [i] in [uu___2] | FStar_Parser_AST.PatName lid -> [lid] | FStar_Parser_AST.PatTvar (i, uu___, uu___1) -> [] | FStar_Parser_AST.PatList ps -> (concat_map ()) lidents_of_pattern ps diff --git a/src/fstar/FStar.Interactive.Ide.Types.fst b/src/fstar/FStar.Interactive.Ide.Types.fst index b518daa81ca..754e8860f29 100755 --- a/src/fstar/FStar.Interactive.Ide.Types.fst +++ b/src/fstar/FStar.Interactive.Ide.Types.fst @@ -95,12 +95,14 @@ type repl_task = | LDInterleaved of timed_fname * timed_fname (* (interface * implementation) *) | LDSingle of timed_fname (* interface or implementation *) | LDInterfaceOfCurrentFile of timed_fname (* interface *) - | PushFragment of either PI.input_frag FStar.Parser.AST.decl * push_kind (* code fragment *) + | PushFragment of either PI.input_frag FStar.Parser.AST.decl (* code fragment *) + * push_kind (* FullCheck, LaxCheck, SyntaxCheck *) + * list json (* any warnings that were raised while checking this fragment *) | Noop (* Used by compute *) type full_buffer_request_kind = | Full : full_buffer_request_kind - | FullBufferWithSymbols : full_buffer_request_kind + | LaxWithSymbols : full_buffer_request_kind | Cache : full_buffer_request_kind | ReloadDeps : full_buffer_request_kind | VerifyToPosition of position @@ -175,9 +177,9 @@ let string_of_repl_task = function Util.format1 "LDSingle %s" (string_of_timed_fname intf_or_impl) | LDInterfaceOfCurrentFile intf -> Util.format1 "LDInterfaceOfCurrentFile %s" (string_of_timed_fname intf) - | PushFragment (Inl frag, _) -> + | PushFragment (Inl frag, _, _) -> Util.format1 "PushFragment { code = %s }" frag.frag_text - | PushFragment (Inr d, _) -> + | PushFragment (Inr d, _, _) -> Util.format1 "PushFragment { decl = %s }" (FStar.Parser.AST.decl_to_string d) | Noop -> "Noop {}" @@ -262,7 +264,8 @@ let query_to_string q = match q.qq with let query_needs_current_module = function | Exit | DescribeProtocol | DescribeRepl | Segment _ | Pop | Push { push_peek_only = false } | VfsAdd _ - | GenericError _ | ProtocolViolation _ | FullBuffer _ | Callback _ | Format _ -> false + | GenericError _ | ProtocolViolation _ + | FullBuffer _ | Callback _ | Format _ | Cancel _ -> false | Push _ | AutoComplete _ | Lookup _ | Compute _ | Search _ -> true let interactive_protocol_vernum = 2 diff --git a/src/fstar/FStar.Interactive.Ide.fst b/src/fstar/FStar.Interactive.Ide.fst index 6ffc9f61688..87ebf735829 100644 --- a/src/fstar/FStar.Interactive.Ide.fst +++ b/src/fstar/FStar.Interactive.Ide.fst @@ -171,7 +171,7 @@ let run_repl_ld_transactions (st: repl_state) (tasks: list repl_task) Options.restore_cmd_line_options false |> ignore; let timestamped_task = update_task_timestamps task in let push_kind = if Options.lax () then LaxCheck else FullCheck in - let success, st = run_repl_transaction st push_kind false timestamped_task in + let success, st = run_repl_transaction st (Some push_kind) false timestamped_task in if success then aux ({ st with repl_deps_stack = !repl_stack }) tasks [] else Inr st @@ -225,7 +225,7 @@ let unpack_interactive_query json = in let parse_full_buffer_kind = function | "full" -> Full - | "full-with-symbols" -> FullBufferWithSymbols + | "lax-with-symbols" -> LaxWithSymbols | "cache" -> Cache | "reload-deps" -> ReloadDeps | "verify-to-position" -> VerifyToPosition (read_to_position ()) @@ -649,7 +649,7 @@ let run_push_without_deps st query Inr decl in let st = set_nosynth_flag st peek_only in - let success, st = run_repl_transaction st push_kind peek_only (PushFragment (frag, push_kind)) in + let success, st = run_repl_transaction st (Some push_kind) peek_only (PushFragment (frag, push_kind, [])) in let st = set_nosynth_flag st false in let status = if success || peek_only then QueryOK else QueryNOK in @@ -671,6 +671,11 @@ let run_push_without_deps st query | _ -> () in let json_errors = JsonList (errs |> List.map json_of_issue) in + let _ = + match errs, status with + | _::_, QueryOK -> add_issues_to_push_fragment [json_errors] + | _ -> () + in let st = if success then { st with repl_line = line; repl_column = column } else st in ((status, json_errors), Inl st) @@ -799,7 +804,7 @@ let run_autocomplete st search_term context = run_module_autocomplete st search_term modules namespaces let run_and_rewind st sigint_default task = - let st = push_repl "run_and_rewind" FullCheck Noop st in + let st = push_repl "run_and_rewind" (Some FullCheck) Noop st in let results = try Util.with_sigint_handler Util.sigint_raise (fun _ -> Inl <| task st) with | Util.SigInt -> Inl sigint_default @@ -1062,10 +1067,10 @@ let rec run_query st (q: query) : (query_status * list json) * either repl_state | Push pquery -> as_json_list (run_push st pquery) | Pop -> as_json_list (run_pop st) | FullBuffer (code, full) -> - let queries = + let queries, issues = FStar.Interactive.Incremental.run_full_buffer st q.qid code full write_full_buffer_fragment_progress in - fold_query validate_and_run_query queries st [] + fold_query validate_and_run_query queries st issues | AutoComplete (search_term, context) -> as_json_list (run_autocomplete st search_term context) | Lookup (symbol, context, pos_opt, rq_info, symrange) -> diff --git a/src/fstar/FStar.Interactive.Incremental.fst b/src/fstar/FStar.Interactive.Incremental.fst index e4ae4433772..7357bad8221 100755 --- a/src/fstar/FStar.Interactive.Incremental.fst +++ b/src/fstar/FStar.Interactive.Incremental.fst @@ -167,7 +167,7 @@ let inspect_repl_stack (s:repl_stack_t) (push_kind : push_kind) (with_symbols:bool) (write_full_buffer_fragment_progress: fragment_progress -> unit) - : qst (list query) + : qst (list query & list json) = let entries = List.rev s in let push_decls = push_decls push_kind with_symbols write_full_buffer_fragment_progress in match BU.prefix_until @@ -176,46 +176,46 @@ let inspect_repl_stack (s:repl_stack_t) with | None -> let! ds = push_decls ds in - return ds + return (ds, []) | Some (prefix, first_push, rest) -> let entries = first_push :: rest in let repl_task (_, (p, _)) = p in - let rec matching_prefix entries (ds:list (decl & code_fragment)) - : qst (list query) + let rec matching_prefix (accum:list json) entries (ds:list (decl & code_fragment)) + : qst (list query & list json) = match entries, ds with | [], [] -> - return [] + return ([], accum) | e::entries, d::ds -> ( match repl_task e with | Noop -> - matching_prefix entries (d::ds) - | PushFragment (Inl frag, _) -> - let! pops = pop_entries (e::entries) in + matching_prefix accum entries (d::ds) + | PushFragment (Inl frag, _, _) -> + let! pops = pop_entries (e::entries) in let! pushes = push_decls (d::ds) in - return (pops @ pushes) - | PushFragment (Inr d', pk) -> + return (pops @ pushes, accum) + | PushFragment (Inr d', pk, issues) -> if eq_decl (fst d) d' then ( let d, s = d in write_full_buffer_fragment_progress (FragmentSuccess (d, s, pk)); - matching_prefix entries ds + matching_prefix (issues@accum) entries ds ) else let! pops = pop_entries (e::entries) in let! pushes = push_decls (d::ds) in - return (pops @ pushes) + return (pops @ pushes, accum) ) | [], ds -> let! pushes = push_decls ds in - return pushes + return (pushes, accum) | es, [] -> let! pops = pop_entries es in - return pops + return (pops, accum) in - matching_prefix entries ds + matching_prefix [] entries ds (* A reload_deps request just pops away the entire stack of PushFragments. We also push on just the `module A` declaration after popping. That's done below. *) @@ -262,7 +262,7 @@ let run_full_buffer (st:repl_state) (code:string) (request_type:full_buffer_request_kind) (write_full_buffer_fragment_progress: fragment_progress -> unit) - : list query + : list query & list json = let parse_result = parse_code code in let log_syntax_issues err = match err with @@ -283,16 +283,18 @@ let run_full_buffer (st:repl_state) decls | _ -> decls in - let with_symbols = request_type = FullBufferWithSymbols in + let with_symbols = request_type = LaxWithSymbols in let qs = match parse_result with | IncrementalFragment (decls, _, err_opt) -> ( + // This is a diagnostic message that is send to the IDE as an info message + // The script test-incremental.py in tests/ide/ depends on this message BU.print1 "Parsed %s declarations\n" (string_of_int (List.length decls)); match request_type, decls with | ReloadDeps, d::_ -> run_qst (let! queries = reload_deps (!repl_stack) in let! push_mod = push_decl FullCheck with_symbols write_full_buffer_fragment_progress d in - return (queries @ push_mod)) + return (queries @ push_mod, [])) qid | _ -> @@ -300,26 +302,26 @@ let run_full_buffer (st:repl_state) let push_kind = match request_type with | LaxToPosition _ -> LaxCheck - | FullBufferWithSymbols -> LaxCheck + | LaxWithSymbols -> LaxCheck | _ -> FullCheck in - let queries = + let queries, issues = run_qst (inspect_repl_stack (!repl_stack) decls push_kind with_symbols write_full_buffer_fragment_progress) qid in - if request_type = Full then log_syntax_issues err_opt; + if request_type <> Cache then log_syntax_issues err_opt; if Options.debug_any() then ( BU.print1 "Generating queries\n%s\n" (String.concat "\n" (List.map query_to_string queries)) ); - if request_type <> Cache then queries else [] + if request_type <> Cache then (queries, issues) else ([] , issues) ) | ParseError err -> if request_type = Full then log_syntax_issues (Some err); - [] + [], [] | _ -> failwith "Unexpected parse result" in diff --git a/src/fstar/FStar.Interactive.Incremental.fsti b/src/fstar/FStar.Interactive.Incremental.fsti index 35494cdaec7..153fa1ad3e6 100755 --- a/src/fstar/FStar.Interactive.Incremental.fsti +++ b/src/fstar/FStar.Interactive.Incremental.fsti @@ -23,6 +23,7 @@ open FStar.Compiler open FStar.Parser.AST open FStar.Errors open FStar.Interactive.Ide.Types +open FStar.Compiler.Util (* Various kinds of progress messages to print back to the client *) type fragment_progress = @@ -46,10 +47,9 @@ val run_full_buffer (st:repl_state) (code:string) (full:full_buffer_request_kind) (write_full_buffer_fragment_progress: fragment_progress -> unit) - : list query + : list query & list json (* Pretty-print the code for reformatting, or return a syntax error *) val format_code (st:repl_state) (code:string) - : either string (list issue) - + : either string (list issue) \ No newline at end of file diff --git a/src/fstar/FStar.Interactive.PushHelper.fst b/src/fstar/FStar.Interactive.PushHelper.fst index 6538435d0e5..5e59b71b2fa 100644 --- a/src/fstar/FStar.Interactive.PushHelper.fst +++ b/src/fstar/FStar.Interactive.PushHelper.fst @@ -106,10 +106,22 @@ let snapshot_env env msg : repl_depth_t * env_t = let opt_depth, () = Options.snapshot () in (ctx_depth, opt_depth), env -let push_repl msg push_kind task st = +let push_repl msg push_kind_opt task st = let depth, env = snapshot_env st.repl_env msg in repl_stack := (depth, (task, st)) :: !repl_stack; - { st with repl_env = set_check_kind env push_kind } // repl_env is the only mutable part of st + match push_kind_opt with + | None -> st + | Some push_kind -> + { st with repl_env = set_check_kind env push_kind } // repl_env is the only mutable part of st + +(* Record the issues that were raised by the last push *) +let add_issues_to_push_fragment (issues: list json) = + match !repl_stack with + | (depth, (PushFragment(frag, push_kind, i), st))::rest -> ( + let pf = PushFragment(frag, push_kind, issues @ i) in + repl_stack := (depth, (pf, st)) :: rest + ) + | _ -> () (** Revert to a previous checkpoint. @@ -166,7 +178,7 @@ let run_repl_task (curmod: optmod_t) (env: env_t) (task: repl_task) : optmod_t * curmod, tc_one env None intf_or_impl.tf_fname | LDInterfaceOfCurrentFile intf -> curmod, Universal.load_interface_decls env intf.tf_fname - | PushFragment (frag, _) -> + | PushFragment (frag, _, _) -> tc_one_fragment curmod env frag | Noop -> curmod, env @@ -251,7 +263,7 @@ let track_name_changes (env: env_t) // variant of run_repl_transaction in IDE let repl_tx st push_kind task = try - let st = push_repl "repl_tx" push_kind task st in + let st = push_repl "repl_tx" (Some push_kind) task st in let env, finish_name_tracking = track_name_changes st.repl_env in // begin name tracking let curmod, env = run_repl_task st.repl_curmod env task in let st = { st with repl_curmod = curmod; repl_env = env } in @@ -371,5 +383,5 @@ let full_lax text st = match ld_deps st with | Inl (st, deps) -> let names = add_module_completions st.repl_fname deps st.repl_names in - repl_tx ({ st with repl_names = names }) LaxCheck (PushFragment (Inl frag, LaxCheck)) + repl_tx ({ st with repl_names = names }) LaxCheck (PushFragment (Inl frag, LaxCheck, [])) | Inr st -> None, st diff --git a/src/fstar/FStar.Interactive.PushHelper.fsti b/src/fstar/FStar.Interactive.PushHelper.fsti index 49f4f1b0167..20ea099d8a7 100644 --- a/src/fstar/FStar.Interactive.PushHelper.fsti +++ b/src/fstar/FStar.Interactive.PushHelper.fsti @@ -47,7 +47,8 @@ val repl_stack : ref repl_stack_t val set_check_kind : env_t -> push_kind -> env_t // Push an Pop, directly copied over from IDE -val push_repl : string -> push_kind -> repl_task -> repl_state -> repl_state +val push_repl : string -> option push_kind -> repl_task -> repl_state -> repl_state +val add_issues_to_push_fragment (issues: list json) : unit val pop_repl : string -> repl_state -> repl_state // Factored out from IDE for use by LSP as well @@ -65,4 +66,4 @@ val commit_name_tracking : repl_state -> list name_tracking_event -> repl_state // Lax-check the whole file; used on didOpen and didSave // returns a diagnostic (only on error) along with the repl_state -val full_lax : string -> repl_state -> option assoct * repl_state \ No newline at end of file +val full_lax : string -> repl_state -> option assoct * repl_state diff --git a/src/parser/FStar.Parser.AST.Util.fst b/src/parser/FStar.Parser.AST.Util.fst index 5e3541fb39a..03e868d1836 100755 --- a/src/parser/FStar.Parser.AST.Util.fst +++ b/src/parser/FStar.Parser.AST.Util.fst @@ -617,7 +617,7 @@ and lidents_of_pattern p = | PatWild _ -> [] | PatConst _ -> [] | PatApp (p, ps) -> lidents_of_pattern p @ concat_map lidents_of_pattern ps - | PatVar (i, _, _) -> [] + | PatVar (i, _, _) -> [FStar.Ident.lid_of_ids [i]] | PatName lid -> [lid] | PatTvar (i, _, _) -> [] | PatList ps -> concat_map lidents_of_pattern ps diff --git a/src/parser/FStar.Parser.AST.fst b/src/parser/FStar.Parser.AST.fst index f990571b82b..a1a25206763 100644 --- a/src/parser/FStar.Parser.AST.fst +++ b/src/parser/FStar.Parser.AST.fst @@ -991,6 +991,8 @@ let idents_of_binders bs r = let decl_syntax_is_delimited (d:decl) = match d.d with + | Pragma (ResetOptions None) -> false + | Pragma (PushOptions None) -> false | Pragma _ | NewEffect (DefineEffect _) | LayeredEffect (DefineEffect _) diff --git a/tests/ide/test-incremental.py b/tests/ide/test-incremental.py index 863c8d14eb5..697a2d80688 100644 --- a/tests/ide/test-incremental.py +++ b/tests/ide/test-incremental.py @@ -10,8 +10,45 @@ # The path to the F* executable fstar = sys.argv[1] +# approximating a test to decide if a string is an F* comment +# some ulib files have trailing comments in a variety of styles +def is_comment(str): + if (str.startswith("(*") and str.endswith("*)")): + return True + lines = str.splitlines() + for line in lines: + if (line.startswith("(*") and line.endswith("*)")): + continue + if not line.startswith("//"): + return False + return True + +def get_contents_in_range(file_lines, start_pos, end_pos): + start_line = start_pos[0] + start_col = start_pos[1] + lines = [] + if (end_pos == None): + lines.append(file_lines[start_line - 1][start_col:]) + lines.extend(file_lines[start_line:]) + lines = "\n".join(lines) + return lines + end_line = end_pos[0] + end_col = end_pos[1] + # Skip the prefix of file_lines until start_line + # take all the lines between start_line and end_line + if (start_line < end_line): + lines.append(file_lines[start_line - 1][start_col:]) + lines.extend(file_lines[start_line:end_line - 1]) + lines.append(file_lines[end_line - 1][:end_col]) + elif (start_line == end_line): + lines.append(file_lines[start_line - 1][start_col:end_col]) + # print (f"lines = {lines}") + lines = "\n".join(lines) + return lines + # A function to validate the response from F* interactive mode -def validate_response(response): +def validate_response(response, file_contents): + file_lines = file_contents.splitlines() # parse the each line of the response into a JSON object # if the line is not valid JSON, print an error message and exit # store the JSON objects in a list @@ -93,6 +130,7 @@ def validate_response(response): # The first message in the pair has contents.stage = "full-buffer-fragment-started" # The second message in the pair has contents.stage = "full-buffer-fragment-lax-ok" num_successes = 1 + last_fragment_end = [0, 0] for j in range(i + 2, len(json_objects) - 1, 2): if json_objects[j]["kind"] != "message": break @@ -102,8 +140,24 @@ def validate_response(response): break if json_objects[j + 1]["contents"]["stage"] != "full-buffer-fragment-lax-ok": break + # {"stage":"full-buffer-fragment-lax-ok","ranges":{"fname":"","beg":[16,0],"end":[16,29]},"code-fragment":{"range":{"fname":"","beg":[16,0],"end":[16,29]},"code":"module FStar.Reflection.Arith"} + lax_ok = json_objects[j + 1]["contents"] + code_frag = lax_ok["code-fragment"] + start_pos = code_frag["range"]["beg"] + end_pos = code_frag["range"]["end"] + last_fragment_end = end_pos + lines = get_contents_in_range(file_lines, start_pos, end_pos) + # join the lines together with newlines + if code_frag["code"] != lines: + print(f"Code fragment does not match the code: Expected {code_frag['code']} but got {lines}") + return False num_successes = num_successes + 1 + remaining_lines = get_contents_in_range(file_lines, last_fragment_end, None) + remaining_lines = remaining_lines.strip() + if remaining_lines != "" and not (is_comment(remaining_lines)): + print(f"Remaining lines are not empty: {remaining_lines}") + return False if num_successes != num_decls: print(f"Number of successes {num_successes} does not match number of declarations {num_decls}") return False @@ -141,13 +195,13 @@ def test_file(file): with open(filepath, "r") as f: contents = f.read() # Escape the contents of the file for JSON - contents = json.dumps(contents) + json_contents = json.dumps(contents) # print(contents) # Format the contents of the file into a request for the F* ide # The first line is a JSON object initializing the ide # The second line is a JSON object containing the contents of the file # The third line is an exit command - request = f'{{"query-id":"1", "query": "vfs-add", "args":{{"filename":null, "contents": {contents}}}}}\n{{"query-id":"2", "query": "full-buffer", "args":{{"kind": "full-with-symbols", "code": {contents}, "line":1, "column":0}}}}\n{{"query-id":"3", "query": "exit", "args":{{}}}}\n' + request = f'{{"query-id":"1", "query": "vfs-add", "args":{{"filename":null, "contents": {json_contents}}}}}\n{{"query-id":"2", "query": "full-buffer", "args":{{"kind": "lax-with-symbols", "code": {json_contents}, "line":1, "column":0}}}}\n{{"query-id":"3", "query": "exit", "args":{{}}}}\n' # print the request to the console for debugging #print(request) # Run fstar on the file with the request as stdin @@ -163,11 +217,11 @@ def test_file(file): print(p.stdout.read()) return False # Parse the response into a list of lines - lines = response.splitlines() + # lines = response.splitlines() # Print the number of lines in the response - #print(f"Response has {len(lines)} lines") + # print(f"Response: {response}") # Validate the response - return validate_response(response) + return validate_response(response, contents) # List all files in ../../ulib files = os.listdir("../../ulib") @@ -178,7 +232,7 @@ def test_file(file): for file in files: # If the test fails, exit with code 1 if not test_file(file): - print(f"Failed test on {file}") + print(f" *** Failed test on {file}") succeeded = False if succeeded: From 8292f2eb11f1de9c06e0d2fccfa0787302178a13 Mon Sep 17 00:00:00 2001 From: Nikhil Swamy Date: Mon, 27 Mar 2023 15:37:26 -0700 Subject: [PATCH 43/48] fix a bug in AST comparison that was ignoring binder qualifiers and attributes --- ocaml/fstar-lib/generated/FStar_Parser_AST_Util.ml | 8 +++++++- src/parser/FStar.Parser.AST.Util.fst | 5 ++++- 2 files changed, 11 insertions(+), 2 deletions(-) diff --git a/ocaml/fstar-lib/generated/FStar_Parser_AST_Util.ml b/ocaml/fstar-lib/generated/FStar_Parser_AST_Util.ml index 9cc75932ae5..7b9a1f2ea3a 100755 --- a/ocaml/fstar-lib/generated/FStar_Parser_AST_Util.ml +++ b/ocaml/fstar-lib/generated/FStar_Parser_AST_Util.ml @@ -360,7 +360,13 @@ and (eq_calc_step : ((eq_term t1 t4) && (eq_term t2 t5)) && (eq_term t3 t6) and (eq_binder : FStar_Parser_AST.binder -> FStar_Parser_AST.binder -> Prims.bool) = - fun b1 -> fun b2 -> eq_binder' b1.FStar_Parser_AST.b b2.FStar_Parser_AST.b + fun b1 -> + fun b2 -> + ((eq_binder' b1.FStar_Parser_AST.b b2.FStar_Parser_AST.b) && + (eq_aqual b1.FStar_Parser_AST.aqual b2.FStar_Parser_AST.aqual)) + && + (eq_list eq_term b1.FStar_Parser_AST.battributes + b2.FStar_Parser_AST.battributes) and (eq_binder' : FStar_Parser_AST.binder' -> FStar_Parser_AST.binder' -> Prims.bool) = fun b1 -> diff --git a/src/parser/FStar.Parser.AST.Util.fst b/src/parser/FStar.Parser.AST.Util.fst index 03e868d1836..390ffd44c0a 100755 --- a/src/parser/FStar.Parser.AST.Util.fst +++ b/src/parser/FStar.Parser.AST.Util.fst @@ -358,7 +358,10 @@ and eq_calc_step (CalcStep (t1, t2, t3)) (CalcStep (t4, t5, t6)) = eq_term t2 t5 && eq_term t3 t6 -and eq_binder (b1 b2:binder) = eq_binder' b1.b b2.b +and eq_binder (b1 b2:binder) = + eq_binder' b1.b b2.b && + eq_aqual b1.aqual b2.aqual && + eq_list eq_term b1.battributes b2.battributes and eq_binder' (b1 b2:binder') = match b1, b2 with From 462c5317805a940e6d71d2ced1ff58f6d09532e0 Mon Sep 17 00:00:00 2001 From: Nikhil Swamy Date: Thu, 30 Mar 2023 10:52:12 -0700 Subject: [PATCH 44/48] remove some debug info --- .../generated/FStar_Interactive_Ide.ml | 34 +++++++------------ src/fstar/FStar.Interactive.Ide.fst | 5 --- 2 files changed, 12 insertions(+), 27 deletions(-) diff --git a/ocaml/fstar-lib/generated/FStar_Interactive_Ide.ml b/ocaml/fstar-lib/generated/FStar_Interactive_Ide.ml index bf4e5598fbd..c400ce346e5 100644 --- a/ocaml/fstar-lib/generated/FStar_Interactive_Ide.ml +++ b/ocaml/fstar-lib/generated/FStar_Interactive_Ide.ml @@ -2829,28 +2829,18 @@ let (js_repl_init_opts : unit -> unit) = | uu___2 -> ())) let rec (go : FStar_Interactive_Ide_Types.repl_state -> Prims.int) = fun st -> - (let uu___1 = FStar_Options.debug_any () in - if uu___1 - then - let uu___2 = - let uu___3 = - FStar_Compiler_Effect.op_Bang - FStar_Interactive_PushHelper.repl_stack in - FStar_Interactive_Ide_Types.string_of_repl_stack uu___3 in - FStar_Compiler_Util.print1 "Repl stack is:\n%s\n" uu___2 - else ()); - (let uu___1 = read_interactive_query st in - match uu___1 with - | (query, st1) -> - let uu___2 = validate_and_run_query st1 query in - (match uu___2 with - | ((status, responses), state_opt) -> - (FStar_Compiler_List.iter - (write_response query.FStar_Interactive_Ide_Types.qid status) - responses; - (match state_opt with - | FStar_Pervasives.Inl st' -> go st' - | FStar_Pervasives.Inr exitcode -> exitcode)))) + let uu___ = read_interactive_query st in + match uu___ with + | (query, st1) -> + let uu___1 = validate_and_run_query st1 query in + (match uu___1 with + | ((status, responses), state_opt) -> + (FStar_Compiler_List.iter + (write_response query.FStar_Interactive_Ide_Types.qid status) + responses; + (match state_opt with + | FStar_Pervasives.Inl st' -> go st' + | FStar_Pervasives.Inr exitcode -> exitcode))) let (interactive_error_handler : FStar_Errors.error_handler) = let issues = FStar_Compiler_Util.mk_ref [] in let add_one e = diff --git a/src/fstar/FStar.Interactive.Ide.fst b/src/fstar/FStar.Interactive.Ide.fst index 87ebf735829..b801ea28ea2 100644 --- a/src/fstar/FStar.Interactive.Ide.fst +++ b/src/fstar/FStar.Interactive.Ide.fst @@ -1123,11 +1123,6 @@ let js_repl_init_opts () = (** This is the main loop for the desktop version **) let rec go st : int = - if Options.debug_any () - then ( - FStar.Compiler.Util.print1 "Repl stack is:\n%s\n" - (string_of_repl_stack (!repl_stack)) - ); let query, st = read_interactive_query st in let (status, responses), state_opt = validate_and_run_query st query in List.iter (write_response query.qid status) responses; From 1b25ea13b49da69de2e2ffe92f1e1518309f32fa Mon Sep 17 00:00:00 2001 From: Nikhil Swamy Date: Thu, 30 Mar 2023 12:44:30 -0700 Subject: [PATCH 45/48] snap --- .../FStar_Interactive_QueryHelper.ml | 17 ++-- ocaml/fstar-lib/generated/FStar_Parser_AST.ml | 16 +++- .../generated/FStar_Parser_AST_Util.ml | 0 .../fstar-lib/generated/FStar_Parser_Parse.ml | 2 +- .../generated/FStar_Syntax_Resugar.ml | 81 ++++++++++++++++--- .../fstar-tests/generated/FStar_Tests_Pars.ml | 4 +- .../fstar-tests/generated/FStar_Tests_Test.ml | 46 ++++++++--- 7 files changed, 135 insertions(+), 31 deletions(-) mode change 100755 => 100644 ocaml/fstar-lib/generated/FStar_Parser_AST_Util.ml diff --git a/ocaml/fstar-lib/generated/FStar_Interactive_QueryHelper.ml b/ocaml/fstar-lib/generated/FStar_Interactive_QueryHelper.ml index 306556d0ba0..a632261d525 100644 --- a/ocaml/fstar-lib/generated/FStar_Interactive_QueryHelper.ml +++ b/ocaml/fstar-lib/generated/FStar_Interactive_QueryHelper.ml @@ -45,10 +45,17 @@ let (term_to_string : fun t -> with_printed_effect_args (fun uu___ -> FStar_TypeChecker_Normalize.term_to_string tcenv t) -let (sigelt_to_string : FStar_Syntax_Syntax.sigelt -> Prims.string) = - fun se -> - with_printed_effect_args - (fun uu___ -> FStar_Syntax_Print.sigelt_to_string se) +let (sigelt_to_string : + FStar_TypeChecker_Env.env -> FStar_Syntax_Syntax.sigelt -> Prims.string) = + fun tcenv -> + fun se -> + with_printed_effect_args + (fun uu___ -> + let uu___1 = + FStar_Syntax_DsEnv.set_current_module + tcenv.FStar_TypeChecker_Env.dsenv + tcenv.FStar_TypeChecker_Env.curmodule in + FStar_Syntax_Print.sigelt_to_string' uu___1 se) let (symlookup : FStar_TypeChecker_Env.env -> Prims.string -> @@ -85,7 +92,7 @@ let (symlookup : (fun uu___1 -> match uu___1 with | (FStar_Pervasives.Inr (se, uu___2), uu___3) -> - let uu___4 = sigelt_to_string se in + let uu___4 = sigelt_to_string tcenv se in FStar_Pervasives_Native.Some uu___4 | uu___2 -> FStar_Pervasives_Native.None) in let info_at_pos_opt = diff --git a/ocaml/fstar-lib/generated/FStar_Parser_AST.ml b/ocaml/fstar-lib/generated/FStar_Parser_AST.ml index e648e23afc2..5661fa2159c 100644 --- a/ocaml/fstar-lib/generated/FStar_Parser_AST.ml +++ b/ocaml/fstar-lib/generated/FStar_Parser_AST.ml @@ -2300,6 +2300,19 @@ let (id_of_tycon : tycon -> Prims.string) = | TyconRecord (i, uu___1, uu___2, uu___3, uu___4) -> FStar_Ident.string_of_id i | TyconVariant (i, uu___1, uu___2, uu___3) -> FStar_Ident.string_of_id i +let (string_of_pragma : pragma -> Prims.string) = + fun uu___ -> + match uu___ with + | SetOptions s -> FStar_Compiler_Util.format1 "set-options \"%s\"" s + | ResetOptions s -> + FStar_Compiler_Util.format1 "reset-options \"%s\"" + (FStar_Compiler_Util.dflt "" s) + | PushOptions s -> + FStar_Compiler_Util.format1 "push-options \"%s\"" + (FStar_Compiler_Util.dflt "" s) + | PopOptions -> "pop-options" + | RestartSolver -> "restart-solver" + | PrintEffectsGraph -> "print-effects-graph" let (decl_to_string : decl -> Prims.string) = fun d -> match d.d with @@ -2381,7 +2394,8 @@ let (decl_to_string : decl -> Prims.string) = Prims.op_Hat uu___1 uu___2 in Prims.op_Hat "splice[" uu___ | SubEffect uu___ -> "sub_effect" - | Pragma uu___ -> "pragma" + | Pragma p -> + let uu___ = string_of_pragma p in Prims.op_Hat "pragma #" uu___ let (modul_to_string : modul -> Prims.string) = fun m -> match m with diff --git a/ocaml/fstar-lib/generated/FStar_Parser_AST_Util.ml b/ocaml/fstar-lib/generated/FStar_Parser_AST_Util.ml old mode 100755 new mode 100644 diff --git a/ocaml/fstar-lib/generated/FStar_Parser_Parse.ml b/ocaml/fstar-lib/generated/FStar_Parser_Parse.ml index 3887a8d6259..094e2cc28c9 100644 --- a/ocaml/fstar-lib/generated/FStar_Parser_Parse.ml +++ b/ocaml/fstar-lib/generated/FStar_Parser_Parse.ml @@ -10134,7 +10134,7 @@ let attrs = ( Some x ) in let op = ( mk_ident ("let" ^ op, rhs parseState 1) ) in ( let lbs = (op, b)::lbs in mk_term (LetOperator ( List.map (fun (op, (pat, tm)) -> (op, pat, tm)) lbs - , e)) (rhs2 parseState 1 4) Expr + , e)) (rhs2 parseState 1 5) Expr )) # 10126 "parse.ml" : 'noSeqTerm)) diff --git a/ocaml/fstar-lib/generated/FStar_Syntax_Resugar.ml b/ocaml/fstar-lib/generated/FStar_Syntax_Resugar.ml index 20910a9aa24..6b976a653d4 100644 --- a/ocaml/fstar-lib/generated/FStar_Syntax_Resugar.ml +++ b/ocaml/fstar-lib/generated/FStar_Syntax_Resugar.ml @@ -353,13 +353,18 @@ let (may_shorten : FStar_Ident.lident -> Prims.bool) = | "Prims.Cons" -> false | uu___1 -> let uu___2 = is_tuple_constructor_lid lid in Prims.op_Negation uu___2 +let (maybe_shorten_lid : + FStar_Syntax_DsEnv.env -> FStar_Ident.lident -> FStar_Ident.lident) = + fun env -> + fun lid -> + let uu___ = may_shorten lid in + if uu___ then FStar_Syntax_DsEnv.shorten_lid env lid else lid let (maybe_shorten_fv : FStar_Syntax_DsEnv.env -> FStar_Syntax_Syntax.fv -> FStar_Ident.lident) = fun env -> fun fv -> let lid = (fv.FStar_Syntax_Syntax.fv_name).FStar_Syntax_Syntax.v in - let uu___ = may_shorten lid in - if uu___ then FStar_Syntax_DsEnv.shorten_lid env lid else lid + maybe_shorten_lid env lid let (serialize_machine_integer_desc : (FStar_Const.signedness * FStar_Const.width) -> Prims.string) = fun uu___ -> @@ -1814,12 +1819,15 @@ and (resugar_comp' : let uu___2 = let uu___3 = let uu___4 = + maybe_shorten_lid env + c1.FStar_Syntax_Syntax.effect_name in + let uu___5 = FStar_Compiler_List.map (fun t -> (t, FStar_Parser_AST.Nothing)) (FStar_Compiler_List.op_At pre2 (FStar_Compiler_List.op_At (post2 :: decrease) pats2)) in - ((c1.FStar_Syntax_Syntax.effect_name), uu___4) in + (uu___4, uu___5) in FStar_Parser_AST.Construct uu___3 in mk uu___2) else @@ -1854,14 +1862,23 @@ and (resugar_comp' : aux (FStar_Compiler_List.op_At es l) tl | uu___4 -> aux l tl) in let decrease = aux [] c1.FStar_Syntax_Syntax.flags in - mk - (FStar_Parser_AST.Construct - ((c1.FStar_Syntax_Syntax.effect_name), - (FStar_Compiler_List.op_At (result :: decrease) args))) + let uu___3 = + let uu___4 = + let uu___5 = + maybe_shorten_lid env c1.FStar_Syntax_Syntax.effect_name in + (uu___5, + (FStar_Compiler_List.op_At (result :: decrease) args)) in + FStar_Parser_AST.Construct uu___4 in + mk uu___3 else - mk - (FStar_Parser_AST.Construct - ((c1.FStar_Syntax_Syntax.effect_name), [result]))) + (let uu___4 = + let uu___5 = + let uu___6 = + maybe_shorten_lid env + c1.FStar_Syntax_Syntax.effect_name in + (uu___6, [result]) in + FStar_Parser_AST.Construct uu___5 in + mk uu___4)) and (resugar_binder' : FStar_Syntax_DsEnv.env -> FStar_Syntax_Syntax.binder -> @@ -2734,7 +2751,49 @@ let (resugar_sigelt' : (let mk e = FStar_Syntax_Syntax.mk e se.FStar_Syntax_Syntax.sigrng in let dummy = mk FStar_Syntax_Syntax.Tm_unknown in - let desugared_let = mk (FStar_Syntax_Syntax.Tm_let (lbs, dummy)) in + let nopath_lbs uu___3 = + match uu___3 with + | (is_rec, lbs1) -> + let nopath fv = + let uu___4 = + let uu___5 = + let uu___6 = + let uu___7 = FStar_Syntax_Syntax.lid_of_fv fv in + FStar_Ident.ident_of_lid uu___7 in + [uu___6] in + FStar_Ident.lid_of_ids uu___5 in + FStar_Syntax_Syntax.lid_as_fv uu___4 + FStar_Syntax_Syntax.delta_constant + FStar_Pervasives_Native.None in + let lbs2 = + FStar_Compiler_List.map + (fun lb -> + let uu___4 = + let uu___5 = + let uu___6 = + FStar_Compiler_Util.right + lb.FStar_Syntax_Syntax.lbname in + FStar_Compiler_Effect.op_Less_Bar nopath uu___6 in + FStar_Pervasives.Inr uu___5 in + { + FStar_Syntax_Syntax.lbname = uu___4; + FStar_Syntax_Syntax.lbunivs = + (lb.FStar_Syntax_Syntax.lbunivs); + FStar_Syntax_Syntax.lbtyp = + (lb.FStar_Syntax_Syntax.lbtyp); + FStar_Syntax_Syntax.lbeff = + (lb.FStar_Syntax_Syntax.lbeff); + FStar_Syntax_Syntax.lbdef = + (lb.FStar_Syntax_Syntax.lbdef); + FStar_Syntax_Syntax.lbattrs = + (lb.FStar_Syntax_Syntax.lbattrs); + FStar_Syntax_Syntax.lbpos = + (lb.FStar_Syntax_Syntax.lbpos) + }) lbs1 in + (is_rec, lbs2) in + let lbs1 = nopath_lbs lbs in + let desugared_let = + mk (FStar_Syntax_Syntax.Tm_let (lbs1, dummy)) in let t = resugar_term' env desugared_let in match t.FStar_Parser_AST.tm with | FStar_Parser_AST.Let (isrec, lets, uu___3) -> diff --git a/ocaml/fstar-tests/generated/FStar_Tests_Pars.ml b/ocaml/fstar-tests/generated/FStar_Tests_Pars.ml index 4ff36f49ff6..d1b2ddcea28 100644 --- a/ocaml/fstar-tests/generated/FStar_Tests_Pars.ml +++ b/ocaml/fstar-tests/generated/FStar_Tests_Pars.ml @@ -448,7 +448,7 @@ let (tc' : match uu___ with | (tm1, uu___1, g) -> (FStar_TypeChecker_Rel.force_trivial_guard tcenv1 g; - (let tm2 = FStar_Syntax_Subst.deep_compress false tm1 in + (let tm2 = FStar_Syntax_Compress.deep_compress false tm1 in (tm2, tcenv1))) let (tc : Prims.string -> FStar_Syntax_Syntax.term) = fun s -> let uu___ = tc' s in match uu___ with | (tm, uu___1) -> tm @@ -545,7 +545,7 @@ let (tc_term : FStar_Syntax_Syntax.term -> FStar_Syntax_Syntax.term) = match uu___ with | (tm1, uu___1, g) -> (FStar_TypeChecker_Rel.force_trivial_guard tcenv1 g; - (let tm2 = FStar_Syntax_Subst.deep_compress false tm1 in tm2)) + (let tm2 = FStar_Syntax_Compress.deep_compress false tm1 in tm2)) let (pars_and_tc_fragment : Prims.string -> unit) = fun s -> FStar_Options.set_option "trace_error" (FStar_Options.Bool true); diff --git a/ocaml/fstar-tests/generated/FStar_Tests_Test.ml b/ocaml/fstar-tests/generated/FStar_Tests_Test.ml index 08c1ce2e28e..d1d8af0de7b 100644 --- a/ocaml/fstar-tests/generated/FStar_Tests_Test.ml +++ b/ocaml/fstar-tests/generated/FStar_Tests_Test.ml @@ -6,17 +6,41 @@ let main : 'uuuuu 'uuuuu1 . 'uuuuu -> 'uuuuu1 = (fun uu___1 -> match () with | () -> - (FStar_Main.setup_hooks (); - (let uu___4 = FStar_Tests_Pars.init () in - FStar_Compiler_Effect.op_Bar_Greater uu___4 - (fun uu___5 -> ())); - FStar_Tests_Pars.parse_incremental_decls (); - FStar_Tests_Norm.run_all (); - (let uu___7 = FStar_Tests_Unif.run_all () in - if uu___7 - then () - else FStar_Compiler_Effect.exit Prims.int_one); - FStar_Compiler_Effect.exit Prims.int_zero)) () + let uu___2 = FStar_Options.parse_cmd_line () in + (match uu___2 with + | (res, fs) -> + (match res with + | FStar_Getopt.Help -> + (FStar_Compiler_Util.print_string + "F* unit tests. This binary can take the same options as F*, but not all of them are meaningful."; + FStar_Compiler_Effect.exit Prims.int_zero) + | FStar_Getopt.Error msg -> + (FStar_Compiler_Util.print_error msg; + FStar_Compiler_Effect.exit Prims.int_one) + | FStar_Getopt.Empty -> + (FStar_Main.setup_hooks (); + (let uu___5 = FStar_Tests_Pars.init () in + FStar_Compiler_Effect.op_Bar_Greater uu___5 + (fun uu___6 -> ())); + FStar_Tests_Pars.parse_incremental_decls (); + FStar_Tests_Norm.run_all (); + (let uu___8 = FStar_Tests_Unif.run_all () in + if uu___8 + then () + else FStar_Compiler_Effect.exit Prims.int_one); + FStar_Compiler_Effect.exit Prims.int_zero) + | FStar_Getopt.Success -> + (FStar_Main.setup_hooks (); + (let uu___5 = FStar_Tests_Pars.init () in + FStar_Compiler_Effect.op_Bar_Greater uu___5 + (fun uu___6 -> ())); + FStar_Tests_Pars.parse_incremental_decls (); + FStar_Tests_Norm.run_all (); + (let uu___8 = FStar_Tests_Unif.run_all () in + if uu___8 + then () + else FStar_Compiler_Effect.exit Prims.int_one); + FStar_Compiler_Effect.exit Prims.int_zero)))) () with | FStar_Errors.Error (err, msg, r, _ctx) when let uu___2 = FStar_Options.trace_error () in From 9c82c5a22f521a04cea39fda4be7634f9fccf688 Mon Sep 17 00:00:00 2001 From: Nikhil Swamy Date: Fri, 31 Mar 2023 09:36:58 -0700 Subject: [PATCH 46/48] send back digests rather than full code in fragment-ok messages --- ocaml/fstar-lib/generated/FStar_Interactive_Ide.ml | 13 ++++++++++--- src/fstar/FStar.Interactive.Ide.fst | 2 +- 2 files changed, 11 insertions(+), 4 deletions(-) diff --git a/ocaml/fstar-lib/generated/FStar_Interactive_Ide.ml b/ocaml/fstar-lib/generated/FStar_Interactive_Ide.ml index c400ce346e5..4533172884f 100644 --- a/ocaml/fstar-lib/generated/FStar_Interactive_Ide.ml +++ b/ocaml/fstar-lib/generated/FStar_Interactive_Ide.ml @@ -1406,9 +1406,16 @@ let (write_full_buffer_fragment_progress : FStar_Compiler_Range.json_of_def_range cf.FStar_Parser_ParseIt.range in ("range", uu___2) in - [uu___1; - ("code", - (FStar_Compiler_Util.JsonStr (cf.FStar_Parser_ParseIt.code)))] in + let uu___2 = + let uu___3 = + let uu___4 = + let uu___5 = + FStar_Compiler_Util.digest_of_string + cf.FStar_Parser_ParseIt.code in + FStar_Compiler_Util.JsonStr uu___5 in + ("code-digest", uu___4) in + [uu___3] in + uu___1 :: uu___2 in FStar_Compiler_Util.JsonAssoc uu___ in match di with | FStar_Interactive_Incremental.FragmentStarted d -> diff --git a/src/fstar/FStar.Interactive.Ide.fst b/src/fstar/FStar.Interactive.Ide.fst index b801ea28ea2..3d3de0d248c 100644 --- a/src/fstar/FStar.Interactive.Ide.fst +++ b/src/fstar/FStar.Interactive.Ide.fst @@ -598,7 +598,7 @@ let write_full_buffer_fragment_progress (di:Incremental.fragment_progress) = let open FStar.Interactive.Incremental in let json_of_code_fragment (cf:FStar.Parser.ParseIt.code_fragment) = JsonAssoc ["range", json_of_def_range cf.range; - "code", JsonStr cf.code] + "code-digest", JsonStr (BU.digest_of_string cf.code)] in match di with | FragmentStarted d -> From 881f40a099784d14a2d25fe3f9bd7fe2c4428980 Mon Sep 17 00:00:00 2001 From: Nikhil Swamy Date: Fri, 31 Mar 2023 15:13:32 -0700 Subject: [PATCH 47/48] add full-buffer-start and finished messages; separate with_symbols from full/lax requests --- .../generated/FStar_Interactive_Ide.ml | 93 +++--- .../generated/FStar_Interactive_Ide_Types.ml | 11 +- .../FStar_Interactive_Incremental.ml | 293 ++++++++++-------- .../generated/FStar_Interactive_JsonHelper.ml | 5 + src/fstar/FStar.Interactive.Ide.Types.fst | 4 +- src/fstar/FStar.Interactive.Ide.fst | 38 ++- src/fstar/FStar.Interactive.Incremental.fst | 21 +- src/fstar/FStar.Interactive.Incremental.fsti | 3 + src/fstar/FStar.Interactive.JsonHelper.fst | 3 + src/fstar/FStar.Interactive.JsonHelper.fsti | 1 + 10 files changed, 271 insertions(+), 201 deletions(-) diff --git a/ocaml/fstar-lib/generated/FStar_Interactive_Ide.ml b/ocaml/fstar-lib/generated/FStar_Interactive_Ide.ml index 4533172884f..7d21c8f9b3c 100644 --- a/ocaml/fstar-lib/generated/FStar_Interactive_Ide.ml +++ b/ocaml/fstar-lib/generated/FStar_Interactive_Ide.ml @@ -357,20 +357,19 @@ let (unpack_interactive_query : FStar_Compiler_Effect.op_Bar_Greater uu___4 FStar_Interactive_JsonHelper.js_int in ("", uu___2, uu___3) in - let parse_full_buffer_kind uu___1 = - match uu___1 with + let parse_full_buffer_kind kind = + match kind with | "full" -> FStar_Interactive_Ide_Types.Full - | "lax-with-symbols" -> - FStar_Interactive_Ide_Types.LaxWithSymbols + | "lax" -> FStar_Interactive_Ide_Types.Lax | "cache" -> FStar_Interactive_Ide_Types.Cache | "reload-deps" -> FStar_Interactive_Ide_Types.ReloadDeps | "verify-to-position" -> - let uu___2 = read_to_position () in - FStar_Interactive_Ide_Types.VerifyToPosition uu___2 + let uu___1 = read_to_position () in + FStar_Interactive_Ide_Types.VerifyToPosition uu___1 | "lax-to-position" -> - let uu___2 = read_to_position () in - FStar_Interactive_Ide_Types.LaxToPosition uu___2 - | uu___2 -> + let uu___1 = read_to_position () in + FStar_Interactive_Ide_Types.LaxToPosition uu___1 + | uu___1 -> FStar_Compiler_Effect.raise (FStar_Interactive_JsonHelper.InvalidQuery "Invalid full-buffer kind") in @@ -457,7 +456,11 @@ let (unpack_interactive_query : FStar_Compiler_Effect.op_Bar_Greater uu___6 FStar_Interactive_JsonHelper.js_str in parse_full_buffer_kind uu___5 in - (uu___3, uu___4) in + let uu___5 = + let uu___6 = arg "with-symbols" in + FStar_Compiler_Effect.op_Bar_Greater uu___6 + FStar_Interactive_JsonHelper.js_bool in + (uu___3, uu___4, uu___5) in FStar_Interactive_Ide_Types.FullBuffer uu___2 | "autocomplete" -> let uu___2 = @@ -1418,6 +1421,9 @@ let (write_full_buffer_fragment_progress : uu___1 :: uu___2 in FStar_Compiler_Util.JsonAssoc uu___ in match di with + | FStar_Interactive_Incremental.FullBufferStarted -> + write_progress (FStar_Pervasives_Native.Some "full-buffer-started") + [] | FStar_Interactive_Incremental.FragmentStarted d -> let uu___ = let uu___1 = @@ -1484,6 +1490,9 @@ let (write_full_buffer_fragment_progress : FStar_Compiler_Util.JsonList uu___2 in json_of_response qid FStar_Interactive_Ide_Types.QueryNOK uu___1 in FStar_Interactive_JsonHelper.write_json uu___ + | FStar_Interactive_Incremental.FullBufferFinished -> + write_progress (FStar_Pervasives_Native.Some "full-buffer-finished") + [] let (run_push_without_deps : FStar_Interactive_Ide_Types.repl_state -> FStar_Interactive_Ide_Types.push_query -> @@ -2640,30 +2649,29 @@ let rec (fold_query : FStar_Interactive_Ide_Types.query -> run_query_result) -> FStar_Interactive_Ide_Types.query Prims.list -> - FStar_Interactive_Ide_Types.repl_state -> - FStar_Compiler_Util.json Prims.list -> run_query_result) + FStar_Interactive_Ide_Types.repl_state -> run_query_result) = fun f -> fun l -> fun st -> - fun responses -> - match l with - | [] -> - ((FStar_Interactive_Ide_Types.QueryOK, responses), - (FStar_Pervasives.Inl st)) - | q::l1 -> - let uu___ = f st q in - (match uu___ with - | ((status, resp), st') -> - let responses1 = FStar_Compiler_List.op_At responses resp in - (match (status, st') with - | (FStar_Interactive_Ide_Types.QueryOK, - FStar_Pervasives.Inl st1) -> - let st2 = buffer_input_queries st1 in - let uu___1 = maybe_cancel_queries st2 l1 in - (match uu___1 with - | (l2, st3) -> fold_query f l2 st3 responses1) - | uu___1 -> ((status, responses1), st'))) + match l with + | [] -> + ((FStar_Interactive_Ide_Types.QueryOK, []), + (FStar_Pervasives.Inl st)) + | q::l1 -> + let uu___ = f st q in + (match uu___ with + | ((status, responses), st') -> + (FStar_Compiler_List.iter + (write_response q.FStar_Interactive_Ide_Types.qid status) + responses; + (match (status, st') with + | (FStar_Interactive_Ide_Types.QueryOK, + FStar_Pervasives.Inl st1) -> + let st2 = buffer_input_queries st1 in + let uu___2 = maybe_cancel_queries st2 l1 in + (match uu___2 with | (l2, st3) -> fold_query f l2 st3) + | uu___2 -> ((status, []), st')))) let (validate_query : FStar_Interactive_Ide_Types.repl_state -> FStar_Interactive_Ide_Types.query -> FStar_Interactive_Ide_Types.query) @@ -2728,14 +2736,23 @@ let rec (run_query : let uu___ = run_push st pquery in as_json_list uu___ | FStar_Interactive_Ide_Types.Pop -> let uu___ = run_pop st in as_json_list uu___ - | FStar_Interactive_Ide_Types.FullBuffer (code, full) -> - let uu___ = - FStar_Interactive_Incremental.run_full_buffer st - q.FStar_Interactive_Ide_Types.qid code full - write_full_buffer_fragment_progress in - (match uu___ with - | (queries, issues) -> - fold_query validate_and_run_query queries st issues) + | FStar_Interactive_Ide_Types.FullBuffer + (code, full_kind, with_symbols) -> + (write_full_buffer_fragment_progress + FStar_Interactive_Incremental.FullBufferStarted; + (let uu___1 = + FStar_Interactive_Incremental.run_full_buffer st + q.FStar_Interactive_Ide_Types.qid code full_kind with_symbols + write_full_buffer_fragment_progress in + match uu___1 with + | (queries, issues) -> + (FStar_Compiler_List.iter + (write_response q.FStar_Interactive_Ide_Types.qid + FStar_Interactive_Ide_Types.QueryOK) issues; + (let res = fold_query validate_and_run_query queries st in + write_full_buffer_fragment_progress + FStar_Interactive_Incremental.FullBufferFinished; + res)))) | FStar_Interactive_Ide_Types.AutoComplete (search_term1, context) -> let uu___ = run_autocomplete st search_term1 context in as_json_list uu___ diff --git a/ocaml/fstar-lib/generated/FStar_Interactive_Ide_Types.ml b/ocaml/fstar-lib/generated/FStar_Interactive_Ide_Types.ml index 0e2fbf3e670..18bf039094c 100644 --- a/ocaml/fstar-lib/generated/FStar_Interactive_Ide_Types.ml +++ b/ocaml/fstar-lib/generated/FStar_Interactive_Ide_Types.ml @@ -151,16 +151,15 @@ let (uu___is_Noop : repl_task -> Prims.bool) = fun projectee -> match projectee with | Noop -> true | uu___ -> false type full_buffer_request_kind = | Full - | LaxWithSymbols + | Lax | Cache | ReloadDeps | VerifyToPosition of position | LaxToPosition of position let (uu___is_Full : full_buffer_request_kind -> Prims.bool) = fun projectee -> match projectee with | Full -> true | uu___ -> false -let (uu___is_LaxWithSymbols : full_buffer_request_kind -> Prims.bool) = - fun projectee -> - match projectee with | LaxWithSymbols -> true | uu___ -> false +let (uu___is_Lax : full_buffer_request_kind -> Prims.bool) = + fun projectee -> match projectee with | Lax -> true | uu___ -> false let (uu___is_Cache : full_buffer_request_kind -> Prims.bool) = fun projectee -> match projectee with | Cache -> true | uu___ -> false let (uu___is_ReloadDeps : full_buffer_request_kind -> Prims.bool) = @@ -193,7 +192,7 @@ type query' = | Search of Prims.string | GenericError of Prims.string | ProtocolViolation of Prims.string - | FullBuffer of (Prims.string * full_buffer_request_kind) + | FullBuffer of (Prims.string * full_buffer_request_kind * Prims.bool) | Callback of (repl_state -> ((query_status * FStar_Compiler_Util.json Prims.list) * (repl_state, @@ -277,7 +276,7 @@ let (uu___is_FullBuffer : query' -> Prims.bool) = fun projectee -> match projectee with | FullBuffer _0 -> true | uu___ -> false let (__proj__FullBuffer__item___0 : - query' -> (Prims.string * full_buffer_request_kind)) = + query' -> (Prims.string * full_buffer_request_kind * Prims.bool)) = fun projectee -> match projectee with | FullBuffer _0 -> _0 let (uu___is_Callback : query' -> Prims.bool) = fun projectee -> diff --git a/ocaml/fstar-lib/generated/FStar_Interactive_Incremental.ml b/ocaml/fstar-lib/generated/FStar_Interactive_Incremental.ml index f0f0ca5589d..ddb8ce1fb74 100644 --- a/ocaml/fstar-lib/generated/FStar_Interactive_Incremental.ml +++ b/ocaml/fstar-lib/generated/FStar_Interactive_Incremental.ml @@ -1,11 +1,16 @@ open Prims type fragment_progress = + | FullBufferStarted | FragmentStarted of FStar_Parser_AST.decl | FragmentSuccess of (FStar_Parser_AST.decl * FStar_Parser_ParseIt.code_fragment * FStar_Interactive_Ide_Types.push_kind) | FragmentFailed of FStar_Parser_AST.decl | FragmentError of FStar_Errors.issue Prims.list + | FullBufferFinished +let (uu___is_FullBufferStarted : fragment_progress -> Prims.bool) = + fun projectee -> + match projectee with | FullBufferStarted -> true | uu___ -> false let (uu___is_FragmentStarted : fragment_progress -> Prims.bool) = fun projectee -> match projectee with | FragmentStarted _0 -> true | uu___ -> false @@ -32,6 +37,9 @@ let (uu___is_FragmentError : fragment_progress -> Prims.bool) = let (__proj__FragmentError__item___0 : fragment_progress -> FStar_Errors.issue Prims.list) = fun projectee -> match projectee with | FragmentError _0 -> _0 +let (uu___is_FullBufferFinished : fragment_progress -> Prims.bool) = + fun projectee -> + match projectee with | FullBufferFinished -> true | uu___ -> false type qid = (Prims.string * Prims.int) type 'a qst = qid -> ('a * qid) let return : 'a . 'a -> 'a qst = fun x -> fun q -> (x, q) @@ -232,13 +240,13 @@ let (inspect_repl_stack : let entries1 = first_push :: rest in let repl_task1 uu___1 = match uu___1 with | (uu___2, (p, uu___3)) -> p in - let rec matching_prefix accum entries2 ds1 = + let rec matching_prefix accum lookups entries2 ds1 = match (entries2, ds1) with | ([], []) -> return ([], accum) | (e::entries3, d::ds2) -> (match repl_task1 e with | FStar_Interactive_Ide_Types.Noop -> - matching_prefix accum entries3 (d :: ds2) + matching_prefix accum lookups entries3 (d :: ds2) | FStar_Interactive_Ide_Types.PushFragment (FStar_Pervasives.Inl frag, uu___1, uu___2) -> let uu___3 = pop_entries (e :: entries3) in @@ -262,9 +270,21 @@ let (inspect_repl_stack : | (d1, s1) -> (write_full_buffer_fragment_progress (FragmentSuccess (d1, s1, pk)); - matching_prefix - (FStar_Compiler_List.op_At issues accum) - entries3 ds2)) + if with_symbols + then + (let uu___4 = dump_symbols d1 in + op_let_Bang uu___4 + (fun lookups' -> + matching_prefix + (FStar_Compiler_List.op_At + issues accum) + (FStar_Compiler_List.op_At + lookups' lookups) entries3 + ds2)) + else + matching_prefix + (FStar_Compiler_List.op_At issues + accum) lookups entries3 ds2)) else (let uu___3 = pop_entries (e :: entries3) in op_let_Bang uu___3 @@ -274,7 +294,8 @@ let (inspect_repl_stack : (fun pushes -> return ((FStar_Compiler_List.op_At pops - pushes), accum))))) + (FStar_Compiler_List.op_At + lookups pushes)), accum))))) | ([], ds2) -> let uu___1 = push_decls1 ds2 in op_let_Bang uu___1 @@ -282,7 +303,7 @@ let (inspect_repl_stack : | (es, []) -> let uu___1 = pop_entries es in op_let_Bang uu___1 (fun pops -> return (pops, accum)) in - matching_prefix [] entries1 ds + matching_prefix [] [] entries1 ds let reload_deps : 'uuuuu 'uuuuu1 . ('uuuuu * (FStar_Interactive_Ide_Types.repl_task * 'uuuuu1)) Prims.list @@ -358,130 +379,134 @@ let (run_full_buffer : Prims.string -> Prims.string -> FStar_Interactive_Ide_Types.full_buffer_request_kind -> - (fragment_progress -> unit) -> - (FStar_Interactive_Ide_Types.query Prims.list * - FStar_Compiler_Util.json Prims.list)) + Prims.bool -> + (fragment_progress -> unit) -> + (FStar_Interactive_Ide_Types.query Prims.list * + FStar_Compiler_Util.json Prims.list)) = fun st -> fun qid1 -> fun code -> fun request_type -> - fun write_full_buffer_fragment_progress -> - let parse_result = parse_code code in - let log_syntax_issues err = - match err with - | FStar_Pervasives_Native.None -> () - | FStar_Pervasives_Native.Some err1 -> - let issue = syntax_issue err1 in - write_full_buffer_fragment_progress (FragmentError [issue]) in - let filter_decls decls = - match request_type with - | FStar_Interactive_Ide_Types.VerifyToPosition - (uu___, line, _col) -> - FStar_Compiler_List.filter - (fun uu___1 -> - match uu___1 with - | (d, uu___2) -> - let start = - FStar_Compiler_Range.start_of_range - d.FStar_Parser_AST.drange in - let start_line = - FStar_Compiler_Range.line_of_pos start in - start_line <= line) decls - | FStar_Interactive_Ide_Types.LaxToPosition (uu___, line, _col) - -> - FStar_Compiler_List.filter - (fun uu___1 -> - match uu___1 with - | (d, uu___2) -> - let start = - FStar_Compiler_Range.start_of_range - d.FStar_Parser_AST.drange in - let start_line = - FStar_Compiler_Range.line_of_pos start in - start_line <= line) decls - | uu___ -> decls in - let with_symbols = - request_type = FStar_Interactive_Ide_Types.LaxWithSymbols in - let qs = - match parse_result with - | FStar_Parser_ParseIt.IncrementalFragment - (decls, uu___, err_opt) -> - ((let uu___2 = - FStar_Compiler_Util.string_of_int - (FStar_Compiler_List.length decls) in - FStar_Compiler_Util.print1 "Parsed %s declarations\n" - uu___2); - (match (request_type, decls) with - | (FStar_Interactive_Ide_Types.ReloadDeps, d::uu___2) -> - let uu___3 = - let uu___4 = - let uu___5 = - FStar_Compiler_Effect.op_Bang - FStar_Interactive_PushHelper.repl_stack in - reload_deps uu___5 in - op_let_Bang uu___4 - (fun queries -> - let uu___5 = - push_decl - FStar_Interactive_Ide_Types.FullCheck - with_symbols - write_full_buffer_fragment_progress d in - op_let_Bang uu___5 - (fun push_mod -> - return - ((FStar_Compiler_List.op_At queries - push_mod), []))) in - run_qst uu___3 qid1 - | uu___2 -> - let decls1 = filter_decls decls in - let push_kind = - match request_type with - | FStar_Interactive_Ide_Types.LaxToPosition uu___3 - -> FStar_Interactive_Ide_Types.LaxCheck - | FStar_Interactive_Ide_Types.LaxWithSymbols -> - FStar_Interactive_Ide_Types.LaxCheck - | uu___3 -> FStar_Interactive_Ide_Types.FullCheck in - let uu___3 = - let uu___4 = - let uu___5 = - FStar_Compiler_Effect.op_Bang - FStar_Interactive_PushHelper.repl_stack in - inspect_repl_stack uu___5 decls1 push_kind - with_symbols - write_full_buffer_fragment_progress in - run_qst uu___4 qid1 in - (match uu___3 with - | (queries, issues) -> - (if - request_type <> - FStar_Interactive_Ide_Types.Cache - then log_syntax_issues err_opt - else (); - (let uu___6 = FStar_Options.debug_any () in - if uu___6 - then - let uu___7 = - let uu___8 = - FStar_Compiler_List.map - FStar_Interactive_Ide_Types.query_to_string - queries in - FStar_String.concat "\n" uu___8 in - FStar_Compiler_Util.print1 - "Generating queries\n%s\n" uu___7 - else ()); - if - request_type <> - FStar_Interactive_Ide_Types.Cache - then (queries, issues) - else ([], issues))))) - | FStar_Parser_ParseIt.ParseError err -> - (if request_type = FStar_Interactive_Ide_Types.Full - then log_syntax_issues (FStar_Pervasives_Native.Some err) - else (); - ([], [])) - | uu___ -> failwith "Unexpected parse result" in - qs + fun with_symbols -> + fun write_full_buffer_fragment_progress -> + let parse_result = parse_code code in + let log_syntax_issues err = + match err with + | FStar_Pervasives_Native.None -> () + | FStar_Pervasives_Native.Some err1 -> + let issue = syntax_issue err1 in + write_full_buffer_fragment_progress + (FragmentError [issue]) in + let filter_decls decls = + match request_type with + | FStar_Interactive_Ide_Types.VerifyToPosition + (uu___, line, _col) -> + FStar_Compiler_List.filter + (fun uu___1 -> + match uu___1 with + | (d, uu___2) -> + let start = + FStar_Compiler_Range.start_of_range + d.FStar_Parser_AST.drange in + let start_line = + FStar_Compiler_Range.line_of_pos start in + start_line <= line) decls + | FStar_Interactive_Ide_Types.LaxToPosition + (uu___, line, _col) -> + FStar_Compiler_List.filter + (fun uu___1 -> + match uu___1 with + | (d, uu___2) -> + let start = + FStar_Compiler_Range.start_of_range + d.FStar_Parser_AST.drange in + let start_line = + FStar_Compiler_Range.line_of_pos start in + start_line <= line) decls + | uu___ -> decls in + let qs = + match parse_result with + | FStar_Parser_ParseIt.IncrementalFragment + (decls, uu___, err_opt) -> + ((let uu___2 = + FStar_Compiler_Util.string_of_int + (FStar_Compiler_List.length decls) in + FStar_Compiler_Util.print1 "Parsed %s declarations\n" + uu___2); + (match (request_type, decls) with + | (FStar_Interactive_Ide_Types.ReloadDeps, d::uu___2) + -> + let uu___3 = + let uu___4 = + let uu___5 = + FStar_Compiler_Effect.op_Bang + FStar_Interactive_PushHelper.repl_stack in + reload_deps uu___5 in + op_let_Bang uu___4 + (fun queries -> + let uu___5 = + push_decl + FStar_Interactive_Ide_Types.FullCheck + with_symbols + write_full_buffer_fragment_progress d in + op_let_Bang uu___5 + (fun push_mod -> + return + ((FStar_Compiler_List.op_At queries + push_mod), []))) in + run_qst uu___3 qid1 + | uu___2 -> + let decls1 = filter_decls decls in + let push_kind = + match request_type with + | FStar_Interactive_Ide_Types.LaxToPosition + uu___3 -> + FStar_Interactive_Ide_Types.LaxCheck + | FStar_Interactive_Ide_Types.Lax -> + FStar_Interactive_Ide_Types.LaxCheck + | uu___3 -> FStar_Interactive_Ide_Types.FullCheck in + let uu___3 = + let uu___4 = + let uu___5 = + FStar_Compiler_Effect.op_Bang + FStar_Interactive_PushHelper.repl_stack in + inspect_repl_stack uu___5 decls1 push_kind + with_symbols + write_full_buffer_fragment_progress in + run_qst uu___4 qid1 in + (match uu___3 with + | (queries, issues) -> + (if + request_type <> + FStar_Interactive_Ide_Types.Cache + then log_syntax_issues err_opt + else (); + (let uu___6 = FStar_Options.debug_any () in + if uu___6 + then + let uu___7 = + let uu___8 = + FStar_Compiler_List.map + FStar_Interactive_Ide_Types.query_to_string + queries in + FStar_String.concat "\n" uu___8 in + FStar_Compiler_Util.print1 + "Generating queries\n%s\n" uu___7 + else ()); + if + request_type <> + FStar_Interactive_Ide_Types.Cache + then (queries, issues) + else ([], issues))))) + | FStar_Parser_ParseIt.ParseError err -> + (if request_type = FStar_Interactive_Ide_Types.Full + then + log_syntax_issues (FStar_Pervasives_Native.Some err) + else (); + ([], [])) + | uu___ -> failwith "Unexpected parse result" in + qs let (format_code : FStar_Interactive_Ide_Types.repl_state -> Prims.string -> @@ -492,21 +517,21 @@ let (format_code : let parse_result = parse_code code in match parse_result with | FStar_Parser_ParseIt.IncrementalFragment - (decls, uu___, FStar_Pervasives_Native.None) -> + (decls, comments, FStar_Pervasives_Native.None) -> let doc_to_string doc = FStar_Pprint.pretty_string (FStar_Compiler_Util.float_of_string "1.0") (Prims.of_int (100)) doc in let formatted_code = - let uu___1 = + let uu___ = FStar_Compiler_List.map - (fun uu___2 -> - match uu___2 with - | (d, uu___3) -> - let uu___4 = + (fun uu___1 -> + match uu___1 with + | (d, uu___2) -> + let uu___3 = FStar_Parser_ToDocument.decl_to_document d in - doc_to_string uu___4) decls in - FStar_Compiler_Effect.op_Bar_Greater uu___1 + doc_to_string uu___3) decls in + FStar_Compiler_Effect.op_Bar_Greater uu___ (FStar_String.concat "\n\n") in FStar_Pervasives.Inl formatted_code | FStar_Parser_ParseIt.IncrementalFragment diff --git a/ocaml/fstar-lib/generated/FStar_Interactive_JsonHelper.ml b/ocaml/fstar-lib/generated/FStar_Interactive_JsonHelper.ml index 6ccd25b8b1a..32dabaaad8c 100644 --- a/ocaml/fstar-lib/generated/FStar_Interactive_JsonHelper.ml +++ b/ocaml/fstar-lib/generated/FStar_Interactive_JsonHelper.ml @@ -69,6 +69,11 @@ let (js_int : FStar_Compiler_Util.json -> Prims.int) = match uu___ with | FStar_Compiler_Util.JsonInt i -> i | other -> js_fail "int" other +let (js_bool : FStar_Compiler_Util.json -> Prims.bool) = + fun uu___ -> + match uu___ with + | FStar_Compiler_Util.JsonBool b -> b + | other -> js_fail "int" other let (js_str : FStar_Compiler_Util.json -> Prims.string) = fun uu___ -> match uu___ with diff --git a/src/fstar/FStar.Interactive.Ide.Types.fst b/src/fstar/FStar.Interactive.Ide.Types.fst index 754e8860f29..0ad6c40cd72 100755 --- a/src/fstar/FStar.Interactive.Ide.Types.fst +++ b/src/fstar/FStar.Interactive.Ide.Types.fst @@ -102,7 +102,7 @@ type repl_task = type full_buffer_request_kind = | Full : full_buffer_request_kind - | LaxWithSymbols : full_buffer_request_kind + | Lax : full_buffer_request_kind | Cache : full_buffer_request_kind | ReloadDeps : full_buffer_request_kind | VerifyToPosition of position @@ -124,7 +124,7 @@ type query' = | ProtocolViolation of string // FullBuffer: To check the full contents of a document. // FStar.Interactive.Incremental parses it into chunks and turns this into several Push queries -| FullBuffer of string & full_buffer_request_kind +| FullBuffer of string & full_buffer_request_kind & bool //bool is with_symbol // Callback: This is an internal query, it cannot be raised by a client. // It is useful to inject operations into the query stream. // e.g., Incremental uses it print progress messages to the client in between diff --git a/src/fstar/FStar.Interactive.Ide.fst b/src/fstar/FStar.Interactive.Ide.fst index 3d3de0d248c..9ef2f35ec8b 100644 --- a/src/fstar/FStar.Interactive.Ide.fst +++ b/src/fstar/FStar.Interactive.Ide.fst @@ -223,9 +223,10 @@ let unpack_interactive_query json = assoc "to-position.line" "line" to_pos |> js_int, assoc "to-position.column" "column" to_pos |> js_int in - let parse_full_buffer_kind = function + let parse_full_buffer_kind (kind:string) = + match kind with | "full" -> Full - | "lax-with-symbols" -> LaxWithSymbols + | "lax" -> Lax | "cache" -> Cache | "reload-deps" -> ReloadDeps | "verify-to-position" -> VerifyToPosition (read_to_position ()) @@ -244,7 +245,9 @@ let unpack_interactive_query json = push_line = arg "line" |> js_int; push_column = arg "column" |> js_int; push_peek_only = query = "peek" }) - | "full-buffer" -> FullBuffer (arg "code" |> js_str, parse_full_buffer_kind (arg "kind" |> js_str)) + | "full-buffer" -> FullBuffer (arg "code" |> js_str, + parse_full_buffer_kind (arg "kind" |> js_str), + arg "with-symbols" |> js_bool) | "autocomplete" -> AutoComplete (arg "partial-symbol" |> js_str, try_arg "context" |> js_optional_completion_context) | "lookup" -> Lookup (arg "symbol" |> js_str, @@ -601,6 +604,9 @@ let write_full_buffer_fragment_progress (di:Incremental.fragment_progress) = "code-digest", JsonStr (BU.digest_of_string cf.code)] in match di with + | FullBufferStarted -> + write_progress (Some "full-buffer-started") [] + | FragmentStarted d -> write_progress (Some "full-buffer-fragment-started") ["ranges", json_of_def_range d.FStar.Parser.AST.drange] @@ -615,6 +621,7 @@ let write_full_buffer_fragment_progress (di:Incremental.fragment_progress) = | FragmentFailed d -> write_progress (Some "full-buffer-fragment-failed") ["ranges", json_of_def_range d.FStar.Parser.AST.drange] + | FragmentError issues -> let qid = match !repl_current_qid with @@ -623,6 +630,9 @@ let write_full_buffer_fragment_progress (di:Incremental.fragment_progress) = in write_json (json_of_response qid QueryNOK (JsonList (List.map json_of_issue issues))) + | FullBufferFinished -> + write_progress (Some "full-buffer-finished") [] + let run_push_without_deps st query : (query_status & json) & either repl_state int = @@ -1030,20 +1040,19 @@ let maybe_cancel_queries st l = let rec fold_query (f:repl_state -> query -> run_query_result) (l:list query) (st:repl_state) - (responses: list json) : run_query_result = match l with - | [] -> (QueryOK, responses), Inl st + | [] -> (QueryOK, []), Inl st | q::l -> - let (status, resp), st' = f st q in - let responses = responses @ resp in + let (status, responses), st' = f st q in + List.iter (write_response q.qid status) responses; match status, st' with | QueryOK, Inl st -> let st = buffer_input_queries st in let l, st = maybe_cancel_queries st l in - fold_query f l st responses + fold_query f l st | _ -> - (status, responses), st' + (status, []), st' let validate_query st (q: query) : query = match q.qq with @@ -1066,11 +1075,16 @@ let rec run_query st (q: query) : (query_status * list json) * either repl_state | VfsAdd (fname, contents) -> as_json_list (run_vfs_add st fname contents) | Push pquery -> as_json_list (run_push st pquery) | Pop -> as_json_list (run_pop st) - | FullBuffer (code, full) -> + | FullBuffer (code, full_kind, with_symbols) -> + let open FStar.Interactive.Incremental in + write_full_buffer_fragment_progress FullBufferStarted; let queries, issues = - FStar.Interactive.Incremental.run_full_buffer st q.qid code full write_full_buffer_fragment_progress + run_full_buffer st q.qid code full_kind with_symbols write_full_buffer_fragment_progress in - fold_query validate_and_run_query queries st issues + List.iter (write_response q.qid QueryOK) issues; + let res = fold_query validate_and_run_query queries st in + write_full_buffer_fragment_progress FullBufferFinished; + res | AutoComplete (search_term, context) -> as_json_list (run_autocomplete st search_term context) | Lookup (symbol, context, pos_opt, rq_info, symrange) -> diff --git a/src/fstar/FStar.Interactive.Incremental.fst b/src/fstar/FStar.Interactive.Incremental.fst index 7357bad8221..eeb408e44c8 100755 --- a/src/fstar/FStar.Interactive.Incremental.fst +++ b/src/fstar/FStar.Interactive.Incremental.fst @@ -181,7 +181,7 @@ let inspect_repl_stack (s:repl_stack_t) | Some (prefix, first_push, rest) -> let entries = first_push :: rest in let repl_task (_, (p, _)) = p in - let rec matching_prefix (accum:list json) entries (ds:list (decl & code_fragment)) + let rec matching_prefix (accum:list json) (lookups:list query) entries (ds:list (decl & code_fragment)) : qst (list query & list json) = match entries, ds with | [], [] -> @@ -190,7 +190,7 @@ let inspect_repl_stack (s:repl_stack_t) | e::entries, d::ds -> ( match repl_task e with | Noop -> - matching_prefix accum entries (d::ds) + matching_prefix accum lookups entries (d::ds) | PushFragment (Inl frag, _, _) -> let! pops = pop_entries (e::entries) in let! pushes = push_decls (d::ds) in @@ -200,11 +200,15 @@ let inspect_repl_stack (s:repl_stack_t) then ( let d, s = d in write_full_buffer_fragment_progress (FragmentSuccess (d, s, pk)); - matching_prefix (issues@accum) entries ds + if with_symbols + then let! lookups' = dump_symbols d in + matching_prefix (issues@accum) (lookups'@lookups) entries ds + else + matching_prefix (issues@accum) lookups entries ds ) else let! pops = pop_entries (e::entries) in let! pushes = push_decls (d::ds) in - return (pops @ pushes, accum) + return (pops @ lookups @ pushes, accum) ) | [], ds -> @@ -215,7 +219,7 @@ let inspect_repl_stack (s:repl_stack_t) let! pops = pop_entries es in return (pops, accum) in - matching_prefix [] entries ds + matching_prefix [] [] entries ds (* A reload_deps request just pops away the entire stack of PushFragments. We also push on just the `module A` declaration after popping. That's done below. *) @@ -261,6 +265,7 @@ let run_full_buffer (st:repl_state) (qid:string) (code:string) (request_type:full_buffer_request_kind) + (with_symbols:bool) (write_full_buffer_fragment_progress: fragment_progress -> unit) : list query & list json = let parse_result = parse_code code in @@ -283,7 +288,6 @@ let run_full_buffer (st:repl_state) decls | _ -> decls in - let with_symbols = request_type = LaxWithSymbols in let qs = match parse_result with | IncrementalFragment (decls, _, err_opt) -> ( @@ -302,10 +306,9 @@ let run_full_buffer (st:repl_state) let push_kind = match request_type with | LaxToPosition _ -> LaxCheck - | LaxWithSymbols -> LaxCheck + | Lax -> LaxCheck | _ -> FullCheck in - let queries, issues = run_qst (inspect_repl_stack (!repl_stack) decls push_kind with_symbols write_full_buffer_fragment_progress) qid in @@ -331,7 +334,7 @@ let run_full_buffer (st:repl_state) let format_code (st:repl_state) (code:string) = let parse_result = parse_code code in match parse_result with - | IncrementalFragment (decls, _, None) -> + | IncrementalFragment (decls, comments, None) -> let doc_to_string doc = FStar.Pprint.pretty_string (float_of_string "1.0") 100 doc in diff --git a/src/fstar/FStar.Interactive.Incremental.fsti b/src/fstar/FStar.Interactive.Incremental.fsti index 153fa1ad3e6..78d21c4c820 100755 --- a/src/fstar/FStar.Interactive.Incremental.fsti +++ b/src/fstar/FStar.Interactive.Incremental.fsti @@ -27,10 +27,12 @@ open FStar.Compiler.Util (* Various kinds of progress messages to print back to the client *) type fragment_progress = + | FullBufferStarted | FragmentStarted of decl | FragmentSuccess of (decl & FStar.Parser.ParseIt.code_fragment & push_kind) | FragmentFailed of decl | FragmentError of list issue + | FullBufferFinished (* Translates a full-buffer(qid, code) query by 1. Parsing the code into its declarations @@ -46,6 +48,7 @@ val run_full_buffer (st:repl_state) (qid:string) (code:string) (full:full_buffer_request_kind) + (with_symbols:bool) (write_full_buffer_fragment_progress: fragment_progress -> unit) : list query & list json diff --git a/src/fstar/FStar.Interactive.JsonHelper.fst b/src/fstar/FStar.Interactive.JsonHelper.fst index f4aa44d5905..e03f4f4c46a 100644 --- a/src/fstar/FStar.Interactive.JsonHelper.fst +++ b/src/fstar/FStar.Interactive.JsonHelper.fst @@ -65,6 +65,9 @@ let js_fail expected got = let js_int : json -> int = function | JsonInt i -> i | other -> js_fail "int" other +let js_bool : json -> bool = function + | JsonBool b -> b + | other -> js_fail "int" other let js_str : json -> string = function | JsonStr s -> s | other -> js_fail "string" other diff --git a/src/fstar/FStar.Interactive.JsonHelper.fsti b/src/fstar/FStar.Interactive.JsonHelper.fsti index 6493a6d24dd..1f9219597ec 100644 --- a/src/fstar/FStar.Interactive.JsonHelper.fsti +++ b/src/fstar/FStar.Interactive.JsonHelper.fsti @@ -41,6 +41,7 @@ val write_jsonrpc : json -> unit // Only used in LSP val js_fail : string -> json -> 'a val js_int : json -> int +val js_bool : json -> bool val js_str : json -> string val js_list : (json -> 'a) -> json -> list 'a val js_assoc : json -> assoct From 269e8102e117768a84f62280864c6c85f1c6a7b2 Mon Sep 17 00:00:00 2001 From: Nikhil Swamy Date: Sat, 1 Apr 2023 08:34:47 -0700 Subject: [PATCH 48/48] fix incremental test suite oracle; do not fail on lookup failures --- .../generated/FStar_Interactive_Ide.ml | 33 ++-- src/fstar/FStar.Interactive.Ide.fst | 16 +- tests/ide/test-incremental.py | 146 +++++++++++------- 3 files changed, 127 insertions(+), 68 deletions(-) diff --git a/ocaml/fstar-lib/generated/FStar_Interactive_Ide.ml b/ocaml/fstar-lib/generated/FStar_Interactive_Ide.ml index 7d21c8f9b3c..e39623625e0 100644 --- a/ocaml/fstar-lib/generated/FStar_Interactive_Ide.ml +++ b/ocaml/fstar-lib/generated/FStar_Interactive_Ide.ml @@ -1944,18 +1944,31 @@ let run_lookup : fun pos_opt -> fun requested_info -> fun symrange -> - let uu___ = - run_lookup' st symbol context pos_opt requested_info symrange in - match uu___ with - | FStar_Pervasives.Inl err_msg -> + try + (fun uu___ -> + match () with + | () -> + let uu___1 = + run_lookup' st symbol context pos_opt requested_info + symrange in + (match uu___1 with + | FStar_Pervasives.Inl err_msg -> + ((FStar_Interactive_Ide_Types.QueryOK, + (FStar_Compiler_Util.JsonStr err_msg)), + (FStar_Pervasives.Inl st)) + | FStar_Pervasives.Inr (kind, info) -> + ((FStar_Interactive_Ide_Types.QueryOK, + (FStar_Compiler_Util.JsonAssoc + (("kind", + (FStar_Compiler_Util.JsonStr kind)) + :: info))), (FStar_Pervasives.Inl st)))) () + with + | uu___ -> ((FStar_Interactive_Ide_Types.QueryOK, - (FStar_Compiler_Util.JsonStr err_msg)), + (FStar_Compiler_Util.JsonStr + (Prims.op_Hat "Lookup of " + (Prims.op_Hat symbol " failed")))), (FStar_Pervasives.Inl st)) - | FStar_Pervasives.Inr (kind, info) -> - ((FStar_Interactive_Ide_Types.QueryOK, - (FStar_Compiler_Util.JsonAssoc - (("kind", (FStar_Compiler_Util.JsonStr kind)) :: - info))), (FStar_Pervasives.Inl st)) let run_code_autocomplete : 'uuuuu . FStar_Interactive_Ide_Types.repl_state -> diff --git a/src/fstar/FStar.Interactive.Ide.fst b/src/fstar/FStar.Interactive.Ide.fst index 9ef2f35ec8b..432e8bccc6d 100644 --- a/src/fstar/FStar.Interactive.Ide.fst +++ b/src/fstar/FStar.Interactive.Ide.fst @@ -746,12 +746,16 @@ let run_lookup' st symbol context pos_opt requested_info symrange = | LKCode -> run_code_lookup st symbol pos_opt requested_info symrange let run_lookup st symbol context pos_opt requested_info symrange = - match run_lookup' st symbol context pos_opt requested_info symrange with - | Inl err_msg -> - // No result found, but don't fail the query - ((QueryOK, JsonStr err_msg), Inl st) - | Inr (kind, info) -> - ((QueryOK, JsonAssoc (("kind", JsonStr kind) :: info)), Inl st) + try + match run_lookup' st symbol context pos_opt requested_info symrange with + | Inl err_msg -> + // No result found, but don't fail the query + ((QueryOK, JsonStr err_msg), Inl st) + | Inr (kind, info) -> + ((QueryOK, JsonAssoc (("kind", JsonStr kind) :: info)), Inl st) + with + | _ -> ((QueryOK, JsonStr ("Lookup of " ^ symbol^ " failed")), Inl st) + let run_code_autocomplete st search_term = let result = QH.ck_completion st search_term in diff --git a/tests/ide/test-incremental.py b/tests/ide/test-incremental.py index 697a2d80688..9f3c1788e61 100644 --- a/tests/ide/test-incremental.py +++ b/tests/ide/test-incremental.py @@ -5,6 +5,7 @@ import subprocess import json import re +import hashlib # The path to the F* executable @@ -46,6 +47,52 @@ def get_contents_in_range(file_lines, start_pos, end_pos): lines = "\n".join(lines) return lines +def check_one_fragment(file_lines, json_objects, from_line): + l = from_line + # The line from_line is the first line of the fragment + if json_objects[l]["kind"] != "message" or json_objects[l]["level"] != "progress" or json_objects[l]["contents"]["stage"] != "full-buffer-fragment-started": + print(f"Expected a full-buffer-fragment-started message at line {l} got {json_objects[l]}") + return None + l += 1 + if json_objects[l]["kind"] != "message" or json_objects[l]["level"] != "progress" or json_objects[l]["contents"]["stage"] != "full-buffer-fragment-lax-ok": + print(f"Expected a full-buffer-fragment-lax-ok message at line {l} got {json_objects[l]}") + return None + lax_ok = json_objects[l]["contents"] + code_frag = lax_ok["code-fragment"] + start_pos = code_frag["range"]["beg"] + end_pos = code_frag["range"]["end"] + last_fragment_end = end_pos + lines = get_contents_in_range(file_lines, start_pos, end_pos) + # compute an MD5 hash of the lines + # if the hash does not match the hash in the message, print an error message + # and return False + hash = hashlib.md5(lines.encode()).hexdigest() + if hash != code_frag["code-digest"]: + print(f"Hash does not match: Expected {hash} but got {code_frag['code-digest']}") + return None + # # join the lines together with newlines + # if code_frag["code"] != lines: + # print(f"Code fragment does not match the code: Expected {code_frag['code']} but got {lines}") + # return None + l += 1 + # Next line has the form {"kind":"response","query-id":"2.91","status":"success","response":[]} + if json_objects[l]["kind"] != "response": + return None + if json_objects[l]["status"] != "success": + return None + l += 1 + # Next several lines have status "success" and kind=response + while (l < len(json_objects)): + # print(f"Checking line {l} contents {json_objects[l]}") + # if the line has a status field + if json_objects[l]["kind"] != "response": + break + if json_objects[l]["status"] != "success": + print(f"Expected a success response at line {l} got {json_objects[l]}") + return None + l += 1 + return (l, last_fragment_end) + # A function to validate the response from F* interactive mode def validate_response(response, file_contents): file_lines = file_contents.splitlines() @@ -74,40 +121,46 @@ def validate_response(response, file_contents): if json_objects[1]["status"] != "success": print("Second line does not have status success") return False - + + l = 2 + if json_objects[l]["level"] != "progress" or json_objects[l]["contents"]["stage"] != "full-buffer-started": + print("Third line is not a message") + return False + + l=3 # Third line has the form {"kind":"message","query-id":"2","level":"info","contents":"Parsed 138 declarations\n"} - if json_objects[2]["kind"] != "message": - print(f"Second line is not a message {json_objects[2]}") + if json_objects[l]["kind"] != "message": + print(f"Third line is not a message {json_objects[2]}") return False - if json_objects[2]["query-id"] != "2": - print("Second line does not have query-id 2") + if json_objects[l]["query-id"] != "2": + print("Third line does not have query-id 2") return False - if json_objects[2]["level"] != "info": - print("Second line does not have level info") + if json_objects[l]["level"] != "info": + print("Third line does not have level info") return False # the contents should match a regular expression of the form "Parsed \d+ declarations" # store the number of declarations in a variable - if not re.match(r"Parsed \d+ declarations", json_objects[2]["contents"]): + if not re.match(r"Parsed \d+ declarations", json_objects[l]["contents"]): print("Second line does not have the correct contents") return False # Check that the number of declarations is 138 - num_decls = int(re.search(r"\d+", json_objects[2]["contents"]).group()) - + num_decls = int(re.search(r"\d+", json_objects[l]["contents"]).group()) + l=4 # Fourth line has kind "message" and level "progress" and contents.stage = "full-buffer-fragment-started" - if json_objects[3]["kind"] != "message": + if json_objects[l]["kind"] != "message": print("Third line is not a message") return False - if json_objects[3]["level"] != "progress": + if json_objects[l]["level"] != "progress": print("Third line does not have level progress") return False - if json_objects[3]["contents"]["stage"] != "full-buffer-fragment-started": + if json_objects[l]["contents"]["stage"] != "full-buffer-fragment-started": print("Third line does not have stage full-buffer-fragment-started") return False - + l=5 # Next several messages are progress messages with contents.stage = "loading-dependency" # Check all of these messages for the correct kind, level, and stage and stop # when the first message with a different kind or level or stage is found - for i in range(4, len(json_objects) - 1): + for i in range(l, len(json_objects) - 1): if json_objects[i]["kind"] != "message": break if json_objects[i]["level"] != "progress": @@ -118,10 +171,17 @@ def validate_response(response, file_contents): if json_objects[i]["contents"]["stage"] != None: print(f"Message {i} has contents {json_objects[i]} does not have stage None") return False + i += 1 # the next message has conents.stage = "full-buffer-fragment-lax-ok" - if json_objects[i + 1]["contents"]["stage"] != "full-buffer-fragment-lax-ok": + if json_objects[i]["contents"]["stage"] != "full-buffer-fragment-lax-ok": print("Message does not have stage full-buffer-fragment-lax-ok") return False + i += 1 + # the next message has conents.stage = "full-buffer-fragment-lax-ok" + if json_objects[i]["status"] != "success": + print(f"Message does not have success message at line {i}") + return False + i += 1 # Then, we have a sequence of pairs of messages # where the first message in the pair has contents.stage = "full-buffer-fragment-started" # and the second message in the pair has contents.stage = "full-buffer-fragment-lax-ok" @@ -131,28 +191,18 @@ def validate_response(response, file_contents): # The second message in the pair has contents.stage = "full-buffer-fragment-lax-ok" num_successes = 1 last_fragment_end = [0, 0] - for j in range(i + 2, len(json_objects) - 1, 2): - if json_objects[j]["kind"] != "message": - break - if json_objects[j]["level"] != "progress": + while (i < len(json_objects) - 1): + if json_objects[i]["kind"] != "message": break - if json_objects[j]["contents"]["stage"] != "full-buffer-fragment-started": + if json_objects[i]["level"] != "progress": break - if json_objects[j + 1]["contents"]["stage"] != "full-buffer-fragment-lax-ok": + if json_objects[i]["contents"]["stage"] != "full-buffer-fragment-started": break - # {"stage":"full-buffer-fragment-lax-ok","ranges":{"fname":"","beg":[16,0],"end":[16,29]},"code-fragment":{"range":{"fname":"","beg":[16,0],"end":[16,29]},"code":"module FStar.Reflection.Arith"} - lax_ok = json_objects[j + 1]["contents"] - code_frag = lax_ok["code-fragment"] - start_pos = code_frag["range"]["beg"] - end_pos = code_frag["range"]["end"] - last_fragment_end = end_pos - lines = get_contents_in_range(file_lines, start_pos, end_pos) - # join the lines together with newlines - if code_frag["code"] != lines: - print(f"Code fragment does not match the code: Expected {code_frag['code']} but got {lines}") + res = check_one_fragment(file_lines, json_objects, i) + if res == None: return False - num_successes = num_successes + 1 - + (i, last_fragment_end) = res + num_successes += 1 remaining_lines = get_contents_in_range(file_lines, last_fragment_end, None) remaining_lines = remaining_lines.strip() if remaining_lines != "" and not (is_comment(remaining_lines)): @@ -162,24 +212,16 @@ def validate_response(response, file_contents): print(f"Number of successes {num_successes} does not match number of declarations {num_decls}") return False - # The next sequence of messages have status success and query-id "2" - # Check all of these messages for the correct kind and status and stop - # when the first message with a different kind or status is found - for k in range(j + 1, len(json_objects)): - if json_objects[k]["kind"] != "response": - break - if json_objects[k]["status"] != "success": - break - if json_objects[k]["query-id"] != "2": - break - - # check that k is the index of the last message - if k != len(json_objects) - 1: - print(f"Unexpected last message at index {k}, contents = {json_objects[k]}") + # check that i is the index of the secod last message + if i != len(json_objects) - 2: + print(f"Unexpected last message at index {i}, contents = {json_objects[i]}") return False - + if json_objects[i]["kind"] != "message" or json_objects[i]["level"] != "progress" or json_objects[i]["contents"]["stage"] != "full-buffer-finished": + print(f"Unexpected last message at index {i}, contents = {json_objects[i]}") + return False + i += 1 # The last message has query-id "3" and status "success" - if json_objects[k]["query-id"] != "3": + if json_objects[i]["query-id"] != "3": print("Last message does not have query-id 3") return False @@ -201,9 +243,9 @@ def test_file(file): # The first line is a JSON object initializing the ide # The second line is a JSON object containing the contents of the file # The third line is an exit command - request = f'{{"query-id":"1", "query": "vfs-add", "args":{{"filename":null, "contents": {json_contents}}}}}\n{{"query-id":"2", "query": "full-buffer", "args":{{"kind": "lax-with-symbols", "code": {json_contents}, "line":1, "column":0}}}}\n{{"query-id":"3", "query": "exit", "args":{{}}}}\n' + request = f'{{"query-id":"1", "query": "vfs-add", "args":{{"filename":null, "contents": {json_contents}}}}}\n{{"query-id":"2", "query": "full-buffer", "args":{{"kind": "lax", "with-symbols":true, "code": {json_contents}, "line":1, "column":0}}}}\n{{"query-id":"3", "query": "exit", "args":{{}}}}\n' # print the request to the console for debugging - #print(request) + # print(request) # Run fstar on the file with the request as stdin p = subprocess.run([fstar, "--admit_smt_queries", "true", "--ide", file], input=request, encoding="utf-8", stdout=subprocess.PIPE, stderr=subprocess.PIPE) # Read the response from stdout