wenzelm@42243: (* Title: Pure/Syntax/syntax_phases.ML wenzelm@42241: Author: Makarius wenzelm@42241: wenzelm@42243: Main phases of inner syntax processing, with standard implementations wenzelm@42243: of parse/unparse operations. wenzelm@42241: *) wenzelm@42241: wenzelm@42243: signature SYNTAX_PHASES = wenzelm@42242: sig wenzelm@42242: val term_sorts: term -> (indexname * sort) list wenzelm@42242: val typ_of_term: (indexname -> sort) -> term -> typ wenzelm@42242: val decode_term: Proof.context -> wenzelm@42242: Position.reports * term Exn.result -> Position.reports * term Exn.result wenzelm@42242: val parse_ast_pattern: Proof.context -> string * string -> Ast.ast wenzelm@42245: val term_of_typ: bool -> typ -> term wenzelm@42242: end wenzelm@42242: wenzelm@42243: structure Syntax_Phases: SYNTAX_PHASES = wenzelm@42241: struct wenzelm@42241: wenzelm@42242: (** decode parse trees **) wenzelm@42242: wenzelm@42242: (* sort_of_term *) wenzelm@42242: wenzelm@42242: fun sort_of_term tm = wenzelm@42242: let wenzelm@42242: fun err () = raise TERM ("sort_of_term: bad encoding of classes", [tm]); wenzelm@42242: wenzelm@42242: fun class s = Lexicon.unmark_class s handle Fail _ => err (); wenzelm@42242: wenzelm@42242: fun classes (Const (s, _)) = [class s] wenzelm@42242: | classes (Const ("_classes", _) $ Const (s, _) $ cs) = class s :: classes cs wenzelm@42242: | classes _ = err (); wenzelm@42242: wenzelm@42242: fun sort (Const ("_topsort", _)) = [] wenzelm@42242: | sort (Const (s, _)) = [class s] wenzelm@42242: | sort (Const ("_sort", _) $ cs) = classes cs wenzelm@42242: | sort _ = err (); wenzelm@42242: in sort tm end; wenzelm@42242: wenzelm@42242: wenzelm@42242: (* term_sorts *) wenzelm@42242: wenzelm@42242: fun term_sorts tm = wenzelm@42242: let wenzelm@42242: val sort_of = sort_of_term; wenzelm@42242: wenzelm@42242: fun add_env (Const ("_ofsort", _) $ Free (x, _) $ cs) = wenzelm@42242: insert (op =) ((x, ~1), sort_of cs) wenzelm@42242: | add_env (Const ("_ofsort", _) $ (Const ("_tfree", _) $ Free (x, _)) $ cs) = wenzelm@42242: insert (op =) ((x, ~1), sort_of cs) wenzelm@42242: | add_env (Const ("_ofsort", _) $ Var (xi, _) $ cs) = wenzelm@42242: insert (op =) (xi, sort_of cs) wenzelm@42242: | add_env (Const ("_ofsort", _) $ (Const ("_tvar", _) $ Var (xi, _)) $ cs) = wenzelm@42242: insert (op =) (xi, sort_of cs) wenzelm@42242: | add_env (Abs (_, _, t)) = add_env t wenzelm@42242: | add_env (t1 $ t2) = add_env t1 #> add_env t2 wenzelm@42242: | add_env _ = I; wenzelm@42242: in add_env tm [] end; wenzelm@42242: wenzelm@42242: wenzelm@42242: (* typ_of_term *) wenzelm@42242: wenzelm@42242: fun typ_of_term get_sort tm = wenzelm@42242: let wenzelm@42242: fun err () = raise TERM ("typ_of_term: bad encoding of type", [tm]); wenzelm@42242: wenzelm@42242: fun typ_of (Free (x, _)) = TFree (x, get_sort (x, ~1)) wenzelm@42242: | typ_of (Var (xi, _)) = TVar (xi, get_sort xi) wenzelm@42242: | typ_of (Const ("_tfree",_) $ (t as Free _)) = typ_of t wenzelm@42242: | typ_of (Const ("_tvar",_) $ (t as Var _)) = typ_of t wenzelm@42242: | typ_of (Const ("_ofsort", _) $ Free (x, _) $ _) = TFree (x, get_sort (x, ~1)) wenzelm@42242: | typ_of (Const ("_ofsort", _) $ (Const ("_tfree",_) $ Free (x, _)) $ _) = wenzelm@42242: TFree (x, get_sort (x, ~1)) wenzelm@42242: | typ_of (Const ("_ofsort", _) $ Var (xi, _) $ _) = TVar (xi, get_sort xi) wenzelm@42242: | typ_of (Const ("_ofsort", _) $ (Const ("_tvar",_) $ Var (xi, _)) $ _) = wenzelm@42242: TVar (xi, get_sort xi) wenzelm@42242: | typ_of (Const ("_dummy_ofsort", _) $ t) = TFree ("'_dummy_", sort_of_term t) wenzelm@42242: | typ_of t = wenzelm@42242: let wenzelm@42242: val (head, args) = Term.strip_comb t; wenzelm@42242: val a = wenzelm@42242: (case head of wenzelm@42242: Const (c, _) => (Lexicon.unmark_type c handle Fail _ => err ()) wenzelm@42242: | _ => err ()); wenzelm@42242: in Type (a, map typ_of args) end; wenzelm@42242: in typ_of tm end; wenzelm@42242: wenzelm@42242: wenzelm@42242: (* parsetree_to_ast *) wenzelm@42242: wenzelm@42242: fun lookup_tr tab c = Option.map fst (Symtab.lookup tab c); wenzelm@42242: wenzelm@42242: fun parsetree_to_ast ctxt constrain_pos trf parsetree = wenzelm@42242: let wenzelm@42242: val {get_class, get_type, markup_class, markup_type} = ProofContext.type_context ctxt; wenzelm@42242: wenzelm@42242: val reports = Unsynchronized.ref ([]: Position.reports); wenzelm@42242: fun report pos = Position.reports reports [pos]; wenzelm@42242: wenzelm@42242: fun trans a args = wenzelm@42242: (case trf a of wenzelm@42242: NONE => Ast.mk_appl (Ast.Constant a) args wenzelm@42242: | SOME f => f ctxt args); wenzelm@42242: wenzelm@42242: fun ast_of (Parser.Node ("_class_name", [Parser.Tip tok])) = wenzelm@42242: let wenzelm@42242: val c = get_class (Lexicon.str_of_token tok); wenzelm@42242: val _ = report (Lexicon.pos_of_token tok) markup_class c; wenzelm@42242: in Ast.Constant (Lexicon.mark_class c) end wenzelm@42242: | ast_of (Parser.Node ("_type_name", [Parser.Tip tok])) = wenzelm@42242: let wenzelm@42242: val c = get_type (Lexicon.str_of_token tok); wenzelm@42242: val _ = report (Lexicon.pos_of_token tok) markup_type c; wenzelm@42242: in Ast.Constant (Lexicon.mark_type c) end wenzelm@42242: | ast_of (Parser.Node ("_constrain_position", [pt as Parser.Tip tok])) = wenzelm@42242: if constrain_pos then wenzelm@42242: Ast.Appl [Ast.Constant "_constrain", ast_of pt, wenzelm@42242: Ast.Variable (Lexicon.encode_position (Lexicon.pos_of_token tok))] wenzelm@42242: else ast_of pt wenzelm@42242: | ast_of (Parser.Node (a, pts)) = trans a (map ast_of pts) wenzelm@42242: | ast_of (Parser.Tip tok) = Ast.Variable (Lexicon.str_of_token tok); wenzelm@42242: wenzelm@42242: val ast = Exn.interruptible_capture ast_of parsetree; wenzelm@42242: in (! reports, ast) end; wenzelm@42242: wenzelm@42242: wenzelm@42242: (* ast_to_term *) wenzelm@42242: wenzelm@42242: fun ast_to_term ctxt trf = wenzelm@42242: let wenzelm@42242: fun trans a args = wenzelm@42242: (case trf a of wenzelm@42242: NONE => Term.list_comb (Lexicon.const a, args) wenzelm@42242: | SOME f => f ctxt args); wenzelm@42242: wenzelm@42242: fun term_of (Ast.Constant a) = trans a [] wenzelm@42242: | term_of (Ast.Variable x) = Lexicon.read_var x wenzelm@42242: | term_of (Ast.Appl (Ast.Constant a :: (asts as _ :: _))) = wenzelm@42242: trans a (map term_of asts) wenzelm@42242: | term_of (Ast.Appl (ast :: (asts as _ :: _))) = wenzelm@42242: Term.list_comb (term_of ast, map term_of asts) wenzelm@42242: | term_of (ast as Ast.Appl _) = raise Ast.AST ("ast_to_term: malformed ast", [ast]); wenzelm@42242: in term_of end; wenzelm@42242: wenzelm@42242: wenzelm@42242: (* decode_term -- transform parse tree into raw term *) wenzelm@42242: wenzelm@42242: fun markup_bound def id = wenzelm@42242: [Markup.properties [(if def then Markup.defN else Markup.refN, id)] Markup.bound]; wenzelm@42242: wenzelm@42242: fun decode_term _ (result as (_: Position.reports, Exn.Exn _)) = result wenzelm@42242: | decode_term ctxt (reports0, Exn.Result tm) = wenzelm@42242: let wenzelm@42242: val {get_const, get_free, markup_const, markup_free, markup_var} = wenzelm@42242: ProofContext.term_context ctxt; wenzelm@42242: val decodeT = typ_of_term (ProofContext.get_sort ctxt (term_sorts tm)); wenzelm@42242: wenzelm@42242: val reports = Unsynchronized.ref reports0; wenzelm@42242: fun report ps = Position.reports reports ps; wenzelm@42242: wenzelm@42242: fun decode ps qs bs (Const ("_constrain", _) $ t $ typ) = wenzelm@42242: (case Syntax.decode_position_term typ of wenzelm@42242: SOME p => decode (p :: ps) qs bs t wenzelm@42242: | NONE => Type.constraint (decodeT typ) (decode ps qs bs t)) wenzelm@42242: | decode ps qs bs (Const ("_constrainAbs", _) $ t $ typ) = wenzelm@42242: (case Syntax.decode_position_term typ of wenzelm@42242: SOME q => decode ps (q :: qs) bs t wenzelm@42242: | NONE => Type.constraint (decodeT typ --> dummyT) (decode ps qs bs t)) wenzelm@42242: | decode _ qs bs (Abs (x, T, t)) = wenzelm@42242: let wenzelm@42242: val id = serial_string (); wenzelm@42242: val _ = report qs (markup_bound true) id; wenzelm@42242: in Abs (x, T, decode [] [] (id :: bs) t) end wenzelm@42242: | decode _ _ bs (t $ u) = decode [] [] bs t $ decode [] [] bs u wenzelm@42242: | decode ps _ _ (Const (a, T)) = wenzelm@42242: (case try Lexicon.unmark_fixed a of wenzelm@42242: SOME x => (report ps markup_free x; Free (x, T)) wenzelm@42242: | NONE => wenzelm@42242: let wenzelm@42242: val c = wenzelm@42242: (case try Lexicon.unmark_const a of wenzelm@42242: SOME c => c wenzelm@42242: | NONE => snd (get_const a)); wenzelm@42242: val _ = report ps markup_const c; wenzelm@42242: in Const (c, T) end) wenzelm@42242: | decode ps _ _ (Free (a, T)) = wenzelm@42242: (case (get_free a, get_const a) of wenzelm@42242: (SOME x, _) => (report ps markup_free x; Free (x, T)) wenzelm@42242: | (_, (true, c)) => (report ps markup_const c; Const (c, T)) wenzelm@42242: | (_, (false, c)) => wenzelm@42242: if Long_Name.is_qualified c wenzelm@42242: then (report ps markup_const c; Const (c, T)) wenzelm@42242: else (report ps markup_free c; Free (c, T))) wenzelm@42242: | decode ps _ _ (Var (xi, T)) = (report ps markup_var xi; Var (xi, T)) wenzelm@42242: | decode ps _ bs (t as Bound i) = wenzelm@42242: (case try (nth bs) i of wenzelm@42242: SOME id => (report ps (markup_bound false) id; t) wenzelm@42242: | NONE => t); wenzelm@42242: wenzelm@42242: val tm' = Exn.interruptible_capture (fn () => decode [] [] [] tm) (); wenzelm@42242: in (! reports, tm') end; wenzelm@42242: wenzelm@42242: wenzelm@42242: wenzelm@42242: (** parse **) wenzelm@42242: wenzelm@42242: (* results *) wenzelm@42242: wenzelm@42242: fun ambiguity_msg pos = "Parse error: ambiguous syntax" ^ Position.str_of pos; wenzelm@42242: wenzelm@42242: fun proper_results results = map_filter (fn (y, Exn.Result x) => SOME (y, x) | _ => NONE) results; wenzelm@42242: fun failed_results results = map_filter (fn (y, Exn.Exn e) => SOME (y, e) | _ => NONE) results; wenzelm@42242: wenzelm@42242: fun report ctxt = List.app (fn (pos, m) => Context_Position.report ctxt pos m); wenzelm@42242: wenzelm@42242: fun report_result ctxt pos results = wenzelm@42242: (case (proper_results results, failed_results results) of wenzelm@42242: ([], (reports, exn) :: _) => (report ctxt reports; reraise exn) wenzelm@42242: | ([(reports, x)], _) => (report ctxt reports; x) wenzelm@42242: | _ => error (ambiguity_msg pos)); wenzelm@42242: wenzelm@42242: wenzelm@42242: (* parse_asts *) wenzelm@42242: wenzelm@42242: fun parse_asts ctxt raw root (syms, pos) = wenzelm@42242: let wenzelm@42242: val {lexicon, gram, parse_ast_trtab, ...} = Syntax.rep_syntax (ProofContext.syn_of ctxt); wenzelm@42242: wenzelm@42242: val toks = Lexicon.tokenize lexicon raw syms; wenzelm@42242: val _ = List.app (Lexicon.report_token ctxt) toks; wenzelm@42242: wenzelm@42242: val pts = Parser.parse ctxt gram root (filter Lexicon.is_proper toks) wenzelm@42242: handle ERROR msg => wenzelm@42242: error (msg ^ wenzelm@42242: implode (map (Markup.markup Markup.report o Lexicon.reported_token_range ctxt) toks)); wenzelm@42242: val len = length pts; wenzelm@42242: wenzelm@42242: val limit = Config.get ctxt Syntax.ambiguity_limit; wenzelm@42242: val _ = wenzelm@42242: if len <= Config.get ctxt Syntax.ambiguity_level then () wenzelm@42242: else if not (Config.get ctxt Syntax.ambiguity_enabled) then error (ambiguity_msg pos) wenzelm@42242: else wenzelm@42242: (Context_Position.if_visible ctxt warning (cat_lines wenzelm@42242: (("Ambiguous input" ^ Position.str_of pos ^ wenzelm@42242: "\nproduces " ^ string_of_int len ^ " parse trees" ^ wenzelm@42242: (if len <= limit then "" else " (" ^ string_of_int limit ^ " displayed)") ^ ":") :: wenzelm@42242: map (Pretty.string_of o Parser.pretty_parsetree) (take limit pts)))); wenzelm@42242: wenzelm@42242: val constrain_pos = not raw andalso Config.get ctxt Syntax.positions; wenzelm@42242: val parsetree_to_ast = parsetree_to_ast ctxt constrain_pos (lookup_tr parse_ast_trtab); wenzelm@42242: in map parsetree_to_ast pts end; wenzelm@42242: wenzelm@42242: wenzelm@42242: (* read_raw *) wenzelm@42242: wenzelm@42242: fun read_raw ctxt root input = wenzelm@42242: let wenzelm@42242: val {parse_ruletab, parse_trtab, ...} = Syntax.rep_syntax (ProofContext.syn_of ctxt); wenzelm@42242: val norm = Ast.normalize ctxt (Symtab.lookup_list parse_ruletab); wenzelm@42242: val ast_to_term = ast_to_term ctxt (lookup_tr parse_trtab); wenzelm@42242: in wenzelm@42242: parse_asts ctxt false root input wenzelm@42242: |> (map o apsnd o Exn.maps_result) (norm #> Exn.interruptible_capture ast_to_term) wenzelm@42242: end; wenzelm@42242: wenzelm@42242: wenzelm@42242: (* read sorts *) wenzelm@42242: wenzelm@42242: fun standard_parse_sort ctxt (syms, pos) = wenzelm@42242: read_raw ctxt "sort" (syms, pos) wenzelm@42242: |> report_result ctxt pos wenzelm@42242: |> sort_of_term; wenzelm@42242: wenzelm@42242: wenzelm@42242: (* read types *) wenzelm@42242: wenzelm@42242: fun standard_parse_typ ctxt (syms, pos) = wenzelm@42242: read_raw ctxt "type" (syms, pos) wenzelm@42242: |> report_result ctxt pos wenzelm@42242: |> (fn t => typ_of_term (ProofContext.get_sort ctxt (term_sorts t)) t); wenzelm@42242: wenzelm@42242: wenzelm@42242: (* read terms -- brute-force disambiguation via type-inference *) wenzelm@42242: wenzelm@42242: fun standard_parse_term check ctxt root (syms, pos) = wenzelm@42242: let wenzelm@42242: val results = read_raw ctxt root (syms, pos) |> map (decode_term ctxt); wenzelm@42242: wenzelm@42242: val level = Config.get ctxt Syntax.ambiguity_level; wenzelm@42242: val limit = Config.get ctxt Syntax.ambiguity_limit; wenzelm@42242: wenzelm@42242: val ambiguity = length (proper_results results); wenzelm@42242: wenzelm@42242: fun ambig_msg () = wenzelm@42242: if ambiguity > 1 andalso ambiguity <= level then wenzelm@42242: "Got more than one parse tree.\n\ wenzelm@42242: \Retry with smaller syntax_ambiguity_level for more information." wenzelm@42242: else ""; wenzelm@42242: wenzelm@42242: val results' = wenzelm@42242: if ambiguity > 1 then wenzelm@42242: (Par_List.map_name "Syntax.disambig" o apsnd o Exn.maps_result) check results wenzelm@42242: else results; wenzelm@42242: val reports' = fst (hd results'); wenzelm@42242: wenzelm@42242: val errs = map snd (failed_results results'); wenzelm@42242: val checked = map snd (proper_results results'); wenzelm@42242: val len = length checked; wenzelm@42242: wenzelm@42242: val show_term = Syntax.string_of_term (Config.put Syntax.show_brackets true ctxt); wenzelm@42242: in wenzelm@42242: if len = 0 then wenzelm@42242: report_result ctxt pos wenzelm@42242: [(reports', Exn.Exn (Exn.EXCEPTIONS (ERROR (ambig_msg ()) :: errs)))] wenzelm@42242: else if len = 1 then wenzelm@42242: (if ambiguity > level then wenzelm@42242: Context_Position.if_visible ctxt warning wenzelm@42242: "Fortunately, only one parse tree is type correct.\n\ wenzelm@42242: \You may still want to disambiguate your grammar or your input." wenzelm@42242: else (); report_result ctxt pos results') wenzelm@42242: else wenzelm@42242: report_result ctxt pos wenzelm@42242: [(reports', Exn.Exn (ERROR (cat_lines (ambig_msg () :: wenzelm@42242: (("Ambiguous input, " ^ string_of_int len ^ " terms are type correct" ^ wenzelm@42242: (if len <= limit then "" else " (" ^ string_of_int limit ^ " displayed)") ^ ":") :: wenzelm@42242: map show_term (take limit checked))))))] wenzelm@42242: end; wenzelm@42242: wenzelm@42242: wenzelm@42242: (* standard operations *) wenzelm@42242: wenzelm@42241: fun parse_failed ctxt pos msg kind = wenzelm@42241: cat_error msg ("Failed to parse " ^ kind ^ wenzelm@42241: Markup.markup Markup.report (Context_Position.reported_text ctxt pos Markup.bad "")); wenzelm@42241: wenzelm@42241: fun parse_sort ctxt text = wenzelm@42241: let wenzelm@42241: val (syms, pos) = Syntax.parse_token ctxt Markup.sort text; wenzelm@42242: val S = standard_parse_sort ctxt (syms, pos) wenzelm@42242: handle ERROR msg => parse_failed ctxt pos msg "sort"; wenzelm@42241: in Type.minimize_sort (ProofContext.tsig_of ctxt) S end; wenzelm@42241: wenzelm@42241: fun parse_typ ctxt text = wenzelm@42241: let wenzelm@42241: val (syms, pos) = Syntax.parse_token ctxt Markup.typ text; wenzelm@42242: val T = standard_parse_typ ctxt (syms, pos) wenzelm@42242: handle ERROR msg => parse_failed ctxt pos msg "type"; wenzelm@42241: in T end; wenzelm@42241: wenzelm@42241: fun parse_term T ctxt text = wenzelm@42241: let wenzelm@42241: val (T', _) = Type_Infer.paramify_dummies T 0; wenzelm@42241: val (markup, kind) = wenzelm@42241: if T' = propT then (Markup.prop, "proposition") else (Markup.term, "term"); wenzelm@42241: val (syms, pos) = Syntax.parse_token ctxt markup text; wenzelm@42241: wenzelm@42241: val default_root = Config.get ctxt Syntax.default_root; wenzelm@42241: val root = wenzelm@42241: (case T' of wenzelm@42241: Type (c, _) => wenzelm@42241: if c <> "prop" andalso Type.is_logtype (ProofContext.tsig_of ctxt) c wenzelm@42241: then default_root else c wenzelm@42241: | _ => default_root); wenzelm@42241: wenzelm@42241: fun check t = (Syntax.check_term ctxt (Type.constraint T' t); Exn.Result t) wenzelm@42241: handle exn as ERROR _ => Exn.Exn exn; wenzelm@42242: val t = standard_parse_term check ctxt root (syms, pos) wenzelm@42241: handle ERROR msg => parse_failed ctxt pos msg kind; wenzelm@42241: in t end; wenzelm@42241: wenzelm@42241: wenzelm@42242: (* parse_ast_pattern *) wenzelm@42242: wenzelm@42242: fun parse_ast_pattern ctxt (root, str) = wenzelm@42242: let wenzelm@42242: val syn = ProofContext.syn_of ctxt; wenzelm@42242: wenzelm@42242: fun constify (ast as Ast.Constant _) = ast wenzelm@42242: | constify (ast as Ast.Variable x) = wenzelm@42242: if Syntax.is_const syn x orelse Long_Name.is_qualified x then Ast.Constant x wenzelm@42242: else ast wenzelm@42242: | constify (Ast.Appl asts) = Ast.Appl (map constify asts); wenzelm@42242: wenzelm@42242: val (syms, pos) = Syntax.read_token str; wenzelm@42242: in wenzelm@42242: parse_asts ctxt true root (syms, pos) wenzelm@42242: |> report_result ctxt pos wenzelm@42242: |> constify wenzelm@42242: end; wenzelm@42242: wenzelm@42242: wenzelm@42242: wenzelm@42245: (** encode parse trees **) wenzelm@42245: wenzelm@42245: (* term_of_sort *) wenzelm@42245: wenzelm@42245: fun term_of_sort S = wenzelm@42245: let wenzelm@42245: val class = Lexicon.const o Lexicon.mark_class; wenzelm@42245: wenzelm@42245: fun classes [c] = class c wenzelm@42245: | classes (c :: cs) = Lexicon.const "_classes" $ class c $ classes cs; wenzelm@42245: in wenzelm@42245: (case S of wenzelm@42245: [] => Lexicon.const "_topsort" wenzelm@42245: | [c] => class c wenzelm@42245: | cs => Lexicon.const "_sort" $ classes cs) wenzelm@42245: end; wenzelm@42245: wenzelm@42245: wenzelm@42245: (* term_of_typ *) wenzelm@42245: wenzelm@42245: fun term_of_typ show_sorts ty = wenzelm@42245: let wenzelm@42245: fun of_sort t S = wenzelm@42245: if show_sorts then Lexicon.const "_ofsort" $ t $ term_of_sort S wenzelm@42245: else t; wenzelm@42245: wenzelm@42245: fun term_of (Type (a, Ts)) = wenzelm@42245: Term.list_comb (Lexicon.const (Lexicon.mark_type a), map term_of Ts) wenzelm@42245: | term_of (TFree (x, S)) = wenzelm@42245: if is_some (Lexicon.decode_position x) then Lexicon.free x wenzelm@42245: else of_sort (Lexicon.const "_tfree" $ Lexicon.free x) S wenzelm@42245: | term_of (TVar (xi, S)) = of_sort (Lexicon.const "_tvar" $ Lexicon.var xi) S; wenzelm@42245: in term_of ty end; wenzelm@42245: wenzelm@42245: wenzelm@42245: (* simple_ast_of *) wenzelm@42245: wenzelm@42245: fun simple_ast_of ctxt = wenzelm@42245: let wenzelm@42245: val tune_var = if Config.get ctxt show_question_marks then I else unprefix "?"; wenzelm@42245: fun ast_of (Const (c, _)) = Ast.Constant c wenzelm@42245: | ast_of (Free (x, _)) = Ast.Variable x wenzelm@42245: | ast_of (Var (xi, _)) = Ast.Variable (tune_var (Term.string_of_vname xi)) wenzelm@42245: | ast_of (t as _ $ _) = wenzelm@42245: let val (f, args) = strip_comb t wenzelm@42245: in Ast.mk_appl (ast_of f) (map ast_of args) end wenzelm@42245: | ast_of (Bound i) = Ast.Variable ("B." ^ string_of_int i) wenzelm@42245: | ast_of (Abs _) = raise Fail "simple_ast_of: Abs"; wenzelm@42245: in ast_of end; wenzelm@42245: wenzelm@42245: wenzelm@42245: (* sort_to_ast and typ_to_ast *) wenzelm@42245: wenzelm@42245: fun apply_typed x fs = map (fn f => fn ctxt => f ctxt (Config.get ctxt show_sorts) x) fs; wenzelm@42245: wenzelm@42245: fun ast_of_termT ctxt trf tm = wenzelm@42245: let wenzelm@42245: val ctxt' = Config.put show_sorts false ctxt; wenzelm@42245: fun ast_of (t as Const ("_tfree", _) $ Free _) = simple_ast_of ctxt t wenzelm@42245: | ast_of (t as Const ("_tvar", _) $ Var _) = simple_ast_of ctxt t wenzelm@42245: | ast_of (Const (a, _)) = trans a [] wenzelm@42245: | ast_of (t as _ $ _) = wenzelm@42245: (case strip_comb t of wenzelm@42245: (Const (a, _), args) => trans a args wenzelm@42245: | (f, args) => Ast.Appl (map ast_of (f :: args))) wenzelm@42245: | ast_of t = simple_ast_of ctxt t wenzelm@42245: and trans a args = wenzelm@42245: ast_of (Syntax.apply_trans ctxt' (apply_typed dummyT (trf a)) args) wenzelm@42245: handle Match => Ast.mk_appl (Ast.Constant a) (map ast_of args); wenzelm@42245: in ast_of tm end; wenzelm@42245: wenzelm@42245: fun sort_to_ast ctxt trf S = ast_of_termT ctxt trf (term_of_sort S); wenzelm@42245: fun typ_to_ast ctxt trf T = ast_of_termT ctxt trf (term_of_typ (Config.get ctxt show_sorts) T); wenzelm@42245: wenzelm@42245: wenzelm@42245: (* term_to_ast *) wenzelm@42245: wenzelm@42245: fun term_to_ast idents consts ctxt trf tm = wenzelm@42245: let wenzelm@42245: val show_types = wenzelm@42245: Config.get ctxt show_types orelse Config.get ctxt show_sorts orelse wenzelm@42245: Config.get ctxt show_all_types; wenzelm@42245: val show_sorts = Config.get ctxt show_sorts; wenzelm@42245: val show_structs = Config.get ctxt show_structs; wenzelm@42245: val show_free_types = Config.get ctxt show_free_types; wenzelm@42245: val show_all_types = Config.get ctxt show_all_types; wenzelm@42245: wenzelm@42245: val {structs, fixes} = idents; wenzelm@42245: wenzelm@42245: fun mark_atoms ((t as Const (c, T)) $ u) = wenzelm@42245: if member (op =) Syntax.standard_token_markers c wenzelm@42245: then t $ u else mark_atoms t $ mark_atoms u wenzelm@42245: | mark_atoms (t $ u) = mark_atoms t $ mark_atoms u wenzelm@42245: | mark_atoms (Abs (x, T, t)) = Abs (x, T, mark_atoms t) wenzelm@42245: | mark_atoms (t as Const (c, T)) = wenzelm@42245: if member (op =) consts c then t wenzelm@42245: else Const (Lexicon.mark_const c, T) wenzelm@42245: | mark_atoms (t as Free (x, T)) = wenzelm@42245: let val i = find_index (fn s => s = x) structs + 1 in wenzelm@42245: if i = 0 andalso member (op =) fixes x then wenzelm@42245: Const (Lexicon.mark_fixed x, T) wenzelm@42245: else if i = 1 andalso not show_structs then wenzelm@42245: Lexicon.const "_struct" $ Lexicon.const "_indexdefault" wenzelm@42245: else Lexicon.const "_free" $ t wenzelm@42245: end wenzelm@42245: | mark_atoms (t as Var (xi, T)) = wenzelm@42245: if xi = Syntax.dddot_indexname then Const ("_DDDOT", T) wenzelm@42245: else Lexicon.const "_var" $ t wenzelm@42245: | mark_atoms a = a; wenzelm@42245: wenzelm@42245: fun prune_typs (t_seen as (Const _, _)) = t_seen wenzelm@42245: | prune_typs (t as Free (x, ty), seen) = wenzelm@42245: if ty = dummyT then (t, seen) wenzelm@42245: else if not show_free_types orelse member (op aconv) seen t then (Lexicon.free x, seen) wenzelm@42245: else (t, t :: seen) wenzelm@42245: | prune_typs (t as Var (xi, ty), seen) = wenzelm@42245: if ty = dummyT then (t, seen) wenzelm@42245: else if not show_free_types orelse member (op aconv) seen t then (Lexicon.var xi, seen) wenzelm@42245: else (t, t :: seen) wenzelm@42245: | prune_typs (t_seen as (Bound _, _)) = t_seen wenzelm@42245: | prune_typs (Abs (x, ty, t), seen) = wenzelm@42245: let val (t', seen') = prune_typs (t, seen); wenzelm@42245: in (Abs (x, ty, t'), seen') end wenzelm@42245: | prune_typs (t1 $ t2, seen) = wenzelm@42245: let wenzelm@42245: val (t1', seen') = prune_typs (t1, seen); wenzelm@42245: val (t2', seen'') = prune_typs (t2, seen'); wenzelm@42245: in (t1' $ t2', seen'') end; wenzelm@42245: wenzelm@42245: fun ast_of tm = wenzelm@42245: (case strip_comb tm of wenzelm@42245: (t as Abs _, ts) => Ast.mk_appl (ast_of (Syn_Trans.abs_tr' ctxt t)) (map ast_of ts) wenzelm@42245: | ((c as Const ("_free", _)), Free (x, T) :: ts) => wenzelm@42245: Ast.mk_appl (constrain (c $ Lexicon.free x) T) (map ast_of ts) wenzelm@42245: | ((c as Const ("_var", _)), Var (xi, T) :: ts) => wenzelm@42245: Ast.mk_appl (constrain (c $ Lexicon.var xi) T) (map ast_of ts) wenzelm@42245: | ((c as Const ("_bound", _)), Free (x, T) :: ts) => wenzelm@42245: Ast.mk_appl (constrain (c $ Lexicon.free x) T) (map ast_of ts) wenzelm@42245: | (Const ("_idtdummy", T), ts) => wenzelm@42245: Ast.mk_appl (constrain (Lexicon.const "_idtdummy") T) (map ast_of ts) wenzelm@42245: | (const as Const (c, T), ts) => wenzelm@42245: if show_all_types wenzelm@42245: then Ast.mk_appl (constrain const T) (map ast_of ts) wenzelm@42245: else trans c T ts wenzelm@42245: | (t, ts) => Ast.mk_appl (simple_ast_of ctxt t) (map ast_of ts)) wenzelm@42245: wenzelm@42245: and trans a T args = wenzelm@42245: ast_of (Syntax.apply_trans ctxt (apply_typed T (trf a)) args) wenzelm@42245: handle Match => Ast.mk_appl (Ast.Constant a) (map ast_of args) wenzelm@42245: wenzelm@42245: and constrain t T = wenzelm@42245: if show_types andalso T <> dummyT then wenzelm@42245: Ast.Appl [Ast.Constant Syntax.constrainC, simple_ast_of ctxt t, wenzelm@42245: ast_of_termT ctxt trf (term_of_typ show_sorts T)] wenzelm@42245: else simple_ast_of ctxt t; wenzelm@42245: in wenzelm@42245: tm wenzelm@42245: |> Syn_Trans.prop_tr' wenzelm@42245: |> show_types ? (#1 o prune_typs o rpair []) wenzelm@42245: |> mark_atoms wenzelm@42245: |> ast_of wenzelm@42245: end; wenzelm@42245: wenzelm@42245: wenzelm@42245: wenzelm@42242: (** unparse **) wenzelm@42242: wenzelm@42245: (** unparse terms, typs, sorts **) wenzelm@42245: wenzelm@42245: local wenzelm@42245: wenzelm@42245: fun lookup_tr' tab c = map fst (Symtab.lookup_list tab c); wenzelm@42245: wenzelm@42245: fun unparse_t t_to_ast prt_t markup ctxt curried t = wenzelm@42245: let wenzelm@42245: val {consts, print_trtab, print_ruletab, print_ast_trtab, tokentrtab, prtabs, ...} = wenzelm@42245: Syntax.rep_syntax (ProofContext.syn_of ctxt); wenzelm@42245: val ast = t_to_ast consts ctxt (lookup_tr' print_trtab) t; wenzelm@42245: in wenzelm@42245: Pretty.markup markup (prt_t ctxt curried prtabs (lookup_tr' print_ast_trtab) wenzelm@42245: (Syntax.lookup_tokentr tokentrtab (print_mode_value ())) wenzelm@42245: (Ast.normalize ctxt (Symtab.lookup_list print_ruletab) ast)) wenzelm@42245: end; wenzelm@42245: wenzelm@42245: in wenzelm@42245: wenzelm@42245: fun standard_unparse_sort {extern_class} ctxt = wenzelm@42245: unparse_t (K sort_to_ast) wenzelm@42245: (Printer.pretty_typ_ast {extern_class = extern_class, extern_type = I}) wenzelm@42245: Markup.sort ctxt false; wenzelm@42245: wenzelm@42245: fun standard_unparse_typ extern ctxt = wenzelm@42245: unparse_t (K typ_to_ast) (Printer.pretty_typ_ast extern) Markup.typ ctxt false; wenzelm@42245: wenzelm@42245: fun standard_unparse_term idents extern = wenzelm@42245: unparse_t (term_to_ast idents) (Printer.pretty_term_ast extern) Markup.term; wenzelm@42245: wenzelm@42245: end; wenzelm@42245: wenzelm@42245: wenzelm@42241: fun unparse_sort ctxt = wenzelm@42245: standard_unparse_sort {extern_class = Type.extern_class (ProofContext.tsig_of ctxt)} ctxt; wenzelm@42241: wenzelm@42241: fun unparse_typ ctxt = wenzelm@42241: let wenzelm@42241: val tsig = ProofContext.tsig_of ctxt; wenzelm@42241: val extern = {extern_class = Type.extern_class tsig, extern_type = Type.extern_type tsig}; wenzelm@42245: in standard_unparse_typ extern ctxt end; wenzelm@42241: wenzelm@42241: fun unparse_term ctxt = wenzelm@42241: let wenzelm@42241: val tsig = ProofContext.tsig_of ctxt; wenzelm@42241: val syntax = ProofContext.syntax_of ctxt; wenzelm@42241: val consts = ProofContext.consts_of ctxt; wenzelm@42241: val extern = wenzelm@42241: {extern_class = Type.extern_class tsig, wenzelm@42241: extern_type = Type.extern_type tsig, wenzelm@42241: extern_const = Consts.extern consts}; wenzelm@42241: in wenzelm@42245: standard_unparse_term (Local_Syntax.idents_of syntax) extern ctxt wenzelm@42245: (not (Pure_Thy.old_appl_syntax (ProofContext.theory_of ctxt))) wenzelm@42241: end; wenzelm@42241: wenzelm@42241: wenzelm@42245: wenzelm@42245: (** translations **) wenzelm@42245: wenzelm@42245: (* type propositions *) wenzelm@42245: wenzelm@42245: fun type_prop_tr' _ T [Const ("\\<^const>Pure.sort_constraint", _)] = wenzelm@42245: Lexicon.const "_sort_constraint" $ term_of_typ true T wenzelm@42245: | type_prop_tr' show_sorts T [t] = wenzelm@42245: Lexicon.const "_ofclass" $ term_of_typ show_sorts T $ t wenzelm@42245: | type_prop_tr' _ T ts = raise TYPE ("type_prop_tr'", [T], ts); wenzelm@42245: wenzelm@42245: wenzelm@42245: (* type reflection *) wenzelm@42245: wenzelm@42245: fun type_tr' show_sorts (Type ("itself", [T])) ts = wenzelm@42245: Term.list_comb (Lexicon.const "_TYPE" $ term_of_typ show_sorts T, ts) wenzelm@42245: | type_tr' _ _ _ = raise Match; wenzelm@42245: wenzelm@42245: wenzelm@42245: (* type constraints *) wenzelm@42245: wenzelm@42245: fun type_constraint_tr' show_sorts (Type ("fun", [T, _])) (t :: ts) = wenzelm@42245: Term.list_comb (Lexicon.const Syntax.constrainC $ t $ term_of_typ show_sorts T, ts) wenzelm@42245: | type_constraint_tr' _ _ _ = raise Match; wenzelm@42245: wenzelm@42245: wenzelm@42245: (* setup translations *) wenzelm@42245: wenzelm@42245: val _ = Context.>> (Context.map_theory wenzelm@42245: (Sign.add_trfunsT wenzelm@42245: [("_type_prop", type_prop_tr'), wenzelm@42245: ("\\<^const>TYPE", type_tr'), wenzelm@42245: ("_type_constraint_", type_constraint_tr')])); wenzelm@42245: wenzelm@42245: wenzelm@42245: wenzelm@42242: (** install operations **) wenzelm@42242: wenzelm@42241: val _ = Syntax.install_operations wenzelm@42241: {parse_sort = parse_sort, wenzelm@42241: parse_typ = parse_typ, wenzelm@42241: parse_term = parse_term dummyT, wenzelm@42241: parse_prop = parse_term propT, wenzelm@42241: unparse_sort = unparse_sort, wenzelm@42241: unparse_typ = unparse_typ, wenzelm@42241: unparse_term = unparse_term}; wenzelm@42241: wenzelm@42241: end;