src/Pure/Syntax/term_position.ML
author wenzelm
Fri, 16 Oct 2015 14:53:26 +0200
changeset 61458 987533262fc2
parent 52175 626a757d3c2d
child 80873 e71cb37c7395
permissions -rw-r--r--
Markdown support in document text;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
42264
b6c1b0c4c511 separate structure Term_Position;
wenzelm
parents:
diff changeset
     1
(*  Title:      Pure/Syntax/term_position.ML
b6c1b0c4c511 separate structure Term_Position;
wenzelm
parents:
diff changeset
     2
    Author:     Makarius
b6c1b0c4c511 separate structure Term_Position;
wenzelm
parents:
diff changeset
     3
b6c1b0c4c511 separate structure Term_Position;
wenzelm
parents:
diff changeset
     4
Encoded position within term syntax trees.
b6c1b0c4c511 separate structure Term_Position;
wenzelm
parents:
diff changeset
     5
*)
b6c1b0c4c511 separate structure Term_Position;
wenzelm
parents:
diff changeset
     6
b6c1b0c4c511 separate structure Term_Position;
wenzelm
parents:
diff changeset
     7
signature TERM_POSITION =
b6c1b0c4c511 separate structure Term_Position;
wenzelm
parents:
diff changeset
     8
sig
b6c1b0c4c511 separate structure Term_Position;
wenzelm
parents:
diff changeset
     9
  val pretty: Position.T -> Pretty.T
b6c1b0c4c511 separate structure Term_Position;
wenzelm
parents:
diff changeset
    10
  val encode: Position.T -> string
b6c1b0c4c511 separate structure Term_Position;
wenzelm
parents:
diff changeset
    11
  val decode: string -> Position.T option
45445
41e641a870de pass term positions into check phase, where resulting types are reported accordingly, and eventually shown as tooltips;
wenzelm
parents: 42264
diff changeset
    12
  val decode_position: term -> (Position.T * typ) option
41e641a870de pass term positions into check phase, where resulting types are reported accordingly, and eventually shown as tooltips;
wenzelm
parents: 42264
diff changeset
    13
  val decode_positionT: typ -> Position.T option
49685
4341e4d86872 allow position constraints to coexist with 0 or 1 sort constraints;
wenzelm
parents: 49674
diff changeset
    14
  val decode_positionS: sort -> Position.T list * sort
42264
b6c1b0c4c511 separate structure Term_Position;
wenzelm
parents:
diff changeset
    15
  val is_position: term -> bool
45448
018f8959c7a6 more efficient prepare_sorts -- bypass encoded positions;
wenzelm
parents: 45445
diff changeset
    16
  val is_positionT: typ -> bool
52175
626a757d3c2d uniform Term_Position.markers (cf. dbadb4d03cbc);
wenzelm
parents: 50201
diff changeset
    17
  val markers: string list
42264
b6c1b0c4c511 separate structure Term_Position;
wenzelm
parents:
diff changeset
    18
  val strip_positions: term -> term
b6c1b0c4c511 separate structure Term_Position;
wenzelm
parents:
diff changeset
    19
end;
b6c1b0c4c511 separate structure Term_Position;
wenzelm
parents:
diff changeset
    20
b6c1b0c4c511 separate structure Term_Position;
wenzelm
parents:
diff changeset
    21
structure Term_Position: TERM_POSITION =
b6c1b0c4c511 separate structure Term_Position;
wenzelm
parents:
diff changeset
    22
struct
b6c1b0c4c511 separate structure Term_Position;
wenzelm
parents:
diff changeset
    23
b6c1b0c4c511 separate structure Term_Position;
wenzelm
parents:
diff changeset
    24
(* markup *)
b6c1b0c4c511 separate structure Term_Position;
wenzelm
parents:
diff changeset
    25
b6c1b0c4c511 separate structure Term_Position;
wenzelm
parents:
diff changeset
    26
val position_dummy = "<position>";
b6c1b0c4c511 separate structure Term_Position;
wenzelm
parents:
diff changeset
    27
val position_text = XML.Text position_dummy;
b6c1b0c4c511 separate structure Term_Position;
wenzelm
parents:
diff changeset
    28
b6c1b0c4c511 separate structure Term_Position;
wenzelm
parents:
diff changeset
    29
fun pretty pos =
50201
c26369c9eda6 Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
wenzelm
parents: 49685
diff changeset
    30
  Pretty.markup (Position.markup pos Markup.position) [Pretty.str position_dummy];
42264
b6c1b0c4c511 separate structure Term_Position;
wenzelm
parents:
diff changeset
    31
b6c1b0c4c511 separate structure Term_Position;
wenzelm
parents:
diff changeset
    32
fun encode pos =
50201
c26369c9eda6 Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
wenzelm
parents: 49685
diff changeset
    33
  YXML.string_of (XML.Elem (Position.markup pos Markup.position, [position_text]));
42264
b6c1b0c4c511 separate structure Term_Position;
wenzelm
parents:
diff changeset
    34
b6c1b0c4c511 separate structure Term_Position;
wenzelm
parents:
diff changeset
    35
fun decode str =
b6c1b0c4c511 separate structure Term_Position;
wenzelm
parents:
diff changeset
    36
  (case YXML.parse_body str handle Fail msg => error msg of
b6c1b0c4c511 separate structure Term_Position;
wenzelm
parents:
diff changeset
    37
    [XML.Elem ((name, props), [arg])] =>
50201
c26369c9eda6 Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
wenzelm
parents: 49685
diff changeset
    38
      if name = Markup.positionN andalso arg = position_text
42264
b6c1b0c4c511 separate structure Term_Position;
wenzelm
parents:
diff changeset
    39
      then SOME (Position.of_properties props)
b6c1b0c4c511 separate structure Term_Position;
wenzelm
parents:
diff changeset
    40
      else NONE
b6c1b0c4c511 separate structure Term_Position;
wenzelm
parents:
diff changeset
    41
  | _ => NONE);
b6c1b0c4c511 separate structure Term_Position;
wenzelm
parents:
diff changeset
    42
b6c1b0c4c511 separate structure Term_Position;
wenzelm
parents:
diff changeset
    43
b6c1b0c4c511 separate structure Term_Position;
wenzelm
parents:
diff changeset
    44
(* positions within parse trees *)
b6c1b0c4c511 separate structure Term_Position;
wenzelm
parents:
diff changeset
    45
45445
41e641a870de pass term positions into check phase, where resulting types are reported accordingly, and eventually shown as tooltips;
wenzelm
parents: 42264
diff changeset
    46
fun decode_position (Free (x, _)) =
41e641a870de pass term positions into check phase, where resulting types are reported accordingly, and eventually shown as tooltips;
wenzelm
parents: 42264
diff changeset
    47
      (case decode x of
41e641a870de pass term positions into check phase, where resulting types are reported accordingly, and eventually shown as tooltips;
wenzelm
parents: 42264
diff changeset
    48
        SOME pos => SOME (pos, TFree (x, dummyS))
41e641a870de pass term positions into check phase, where resulting types are reported accordingly, and eventually shown as tooltips;
wenzelm
parents: 42264
diff changeset
    49
      | NONE => NONE)
42264
b6c1b0c4c511 separate structure Term_Position;
wenzelm
parents:
diff changeset
    50
  | decode_position _ = NONE;
b6c1b0c4c511 separate structure Term_Position;
wenzelm
parents:
diff changeset
    51
45445
41e641a870de pass term positions into check phase, where resulting types are reported accordingly, and eventually shown as tooltips;
wenzelm
parents: 42264
diff changeset
    52
fun decode_positionT (TFree (x, _)) = decode x
41e641a870de pass term positions into check phase, where resulting types are reported accordingly, and eventually shown as tooltips;
wenzelm
parents: 42264
diff changeset
    53
  | decode_positionT _ = NONE;
41e641a870de pass term positions into check phase, where resulting types are reported accordingly, and eventually shown as tooltips;
wenzelm
parents: 42264
diff changeset
    54
49685
4341e4d86872 allow position constraints to coexist with 0 or 1 sort constraints;
wenzelm
parents: 49674
diff changeset
    55
fun decode_positionS cs =
4341e4d86872 allow position constraints to coexist with 0 or 1 sort constraints;
wenzelm
parents: 49674
diff changeset
    56
  let val (ps, sort) = List.partition (is_some o decode) cs
4341e4d86872 allow position constraints to coexist with 0 or 1 sort constraints;
wenzelm
parents: 49674
diff changeset
    57
  in (map (the o decode) ps, sort) end;
49674
dbadb4d03cbc report sort assignment of visible type variables;
wenzelm
parents: 45666
diff changeset
    58
42264
b6c1b0c4c511 separate structure Term_Position;
wenzelm
parents:
diff changeset
    59
val is_position = is_some o decode_position;
45448
018f8959c7a6 more efficient prepare_sorts -- bypass encoded positions;
wenzelm
parents: 45445
diff changeset
    60
val is_positionT = is_some o decode_positionT;
42264
b6c1b0c4c511 separate structure Term_Position;
wenzelm
parents:
diff changeset
    61
52175
626a757d3c2d uniform Term_Position.markers (cf. dbadb4d03cbc);
wenzelm
parents: 50201
diff changeset
    62
val markers = ["_constrain", "_constrainAbs", "_ofsort"];
626a757d3c2d uniform Term_Position.markers (cf. dbadb4d03cbc);
wenzelm
parents: 50201
diff changeset
    63
42264
b6c1b0c4c511 separate structure Term_Position;
wenzelm
parents:
diff changeset
    64
fun strip_positions ((t as Const (c, _)) $ u $ v) =
52175
626a757d3c2d uniform Term_Position.markers (cf. dbadb4d03cbc);
wenzelm
parents: 50201
diff changeset
    65
      if member (op =) markers c andalso is_position v
42264
b6c1b0c4c511 separate structure Term_Position;
wenzelm
parents:
diff changeset
    66
      then strip_positions u
b6c1b0c4c511 separate structure Term_Position;
wenzelm
parents:
diff changeset
    67
      else t $ strip_positions u $ strip_positions v
b6c1b0c4c511 separate structure Term_Position;
wenzelm
parents:
diff changeset
    68
  | strip_positions (t $ u) = strip_positions t $ strip_positions u
b6c1b0c4c511 separate structure Term_Position;
wenzelm
parents:
diff changeset
    69
  | strip_positions (Abs (x, T, t)) = Abs (x, T, strip_positions t)
b6c1b0c4c511 separate structure Term_Position;
wenzelm
parents:
diff changeset
    70
  | strip_positions t = t;
b6c1b0c4c511 separate structure Term_Position;
wenzelm
parents:
diff changeset
    71
b6c1b0c4c511 separate structure Term_Position;
wenzelm
parents:
diff changeset
    72
end;