src/Pure/Isar/constdefs.ML
author skalberg
Thu, 03 Mar 2005 12:43:01 +0100
changeset 15570 8d8c70b41bab
parent 15531 08c8dad8e399
child 15703 727ef1b8b3ee
permissions -rw-r--r--
Move towards standard functions.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
14657
c7cc01735801 'constdefs' with automatic type-inference and structure context;
wenzelm
parents:
diff changeset
     1
(*  Title:      Pure/Isar/constdefs.ML
c7cc01735801 'constdefs' with automatic type-inference and structure context;
wenzelm
parents:
diff changeset
     2
    ID:         $Id$
c7cc01735801 'constdefs' with automatic type-inference and structure context;
wenzelm
parents:
diff changeset
     3
    Author:     Makarius, Hagia Maria Sion Abbey (Jerusalem)
c7cc01735801 'constdefs' with automatic type-inference and structure context;
wenzelm
parents:
diff changeset
     4
c7cc01735801 'constdefs' with automatic type-inference and structure context;
wenzelm
parents:
diff changeset
     5
Standard constant definitions, with type-inference and optional
c7cc01735801 'constdefs' with automatic type-inference and structure context;
wenzelm
parents:
diff changeset
     6
structure context; specifications need to observe strictly sequential
c7cc01735801 'constdefs' with automatic type-inference and structure context;
wenzelm
parents:
diff changeset
     7
dependencies; no support for overloading.
c7cc01735801 'constdefs' with automatic type-inference and structure context;
wenzelm
parents:
diff changeset
     8
*)
c7cc01735801 'constdefs' with automatic type-inference and structure context;
wenzelm
parents:
diff changeset
     9
c7cc01735801 'constdefs' with automatic type-inference and structure context;
wenzelm
parents:
diff changeset
    10
signature CONSTDEFS =
c7cc01735801 'constdefs' with automatic type-inference and structure context;
wenzelm
parents:
diff changeset
    11
sig
c7cc01735801 'constdefs' with automatic type-inference and structure context;
wenzelm
parents:
diff changeset
    12
  val add_constdefs: (string list * string option) list *
c7cc01735801 'constdefs' with automatic type-inference and structure context;
wenzelm
parents:
diff changeset
    13
    ((bstring * string option * Syntax.mixfix) option *
c7cc01735801 'constdefs' with automatic type-inference and structure context;
wenzelm
parents:
diff changeset
    14
      ((bstring * Args.src list) * string)) list
c7cc01735801 'constdefs' with automatic type-inference and structure context;
wenzelm
parents:
diff changeset
    15
    -> theory -> theory
c7cc01735801 'constdefs' with automatic type-inference and structure context;
wenzelm
parents:
diff changeset
    16
  val add_constdefs_i: (string list * typ option) list *
c7cc01735801 'constdefs' with automatic type-inference and structure context;
wenzelm
parents:
diff changeset
    17
    ((bstring * typ option * Syntax.mixfix) option *
c7cc01735801 'constdefs' with automatic type-inference and structure context;
wenzelm
parents:
diff changeset
    18
      ((bstring * theory attribute list) * term)) list
c7cc01735801 'constdefs' with automatic type-inference and structure context;
wenzelm
parents:
diff changeset
    19
    -> theory -> theory
c7cc01735801 'constdefs' with automatic type-inference and structure context;
wenzelm
parents:
diff changeset
    20
end;
c7cc01735801 'constdefs' with automatic type-inference and structure context;
wenzelm
parents:
diff changeset
    21
c7cc01735801 'constdefs' with automatic type-inference and structure context;
wenzelm
parents:
diff changeset
    22
structure Constdefs: CONSTDEFS =
c7cc01735801 'constdefs' with automatic type-inference and structure context;
wenzelm
parents:
diff changeset
    23
struct
c7cc01735801 'constdefs' with automatic type-inference and structure context;
wenzelm
parents:
diff changeset
    24
c7cc01735801 'constdefs' with automatic type-inference and structure context;
wenzelm
parents:
diff changeset
    25
(** add_constdefs(_i) **)
c7cc01735801 'constdefs' with automatic type-inference and structure context;
wenzelm
parents:
diff changeset
    26
14664
148f6175fa78 improved messages;
wenzelm
parents: 14657
diff changeset
    27
fun pretty_const sg (c, T) =
148f6175fa78 improved messages;
wenzelm
parents: 14657
diff changeset
    28
  Pretty.block [Pretty.str c, Pretty.str " ::", Pretty.brk 1,
148f6175fa78 improved messages;
wenzelm
parents: 14657
diff changeset
    29
    Pretty.quote (Sign.pretty_typ sg T)];
148f6175fa78 improved messages;
wenzelm
parents: 14657
diff changeset
    30
148f6175fa78 improved messages;
wenzelm
parents: 14657
diff changeset
    31
fun pretty_constdefs sg decls =
148f6175fa78 improved messages;
wenzelm
parents: 14657
diff changeset
    32
  Pretty.big_list "constants" (map (pretty_const sg) decls);
148f6175fa78 improved messages;
wenzelm
parents: 14657
diff changeset
    33
14657
c7cc01735801 'constdefs' with automatic type-inference and structure context;
wenzelm
parents:
diff changeset
    34
fun gen_constdef prep_typ prep_term prep_att
c7cc01735801 'constdefs' with automatic type-inference and structure context;
wenzelm
parents:
diff changeset
    35
    structs (thy, (decl, ((raw_name, raw_atts), raw_prop))) =
c7cc01735801 'constdefs' with automatic type-inference and structure context;
wenzelm
parents:
diff changeset
    36
  let
14664
148f6175fa78 improved messages;
wenzelm
parents: 14657
diff changeset
    37
    val sign = Theory.sign_of thy;
148f6175fa78 improved messages;
wenzelm
parents: 14657
diff changeset
    38
    fun err msg ts =
148f6175fa78 improved messages;
wenzelm
parents: 14657
diff changeset
    39
      error (cat_lines (msg :: map (Sign.string_of_term sign) ts));
148f6175fa78 improved messages;
wenzelm
parents: 14657
diff changeset
    40
14657
c7cc01735801 'constdefs' with automatic type-inference and structure context;
wenzelm
parents:
diff changeset
    41
    val ctxt =
c7cc01735801 'constdefs' with automatic type-inference and structure context;
wenzelm
parents:
diff changeset
    42
      ProofContext.init thy |> ProofContext.add_fixes
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
    43
        (List.concat (map (fn (As, T) => map (fn A => (A, T, NONE)) As) structs));
14719
d1157ff6ffcb be liberal about constant names;
wenzelm
parents: 14664
diff changeset
    44
    val (ctxt', d, mx) =
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15010
diff changeset
    45
      (case decl of NONE => (ctxt, NONE, Syntax.NoSyn) | SOME (x, raw_T, mx) =>
14719
d1157ff6ffcb be liberal about constant names;
wenzelm
parents: 14664
diff changeset
    46
        let
d1157ff6ffcb be liberal about constant names;
wenzelm
parents: 14664
diff changeset
    47
          val x' = Syntax.const_name x mx and mx' = Syntax.fix_mixfix x mx;
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
    48
          val T = Option.map (prep_typ ctxt) raw_T;
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15010
diff changeset
    49
        in (ProofContext.add_fixes_liberal [(x', T, SOME mx')] ctxt, SOME x', mx') end);
14657
c7cc01735801 'constdefs' with automatic type-inference and structure context;
wenzelm
parents:
diff changeset
    50
c7cc01735801 'constdefs' with automatic type-inference and structure context;
wenzelm
parents:
diff changeset
    51
    val prop = prep_term ctxt' raw_prop;
14664
148f6175fa78 improved messages;
wenzelm
parents: 14657
diff changeset
    52
    val concl = Logic.strip_imp_concl prop;
148f6175fa78 improved messages;
wenzelm
parents: 14657
diff changeset
    53
    val (lhs, _) = Logic.dest_equals concl handle TERM _ =>
148f6175fa78 improved messages;
wenzelm
parents: 14657
diff changeset
    54
      err "Not a meta-equality (==):" [concl];
148f6175fa78 improved messages;
wenzelm
parents: 14657
diff changeset
    55
    val head = Term.head_of lhs;
148f6175fa78 improved messages;
wenzelm
parents: 14657
diff changeset
    56
    val (c, T) = Term.dest_Free head handle TERM _ =>
148f6175fa78 improved messages;
wenzelm
parents: 14657
diff changeset
    57
      err "Locally fixed variable required as head of definition:" [head];
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15010
diff changeset
    58
    val _ = (case d of NONE => () | SOME c' =>
14719
d1157ff6ffcb be liberal about constant names;
wenzelm
parents: 14664
diff changeset
    59
      if c <> c' then
d1157ff6ffcb be liberal about constant names;
wenzelm
parents: 14664
diff changeset
    60
        err ("Head of definition " ^ quote c ^ " differs from declaration " ^ quote c') []
14664
148f6175fa78 improved messages;
wenzelm
parents: 14657
diff changeset
    61
      else ());
14657
c7cc01735801 'constdefs' with automatic type-inference and structure context;
wenzelm
parents:
diff changeset
    62
14664
148f6175fa78 improved messages;
wenzelm
parents: 14657
diff changeset
    63
    val def = Term.subst_atomic [(Free (c, T), Const (Sign.full_name sign c, T))] prop;
14657
c7cc01735801 'constdefs' with automatic type-inference and structure context;
wenzelm
parents:
diff changeset
    64
    val name = if raw_name = "" then Thm.def_name c else raw_name;
c7cc01735801 'constdefs' with automatic type-inference and structure context;
wenzelm
parents:
diff changeset
    65
    val atts = map (prep_att thy) raw_atts;
14664
148f6175fa78 improved messages;
wenzelm
parents: 14657
diff changeset
    66
148f6175fa78 improved messages;
wenzelm
parents: 14657
diff changeset
    67
    val thy' =
148f6175fa78 improved messages;
wenzelm
parents: 14657
diff changeset
    68
      thy
148f6175fa78 improved messages;
wenzelm
parents: 14657
diff changeset
    69
      |> Theory.add_consts_i [(c, T, mx)]
148f6175fa78 improved messages;
wenzelm
parents: 14657
diff changeset
    70
      |> PureThy.add_defs_i false [((name, def), atts)] |> #1;
148f6175fa78 improved messages;
wenzelm
parents: 14657
diff changeset
    71
  in (thy', (c, T)) end;
14657
c7cc01735801 'constdefs' with automatic type-inference and structure context;
wenzelm
parents:
diff changeset
    72
c7cc01735801 'constdefs' with automatic type-inference and structure context;
wenzelm
parents:
diff changeset
    73
fun gen_constdefs prep_vars prep_typ prep_term prep_att (raw_structs, specs) thy =
14664
148f6175fa78 improved messages;
wenzelm
parents: 14657
diff changeset
    74
  let
148f6175fa78 improved messages;
wenzelm
parents: 14657
diff changeset
    75
    val structs = #2 (foldl_map prep_vars (ProofContext.init thy, raw_structs));
148f6175fa78 improved messages;
wenzelm
parents: 14657
diff changeset
    76
    val (thy', decls) = (thy, specs)
148f6175fa78 improved messages;
wenzelm
parents: 14657
diff changeset
    77
      |> foldl_map (gen_constdef prep_typ prep_term prep_att structs);
148f6175fa78 improved messages;
wenzelm
parents: 14657
diff changeset
    78
  in Pretty.writeln (pretty_constdefs (Theory.sign_of thy') decls); thy' end;
14657
c7cc01735801 'constdefs' with automatic type-inference and structure context;
wenzelm
parents:
diff changeset
    79
14719
d1157ff6ffcb be liberal about constant names;
wenzelm
parents: 14664
diff changeset
    80
val add_constdefs = gen_constdefs ProofContext.read_vars_liberal
d1157ff6ffcb be liberal about constant names;
wenzelm
parents: 14664
diff changeset
    81
  ProofContext.read_typ ProofContext.read_term_liberal Attrib.global_attribute;
d1157ff6ffcb be liberal about constant names;
wenzelm
parents: 14664
diff changeset
    82
val add_constdefs_i = gen_constdefs ProofContext.cert_vars_liberal
14657
c7cc01735801 'constdefs' with automatic type-inference and structure context;
wenzelm
parents:
diff changeset
    83
  ProofContext.cert_typ ProofContext.cert_term (K I);
c7cc01735801 'constdefs' with automatic type-inference and structure context;
wenzelm
parents:
diff changeset
    84
c7cc01735801 'constdefs' with automatic type-inference and structure context;
wenzelm
parents:
diff changeset
    85
end;