src/HOL/HOLCF/Tools/cont_consts.ML
author wenzelm
Sun, 12 Sep 2021 22:31:51 +0200
changeset 74305 28a582aa25dd
parent 69597 ff784d5a5bfb
permissions -rw-r--r--
more antiquotations;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
42151
4da4fc77664b tuned headers;
wenzelm
parents: 40832
diff changeset
     1
(*  Title:      HOL/HOLCF/Tools/cont_consts.ML
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
     2
    Author:     Tobias Mayr, David von Oheimb, and Markus Wenzel
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
     3
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
     4
HOLCF version of consts: handle continuous function types in mixfix
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
     5
syntax.
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
     6
*)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
     7
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
     8
signature CONT_CONSTS =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
     9
sig
35257
3e5980f612d9 adapted to authentic syntax;
wenzelm
parents: 35130
diff changeset
    10
  val add_consts: (binding * typ * mixfix) list -> theory -> theory
3e5980f612d9 adapted to authentic syntax;
wenzelm
parents: 35130
diff changeset
    11
  val add_consts_cmd: (binding * string * mixfix) list -> theory -> theory
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    12
end
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    13
35257
3e5980f612d9 adapted to authentic syntax;
wenzelm
parents: 35130
diff changeset
    14
structure Cont_Consts: CONT_CONSTS =
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    15
struct
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    16
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    17
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    18
(* misc utils *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    19
35257
3e5980f612d9 adapted to authentic syntax;
wenzelm
parents: 35130
diff changeset
    20
fun change_arrow 0 T = T
74305
28a582aa25dd more antiquotations;
wenzelm
parents: 69597
diff changeset
    21
  | change_arrow n (Type (_, [S, T])) = \<^Type>\<open>fun S \<open>change_arrow (n - 1) T\<close>\<close>
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    22
  | change_arrow _ T = raise TYPE ("cont_consts: change_arrow", [T], [])
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    23
33245
65232054ffd0 eliminated some old folds;
wenzelm
parents: 31023
diff changeset
    24
fun trans_rules name2 name1 n mx =
65232054ffd0 eliminated some old folds;
wenzelm
parents: 31023
diff changeset
    25
  let
43329
84472e198515 tuned signature: Name.invent and Name.invent_names;
wenzelm
parents: 42381
diff changeset
    26
    val vnames = Name.invent Name.context "a" n
42224
578a51fae383 discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
wenzelm
parents: 42204
diff changeset
    27
    val extra_parse_rule = Syntax.Parse_Rule (Ast.Constant name2, Ast.Constant name1)
33245
65232054ffd0 eliminated some old folds;
wenzelm
parents: 31023
diff changeset
    28
  in
42204
b3277168c1e7 added Position.reports convenience;
wenzelm
parents: 42151
diff changeset
    29
    [Syntax.Parse_Print_Rule
42224
578a51fae383 discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
wenzelm
parents: 42204
diff changeset
    30
      (Ast.mk_appl (Ast.Constant name2) (map Ast.Variable vnames),
578a51fae383 discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
wenzelm
parents: 42204
diff changeset
    31
        fold (fn a => fn t =>
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 56239
diff changeset
    32
            Ast.mk_appl (Ast.Constant \<^const_syntax>\<open>Rep_cfun\<close>) [t, Ast.Variable a])
42224
578a51fae383 discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
wenzelm
parents: 42204
diff changeset
    33
          vnames (Ast.Constant name1))] @
33245
65232054ffd0 eliminated some old folds;
wenzelm
parents: 31023
diff changeset
    34
    (case mx of
35130
0991c84e8dcf renamed InfixName to Infix etc.;
wenzelm
parents: 35129
diff changeset
    35
      Infix _ => [extra_parse_rule]
0991c84e8dcf renamed InfixName to Infix etc.;
wenzelm
parents: 35129
diff changeset
    36
    | Infixl _ => [extra_parse_rule]
0991c84e8dcf renamed InfixName to Infix etc.;
wenzelm
parents: 35129
diff changeset
    37
    | Infixr _ => [extra_parse_rule]
33245
65232054ffd0 eliminated some old folds;
wenzelm
parents: 31023
diff changeset
    38
    | _ => [])
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    39
  end
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    40
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    41
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    42
(* transforming infix/mixfix declarations of constants with type ...->...
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    43
   a declaration of such a constant is transformed to a normal declaration with
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    44
   an internal name, the same type, and nofix. Additionally, a purely syntactic
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    45
   declaration with the original name, type ...=>..., and the original mixfix
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    46
   is generated and connected to the other declaration via some translation.
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    47
*)
35257
3e5980f612d9 adapted to authentic syntax;
wenzelm
parents: 35130
diff changeset
    48
fun transform thy (c, T, mx) =
3e5980f612d9 adapted to authentic syntax;
wenzelm
parents: 35130
diff changeset
    49
  let
42290
b1f544c84040 discontinued special treatment of structure Lexicon;
wenzelm
parents: 42287
diff changeset
    50
    fun syntax b = Lexicon.mark_const (Sign.full_bname thy b)
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    51
    val c1 = Binding.name_of c
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    52
    val c2 = c1 ^ "_cont_syntax"
42287
d98eb048a2e4 discontinued special treatment of structure Mixfix;
wenzelm
parents: 42224
diff changeset
    53
    val n = Mixfix.mixfix_args mx
35257
3e5980f612d9 adapted to authentic syntax;
wenzelm
parents: 35130
diff changeset
    54
  in
3e5980f612d9 adapted to authentic syntax;
wenzelm
parents: 35130
diff changeset
    55
    ((c, T, NoSyn),
3e5980f612d9 adapted to authentic syntax;
wenzelm
parents: 35130
diff changeset
    56
      (Binding.name c2, change_arrow n T, mx),
3e5980f612d9 adapted to authentic syntax;
wenzelm
parents: 35130
diff changeset
    57
      trans_rules (syntax c2) (syntax c1) n mx)
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    58
  end
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    59
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 56239
diff changeset
    60
fun cfun_arity (Type (n, [_, T])) = if n = \<^type_name>\<open>cfun\<close> then 1 + cfun_arity T else 0
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    61
  | cfun_arity _ = 0
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    62
35257
3e5980f612d9 adapted to authentic syntax;
wenzelm
parents: 35130
diff changeset
    63
fun is_contconst (_, _, NoSyn) = false
3e5980f612d9 adapted to authentic syntax;
wenzelm
parents: 35130
diff changeset
    64
  | is_contconst (_, _, Binder _) = false    (* FIXME ? *)
3e5980f612d9 adapted to authentic syntax;
wenzelm
parents: 35130
diff changeset
    65
  | is_contconst (c, T, mx) =
3e5980f612d9 adapted to authentic syntax;
wenzelm
parents: 35130
diff changeset
    66
      let
42287
d98eb048a2e4 discontinued special treatment of structure Mixfix;
wenzelm
parents: 42224
diff changeset
    67
        val n = Mixfix.mixfix_args mx handle ERROR msg =>
42381
309ec68442c6 added Binding.print convenience, which includes quote already;
wenzelm
parents: 42290
diff changeset
    68
          cat_error msg ("in mixfix annotation for " ^ Binding.print c)
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    69
      in cfun_arity T >= n end
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    70
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    71
35257
3e5980f612d9 adapted to authentic syntax;
wenzelm
parents: 35130
diff changeset
    72
(* add_consts *)
3e5980f612d9 adapted to authentic syntax;
wenzelm
parents: 35130
diff changeset
    73
3e5980f612d9 adapted to authentic syntax;
wenzelm
parents: 35130
diff changeset
    74
local
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    75
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    76
fun gen_add_consts prep_typ raw_decls thy =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    77
  let
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    78
    val decls = map (fn (c, T, mx) => (c, prep_typ thy T, mx)) raw_decls
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    79
    val (contconst_decls, normal_decls) = List.partition is_contconst decls
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    80
    val transformed_decls = map (transform thy) contconst_decls
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    81
  in
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    82
    thy
56239
17df7145a871 tuned signature;
wenzelm
parents: 43329
diff changeset
    83
    |> Sign.add_consts (normal_decls @ map #1 transformed_decls @ map #2 transformed_decls)
42204
b3277168c1e7 added Position.reports convenience;
wenzelm
parents: 42151
diff changeset
    84
    |> Sign.add_trrules (maps #3 transformed_decls)
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    85
  end
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    86
35257
3e5980f612d9 adapted to authentic syntax;
wenzelm
parents: 35130
diff changeset
    87
in
3e5980f612d9 adapted to authentic syntax;
wenzelm
parents: 35130
diff changeset
    88
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    89
val add_consts = gen_add_consts Sign.certify_typ
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    90
val add_consts_cmd = gen_add_consts Syntax.read_typ_global
35257
3e5980f612d9 adapted to authentic syntax;
wenzelm
parents: 35130
diff changeset
    91
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    92
end
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    93
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    94
end