src/Pure/defs.ML
author wenzelm
Sat, 13 May 2006 02:51:37 +0200
changeset 19628 de019ddcd89e
parent 19624 3b6629701a79
child 19692 bad13b32c0f3
permissions -rw-r--r--
actually reject malformed defs; added unchecked flag; tuned;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
17707
bc0270e9d27f back to simple 'defs' (cf. revision 1.79 of theory.ML);
wenzelm
parents: 17670
diff changeset
     1
(*  Title:      Pure/defs.ML
16108
cf468b93a02e Implement cycle-free overloading, so that definitions cannot harm consistency any more (except of course via interaction with axioms).
obua
parents:
diff changeset
     2
    ID:         $Id$
17707
bc0270e9d27f back to simple 'defs' (cf. revision 1.79 of theory.ML);
wenzelm
parents: 17670
diff changeset
     3
    Author:     Makarius
16108
cf468b93a02e Implement cycle-free overloading, so that definitions cannot harm consistency any more (except of course via interaction with axioms).
obua
parents:
diff changeset
     4
19613
9bf274ec94cf allow dependencies of disjoint collections of instances;
wenzelm
parents: 19590
diff changeset
     5
Global well-formedness checks for constant definitions.  Covers
9bf274ec94cf allow dependencies of disjoint collections of instances;
wenzelm
parents: 19590
diff changeset
     6
dependencies of simple sub-structural overloading, where type
9bf274ec94cf allow dependencies of disjoint collections of instances;
wenzelm
parents: 19590
diff changeset
     7
arguments are approximated by the outermost type constructor.
16108
cf468b93a02e Implement cycle-free overloading, so that definitions cannot harm consistency any more (except of course via interaction with axioms).
obua
parents:
diff changeset
     8
*)
cf468b93a02e Implement cycle-free overloading, so that definitions cannot harm consistency any more (except of course via interaction with axioms).
obua
parents:
diff changeset
     9
16877
e92cba1d4842 tuned interfaces declare, define, finalize, merge:
wenzelm
parents: 16838
diff changeset
    10
signature DEFS =
e92cba1d4842 tuned interfaces declare, define, finalize, merge:
wenzelm
parents: 16838
diff changeset
    11
sig
17707
bc0270e9d27f back to simple 'defs' (cf. revision 1.79 of theory.ML);
wenzelm
parents: 17670
diff changeset
    12
  type T
19569
1c23356a8ea8 specifications_of: more detailed information;
wenzelm
parents: 19482
diff changeset
    13
  val specifications_of: T -> string ->
1c23356a8ea8 specifications_of: more detailed information;
wenzelm
parents: 19482
diff changeset
    14
   (serial * {is_def: bool, module: string, name: string, lhs: typ, rhs: (string * typ) list}) list
19590
12af4942923d simple substructure test for typargs (independent type constructors);
wenzelm
parents: 19569
diff changeset
    15
  val empty: T
19613
9bf274ec94cf allow dependencies of disjoint collections of instances;
wenzelm
parents: 19590
diff changeset
    16
  val merge: T * T -> T
9bf274ec94cf allow dependencies of disjoint collections of instances;
wenzelm
parents: 19590
diff changeset
    17
  val define: (string * typ -> typ list) ->
19628
de019ddcd89e actually reject malformed defs;
wenzelm
parents: 19624
diff changeset
    18
    bool -> bool -> string -> string -> string * typ -> (string * typ) list -> T -> T
16108
cf468b93a02e Implement cycle-free overloading, so that definitions cannot harm consistency any more (except of course via interaction with axioms).
obua
parents:
diff changeset
    19
end
cf468b93a02e Implement cycle-free overloading, so that definitions cannot harm consistency any more (except of course via interaction with axioms).
obua
parents:
diff changeset
    20
17711
c16cbe73798c activate signature constraints;
wenzelm
parents: 17707
diff changeset
    21
structure Defs: DEFS =
17707
bc0270e9d27f back to simple 'defs' (cf. revision 1.79 of theory.ML);
wenzelm
parents: 17670
diff changeset
    22
struct
16108
cf468b93a02e Implement cycle-free overloading, so that definitions cannot harm consistency any more (except of course via interaction with axioms).
obua
parents:
diff changeset
    23
19613
9bf274ec94cf allow dependencies of disjoint collections of instances;
wenzelm
parents: 19590
diff changeset
    24
(* dependency items *)
19569
1c23356a8ea8 specifications_of: more detailed information;
wenzelm
parents: 19482
diff changeset
    25
19613
9bf274ec94cf allow dependencies of disjoint collections of instances;
wenzelm
parents: 19590
diff changeset
    26
(*
9bf274ec94cf allow dependencies of disjoint collections of instances;
wenzelm
parents: 19590
diff changeset
    27
  Constant c covers all instances of c
9bf274ec94cf allow dependencies of disjoint collections of instances;
wenzelm
parents: 19590
diff changeset
    28
9bf274ec94cf allow dependencies of disjoint collections of instances;
wenzelm
parents: 19590
diff changeset
    29
  Instance (c, a) covers all instances of applications (c, [Type (a, _)])
9bf274ec94cf allow dependencies of disjoint collections of instances;
wenzelm
parents: 19590
diff changeset
    30
9bf274ec94cf allow dependencies of disjoint collections of instances;
wenzelm
parents: 19590
diff changeset
    31
  Different Constant/Constant or Instance/Instance items represent
9bf274ec94cf allow dependencies of disjoint collections of instances;
wenzelm
parents: 19590
diff changeset
    32
  disjoint sets of instances.  The set Constant c subsumes any
9bf274ec94cf allow dependencies of disjoint collections of instances;
wenzelm
parents: 19590
diff changeset
    33
  Instance (c, a) -- dependencies are propagated accordingly.
9bf274ec94cf allow dependencies of disjoint collections of instances;
wenzelm
parents: 19590
diff changeset
    34
*)
9bf274ec94cf allow dependencies of disjoint collections of instances;
wenzelm
parents: 19590
diff changeset
    35
9bf274ec94cf allow dependencies of disjoint collections of instances;
wenzelm
parents: 19590
diff changeset
    36
datatype item =
9bf274ec94cf allow dependencies of disjoint collections of instances;
wenzelm
parents: 19590
diff changeset
    37
  Constant of string |
9bf274ec94cf allow dependencies of disjoint collections of instances;
wenzelm
parents: 19590
diff changeset
    38
  Instance of string * string;
9bf274ec94cf allow dependencies of disjoint collections of instances;
wenzelm
parents: 19590
diff changeset
    39
9bf274ec94cf allow dependencies of disjoint collections of instances;
wenzelm
parents: 19590
diff changeset
    40
fun make_item (c, [Type (a, _)]) = Instance (c, a)
9bf274ec94cf allow dependencies of disjoint collections of instances;
wenzelm
parents: 19590
diff changeset
    41
  | make_item (c, _) = Constant c;
9bf274ec94cf allow dependencies of disjoint collections of instances;
wenzelm
parents: 19590
diff changeset
    42
9bf274ec94cf allow dependencies of disjoint collections of instances;
wenzelm
parents: 19590
diff changeset
    43
fun pretty_item (Constant c) = Pretty.str (quote c)
9bf274ec94cf allow dependencies of disjoint collections of instances;
wenzelm
parents: 19590
diff changeset
    44
  | pretty_item (Instance (c, a)) = Pretty.str (quote c ^ " (type " ^ quote a ^ ")");
9bf274ec94cf allow dependencies of disjoint collections of instances;
wenzelm
parents: 19590
diff changeset
    45
9bf274ec94cf allow dependencies of disjoint collections of instances;
wenzelm
parents: 19590
diff changeset
    46
fun item_ord (Constant c, Constant c') = fast_string_ord (c, c')
9bf274ec94cf allow dependencies of disjoint collections of instances;
wenzelm
parents: 19590
diff changeset
    47
  | item_ord (Instance ca, Instance ca') = prod_ord fast_string_ord fast_string_ord (ca, ca')
9bf274ec94cf allow dependencies of disjoint collections of instances;
wenzelm
parents: 19590
diff changeset
    48
  | item_ord (Constant _, Instance _) = LESS
9bf274ec94cf allow dependencies of disjoint collections of instances;
wenzelm
parents: 19590
diff changeset
    49
  | item_ord (Instance _, Constant _) = GREATER;
9bf274ec94cf allow dependencies of disjoint collections of instances;
wenzelm
parents: 19590
diff changeset
    50
9bf274ec94cf allow dependencies of disjoint collections of instances;
wenzelm
parents: 19590
diff changeset
    51
structure Items = GraphFun(type key = item val ord = item_ord);
9bf274ec94cf allow dependencies of disjoint collections of instances;
wenzelm
parents: 19590
diff changeset
    52
9bf274ec94cf allow dependencies of disjoint collections of instances;
wenzelm
parents: 19590
diff changeset
    53
fun propagate_deps insts deps =
9bf274ec94cf allow dependencies of disjoint collections of instances;
wenzelm
parents: 19590
diff changeset
    54
  let
19620
ccd6de95f4a6 improved propagate_deps;
wenzelm
parents: 19613
diff changeset
    55
    fun inst_item (Constant c) = Symtab.lookup_list insts c
ccd6de95f4a6 improved propagate_deps;
wenzelm
parents: 19613
diff changeset
    56
      | inst_item (Instance _) = [];
ccd6de95f4a6 improved propagate_deps;
wenzelm
parents: 19613
diff changeset
    57
    fun inst_edge i j =
19628
de019ddcd89e actually reject malformed defs;
wenzelm
parents: 19624
diff changeset
    58
      fold Items.add_edge_acyclic (tl (product (i :: inst_item i) (j :: inst_item j)));
19613
9bf274ec94cf allow dependencies of disjoint collections of instances;
wenzelm
parents: 19590
diff changeset
    59
  in Items.fold (fn (i, (_, (_, js))) => fold (inst_edge i) js) deps deps end;
9bf274ec94cf allow dependencies of disjoint collections of instances;
wenzelm
parents: 19590
diff changeset
    60
9bf274ec94cf allow dependencies of disjoint collections of instances;
wenzelm
parents: 19590
diff changeset
    61
9bf274ec94cf allow dependencies of disjoint collections of instances;
wenzelm
parents: 19590
diff changeset
    62
(* specifications *)
16877
e92cba1d4842 tuned interfaces declare, define, finalize, merge:
wenzelm
parents: 16838
diff changeset
    63
19569
1c23356a8ea8 specifications_of: more detailed information;
wenzelm
parents: 19482
diff changeset
    64
type spec = {is_def: bool, module: string, name: string, lhs: typ, rhs: (string * typ) list};
1c23356a8ea8 specifications_of: more detailed information;
wenzelm
parents: 19482
diff changeset
    65
17707
bc0270e9d27f back to simple 'defs' (cf. revision 1.79 of theory.ML);
wenzelm
parents: 17670
diff changeset
    66
datatype T = Defs of
19613
9bf274ec94cf allow dependencies of disjoint collections of instances;
wenzelm
parents: 19590
diff changeset
    67
 {specs: (bool * spec Inttab.table) Symtab.table,
19620
ccd6de95f4a6 improved propagate_deps;
wenzelm
parents: 19613
diff changeset
    68
  insts: item list Symtab.table,
19613
9bf274ec94cf allow dependencies of disjoint collections of instances;
wenzelm
parents: 19590
diff changeset
    69
  deps: unit Items.T};
9bf274ec94cf allow dependencies of disjoint collections of instances;
wenzelm
parents: 19590
diff changeset
    70
9bf274ec94cf allow dependencies of disjoint collections of instances;
wenzelm
parents: 19590
diff changeset
    71
fun no_overloading_of (Defs {specs, ...}) c =
9bf274ec94cf allow dependencies of disjoint collections of instances;
wenzelm
parents: 19590
diff changeset
    72
  (case Symtab.lookup specs c of
9bf274ec94cf allow dependencies of disjoint collections of instances;
wenzelm
parents: 19590
diff changeset
    73
    SOME (b, _) => b
9bf274ec94cf allow dependencies of disjoint collections of instances;
wenzelm
parents: 19590
diff changeset
    74
  | NONE => false);
16877
e92cba1d4842 tuned interfaces declare, define, finalize, merge:
wenzelm
parents: 16838
diff changeset
    75
19613
9bf274ec94cf allow dependencies of disjoint collections of instances;
wenzelm
parents: 19590
diff changeset
    76
fun specifications_of (Defs {specs, ...}) c =
9bf274ec94cf allow dependencies of disjoint collections of instances;
wenzelm
parents: 19590
diff changeset
    77
  (case Symtab.lookup specs c of
9bf274ec94cf allow dependencies of disjoint collections of instances;
wenzelm
parents: 19590
diff changeset
    78
    SOME (_, sps) => Inttab.dest sps
9bf274ec94cf allow dependencies of disjoint collections of instances;
wenzelm
parents: 19590
diff changeset
    79
  | NONE => []);
16877
e92cba1d4842 tuned interfaces declare, define, finalize, merge:
wenzelm
parents: 16838
diff changeset
    80
19613
9bf274ec94cf allow dependencies of disjoint collections of instances;
wenzelm
parents: 19590
diff changeset
    81
fun make_defs (specs, insts, deps) = Defs {specs = specs, insts = insts, deps = deps};
9bf274ec94cf allow dependencies of disjoint collections of instances;
wenzelm
parents: 19590
diff changeset
    82
fun map_defs f (Defs {specs, insts, deps}) = make_defs (f (specs, insts, deps));
9bf274ec94cf allow dependencies of disjoint collections of instances;
wenzelm
parents: 19590
diff changeset
    83
9bf274ec94cf allow dependencies of disjoint collections of instances;
wenzelm
parents: 19590
diff changeset
    84
val empty = make_defs (Symtab.empty, Symtab.empty, Items.empty);
19590
12af4942923d simple substructure test for typargs (independent type constructors);
wenzelm
parents: 19569
diff changeset
    85
16877
e92cba1d4842 tuned interfaces declare, define, finalize, merge:
wenzelm
parents: 16838
diff changeset
    86
19624
wenzelm
parents: 19620
diff changeset
    87
(* disjoint specs *)
16877
e92cba1d4842 tuned interfaces declare, define, finalize, merge:
wenzelm
parents: 16838
diff changeset
    88
17707
bc0270e9d27f back to simple 'defs' (cf. revision 1.79 of theory.ML);
wenzelm
parents: 17670
diff changeset
    89
fun disjoint_types T U =
bc0270e9d27f back to simple 'defs' (cf. revision 1.79 of theory.ML);
wenzelm
parents: 17670
diff changeset
    90
  (Type.raw_unify (T, Logic.incr_tvar (maxidx_of_typ T + 1) U) Vartab.empty; false)
bc0270e9d27f back to simple 'defs' (cf. revision 1.79 of theory.ML);
wenzelm
parents: 17670
diff changeset
    91
    handle Type.TUNIFY => true;
16308
636a1a84977a 1) Fixed bug in Defs.merge_edges_1.
obua
parents: 16198
diff changeset
    92
19613
9bf274ec94cf allow dependencies of disjoint collections of instances;
wenzelm
parents: 19590
diff changeset
    93
fun disjoint_specs c (i, {lhs = T, name = a, ...}: spec) =
19569
1c23356a8ea8 specifications_of: more detailed information;
wenzelm
parents: 19482
diff changeset
    94
  Inttab.forall (fn (j, {lhs = U, name = b, ...}: spec) =>
1c23356a8ea8 specifications_of: more detailed information;
wenzelm
parents: 19482
diff changeset
    95
    i = j orelse not (Type.could_unify (T, U)) orelse disjoint_types T U orelse
1c23356a8ea8 specifications_of: more detailed information;
wenzelm
parents: 19482
diff changeset
    96
      error ("Type clash in specifications " ^ quote a ^ " and " ^ quote b ^
1c23356a8ea8 specifications_of: more detailed information;
wenzelm
parents: 19482
diff changeset
    97
        " for constant " ^ quote c));
16877
e92cba1d4842 tuned interfaces declare, define, finalize, merge:
wenzelm
parents: 16838
diff changeset
    98
19624
wenzelm
parents: 19620
diff changeset
    99
wenzelm
parents: 19620
diff changeset
   100
(* merge *)
wenzelm
parents: 19620
diff changeset
   101
19613
9bf274ec94cf allow dependencies of disjoint collections of instances;
wenzelm
parents: 19590
diff changeset
   102
fun cycle_msg css =
9bf274ec94cf allow dependencies of disjoint collections of instances;
wenzelm
parents: 19590
diff changeset
   103
  let
9bf274ec94cf allow dependencies of disjoint collections of instances;
wenzelm
parents: 19590
diff changeset
   104
    fun prt_cycle items = Pretty.block (flat
19620
ccd6de95f4a6 improved propagate_deps;
wenzelm
parents: 19613
diff changeset
   105
      (separate [Pretty.str " ->", Pretty.brk 1] (map (single o pretty_item) items)));
19613
9bf274ec94cf allow dependencies of disjoint collections of instances;
wenzelm
parents: 19590
diff changeset
   106
  in Pretty.string_of (Pretty.big_list "Cyclic dependency of constants:" (map prt_cycle css)) end;
16982
4600e74aeb0d chain_history: turned into runtime flag;
wenzelm
parents: 16936
diff changeset
   107
19613
9bf274ec94cf allow dependencies of disjoint collections of instances;
wenzelm
parents: 19590
diff changeset
   108
fun merge
9bf274ec94cf allow dependencies of disjoint collections of instances;
wenzelm
parents: 19590
diff changeset
   109
   (Defs {specs = specs1, insts = insts1, deps = deps1},
9bf274ec94cf allow dependencies of disjoint collections of instances;
wenzelm
parents: 19590
diff changeset
   110
    Defs {specs = specs2, insts = insts2, deps = deps2}) =
9bf274ec94cf allow dependencies of disjoint collections of instances;
wenzelm
parents: 19590
diff changeset
   111
  let
9bf274ec94cf allow dependencies of disjoint collections of instances;
wenzelm
parents: 19590
diff changeset
   112
    val specs' = (specs1, specs2) |> Symtab.join (fn c => fn ((b, sps1), (_, sps2)) =>
9bf274ec94cf allow dependencies of disjoint collections of instances;
wenzelm
parents: 19590
diff changeset
   113
      (b, Inttab.fold (fn sp2 => (disjoint_specs c sp2 sps1; Inttab.update sp2)) sps2 sps1));
9bf274ec94cf allow dependencies of disjoint collections of instances;
wenzelm
parents: 19590
diff changeset
   114
    val insts' = Symtab.merge_list (op =) (insts1, insts2);
9bf274ec94cf allow dependencies of disjoint collections of instances;
wenzelm
parents: 19590
diff changeset
   115
    val items' = propagate_deps insts' (Items.merge_acyclic (K true) (deps1, deps2))
9bf274ec94cf allow dependencies of disjoint collections of instances;
wenzelm
parents: 19590
diff changeset
   116
      handle Items.CYCLES cycles => error (cycle_msg cycles);
9bf274ec94cf allow dependencies of disjoint collections of instances;
wenzelm
parents: 19590
diff changeset
   117
  in make_defs (specs', insts', items') end;
9bf274ec94cf allow dependencies of disjoint collections of instances;
wenzelm
parents: 19590
diff changeset
   118
9bf274ec94cf allow dependencies of disjoint collections of instances;
wenzelm
parents: 19590
diff changeset
   119
9bf274ec94cf allow dependencies of disjoint collections of instances;
wenzelm
parents: 19590
diff changeset
   120
(* define *)
19590
12af4942923d simple substructure test for typargs (independent type constructors);
wenzelm
parents: 19569
diff changeset
   121
19624
wenzelm
parents: 19620
diff changeset
   122
fun pure_args args = forall Term.is_TVar args andalso not (has_duplicates (op =) args);
wenzelm
parents: 19620
diff changeset
   123
19628
de019ddcd89e actually reject malformed defs;
wenzelm
parents: 19624
diff changeset
   124
fun define const_typargs unchecked is_def module name lhs rhs defs = defs
19613
9bf274ec94cf allow dependencies of disjoint collections of instances;
wenzelm
parents: 19590
diff changeset
   125
    |> map_defs (fn (specs, insts, deps) =>
17707
bc0270e9d27f back to simple 'defs' (cf. revision 1.79 of theory.ML);
wenzelm
parents: 17670
diff changeset
   126
  let
19590
12af4942923d simple substructure test for typargs (independent type constructors);
wenzelm
parents: 19569
diff changeset
   127
    val (c, T) = lhs;
12af4942923d simple substructure test for typargs (independent type constructors);
wenzelm
parents: 19569
diff changeset
   128
    val args = const_typargs lhs;
19624
wenzelm
parents: 19620
diff changeset
   129
    val no_overloading = pure_args args;
wenzelm
parents: 19620
diff changeset
   130
    val rec_args = (case args of [Type (_, Ts)] => if pure_args Ts then Ts else [] | _ => []);
17707
bc0270e9d27f back to simple 'defs' (cf. revision 1.79 of theory.ML);
wenzelm
parents: 17670
diff changeset
   131
19628
de019ddcd89e actually reject malformed defs;
wenzelm
parents: 19624
diff changeset
   132
    val lhs' = make_item (c, args);
de019ddcd89e actually reject malformed defs;
wenzelm
parents: 19624
diff changeset
   133
    val rhs' =
de019ddcd89e actually reject malformed defs;
wenzelm
parents: 19624
diff changeset
   134
      if unchecked then []
de019ddcd89e actually reject malformed defs;
wenzelm
parents: 19624
diff changeset
   135
      else rhs |> map_filter (fn (c', T') =>
de019ddcd89e actually reject malformed defs;
wenzelm
parents: 19624
diff changeset
   136
        let val args' = const_typargs (c', T') in
de019ddcd89e actually reject malformed defs;
wenzelm
parents: 19624
diff changeset
   137
          if gen_subset (op =) (args', rec_args) then NONE
de019ddcd89e actually reject malformed defs;
wenzelm
parents: 19624
diff changeset
   138
          else SOME (make_item (c', if no_overloading_of defs c' then [] else args'))
de019ddcd89e actually reject malformed defs;
wenzelm
parents: 19624
diff changeset
   139
        end);
de019ddcd89e actually reject malformed defs;
wenzelm
parents: 19624
diff changeset
   140
19569
1c23356a8ea8 specifications_of: more detailed information;
wenzelm
parents: 19482
diff changeset
   141
    val spec = (serial (), {is_def = is_def, module = module, name = name, lhs = T, rhs = rhs});
19613
9bf274ec94cf allow dependencies of disjoint collections of instances;
wenzelm
parents: 19590
diff changeset
   142
    val specs' = specs
9bf274ec94cf allow dependencies of disjoint collections of instances;
wenzelm
parents: 19590
diff changeset
   143
      |> Symtab.default (c, (false, Inttab.empty))
9bf274ec94cf allow dependencies of disjoint collections of instances;
wenzelm
parents: 19590
diff changeset
   144
      |> Symtab.map_entry c (fn (_, sps) =>
9bf274ec94cf allow dependencies of disjoint collections of instances;
wenzelm
parents: 19590
diff changeset
   145
          (disjoint_specs c spec sps; (no_overloading, Inttab.update spec sps)));
19620
ccd6de95f4a6 improved propagate_deps;
wenzelm
parents: 19613
diff changeset
   146
    val insts' = insts |> fold (fn i as Instance (c, _) =>
ccd6de95f4a6 improved propagate_deps;
wenzelm
parents: 19613
diff changeset
   147
        Symtab.insert_list (op =) (c, i) | _ => I) (lhs' :: rhs');
19613
9bf274ec94cf allow dependencies of disjoint collections of instances;
wenzelm
parents: 19590
diff changeset
   148
    val deps' = deps
19628
de019ddcd89e actually reject malformed defs;
wenzelm
parents: 19624
diff changeset
   149
      |> fold (Items.default_node o rpair ()) (lhs' :: rhs')
de019ddcd89e actually reject malformed defs;
wenzelm
parents: 19624
diff changeset
   150
      |> Items.add_deps_acyclic (lhs', rhs')
19613
9bf274ec94cf allow dependencies of disjoint collections of instances;
wenzelm
parents: 19590
diff changeset
   151
      |> propagate_deps insts'
19628
de019ddcd89e actually reject malformed defs;
wenzelm
parents: 19624
diff changeset
   152
      handle Items.CYCLES cycles => error (cycle_msg cycles);
19613
9bf274ec94cf allow dependencies of disjoint collections of instances;
wenzelm
parents: 19590
diff changeset
   153
9bf274ec94cf allow dependencies of disjoint collections of instances;
wenzelm
parents: 19590
diff changeset
   154
  in (specs', insts', deps') end);
16877
e92cba1d4842 tuned interfaces declare, define, finalize, merge:
wenzelm
parents: 16838
diff changeset
   155
16108
cf468b93a02e Implement cycle-free overloading, so that definitions cannot harm consistency any more (except of course via interaction with axioms).
obua
parents:
diff changeset
   156
end;