src/HOL/Tools/BNF/bnf_def.ML
author kuncar
Thu, 10 Apr 2014 17:48:16 +0200
changeset 57864 f54003750e7d
parent 57718 5a93b8f928a2
child 57865 2ae16e3d8b6d
permissions -rw-r--r--
don't forget to init Interpretation and transfer theorems in the interpretation hook
blanchet@56403
     1
(*  Title:      HOL/Tools/BNF/bnf_def.ML
blanchet@49990
     2
    Author:     Dmitriy Traytel, TU Muenchen
blanchet@49990
     3
    Author:     Jasmin Blanchette, TU Muenchen
blanchet@49990
     4
    Copyright   2012
blanchet@49990
     5
blanchet@49990
     6
Definition of bounded natural functors.
blanchet@49990
     7
*)
blanchet@49990
     8
blanchet@49990
     9
signature BNF_DEF =
blanchet@49990
    10
sig
blanchet@52974
    11
  type bnf
blanchet@49990
    12
  type nonemptiness_witness = {I: int list, wit: term, prop: thm list}
blanchet@49990
    13
blanchet@52974
    14
  val morph_bnf: morphism -> bnf -> bnf
traytel@57358
    15
  val morph_bnf_defs: morphism -> bnf -> bnf
blanchet@52974
    16
  val bnf_of: Proof.context -> string -> bnf option
blanchet@57688
    17
  val bnf_interpretation: (bnf -> theory -> theory) -> theory -> theory
blanchet@57688
    18
  val register_bnf: string -> bnf -> local_theory -> local_theory
traytel@50449
    19
blanchet@52974
    20
  val name_of_bnf: bnf -> binding
blanchet@52974
    21
  val T_of_bnf: bnf -> typ
blanchet@52974
    22
  val live_of_bnf: bnf -> int
blanchet@52974
    23
  val lives_of_bnf: bnf -> typ list
blanchet@52974
    24
  val dead_of_bnf: bnf -> int
blanchet@52974
    25
  val deads_of_bnf: bnf -> typ list
blanchet@57688
    26
  val bd_of_bnf: bnf -> term
blanchet@52974
    27
  val nwits_of_bnf: bnf -> int
blanchet@49990
    28
blanchet@49990
    29
  val mapN: string
blanchet@50522
    30
  val relN: string
blanchet@49990
    31
  val setN: string
blanchet@49990
    32
  val mk_setN: int -> string
traytel@56367
    33
  val mk_witN: int -> string
blanchet@49990
    34
blanchet@52974
    35
  val map_of_bnf: bnf -> term
blanchet@52974
    36
  val sets_of_bnf: bnf -> term list
blanchet@52974
    37
  val rel_of_bnf: bnf -> term
blanchet@50229
    38
blanchet@52974
    39
  val mk_T_of_bnf: typ list -> typ list -> bnf -> typ
blanchet@52974
    40
  val mk_bd_of_bnf: typ list -> typ list -> bnf -> term
blanchet@52974
    41
  val mk_map_of_bnf: typ list -> typ list -> typ list -> bnf -> term
blanchet@52974
    42
  val mk_rel_of_bnf: typ list -> typ list -> typ list -> bnf -> term
blanchet@52974
    43
  val mk_sets_of_bnf: typ list list -> typ list list -> bnf -> term list
blanchet@52974
    44
  val mk_wits_of_bnf: typ list list -> typ list list -> bnf -> (int list * term) list
blanchet@49990
    45
blanchet@52974
    46
  val bd_Card_order_of_bnf: bnf -> thm
blanchet@52974
    47
  val bd_Cinfinite_of_bnf: bnf -> thm
blanchet@52974
    48
  val bd_Cnotzero_of_bnf: bnf -> thm
blanchet@52974
    49
  val bd_card_order_of_bnf: bnf -> thm
blanchet@52974
    50
  val bd_cinfinite_of_bnf: bnf -> thm
blanchet@52974
    51
  val collect_set_map_of_bnf: bnf -> thm
blanchet@52974
    52
  val in_bd_of_bnf: bnf -> thm
blanchet@52974
    53
  val in_cong_of_bnf: bnf -> thm
blanchet@52974
    54
  val in_mono_of_bnf: bnf -> thm
traytel@53030
    55
  val in_rel_of_bnf: bnf -> thm
blanchet@54424
    56
  val map_comp0_of_bnf: bnf -> thm
blanchet@54425
    57
  val map_comp_of_bnf: bnf -> thm
blanchet@52974
    58
  val map_cong0_of_bnf: bnf -> thm
blanchet@52974
    59
  val map_cong_of_bnf: bnf -> thm
blanchet@52974
    60
  val map_def_of_bnf: bnf -> thm
blanchet@54407
    61
  val map_id0_of_bnf: bnf -> thm
blanchet@54422
    62
  val map_id_of_bnf: bnf -> thm
traytel@53856
    63
  val map_transfer_of_bnf: bnf -> thm
traytel@56183
    64
  val le_rel_OO_of_bnf: bnf -> thm
blanchet@52974
    65
  val rel_def_of_bnf: bnf -> thm
traytel@53030
    66
  val rel_Grp_of_bnf: bnf -> thm
traytel@53030
    67
  val rel_OO_of_bnf: bnf -> thm
traytel@53030
    68
  val rel_OO_Grp_of_bnf: bnf -> thm
traytel@53030
    69
  val rel_cong_of_bnf: bnf -> thm
traytel@53030
    70
  val rel_conversep_of_bnf: bnf -> thm
traytel@53030
    71
  val rel_mono_of_bnf: bnf -> thm
traytel@53053
    72
  val rel_mono_strong_of_bnf: bnf -> thm
blanchet@52974
    73
  val rel_eq_of_bnf: bnf -> thm
blanchet@52974
    74
  val rel_flip_of_bnf: bnf -> thm
blanchet@52974
    75
  val set_bd_of_bnf: bnf -> thm list
blanchet@52974
    76
  val set_defs_of_bnf: bnf -> thm list
blanchet@54426
    77
  val set_map0_of_bnf: bnf -> thm list
blanchet@54427
    78
  val set_map_of_bnf: bnf -> thm list
blanchet@52974
    79
  val wit_thms_of_bnf: bnf -> thm list
blanchet@52974
    80
  val wit_thmss_of_bnf: bnf -> thm list list
blanchet@49990
    81
blanchet@55688
    82
  val mk_map: int -> typ list -> typ list -> term -> term
blanchet@55688
    83
  val mk_rel: int -> typ list -> typ list -> term -> term
blanchet@55689
    84
  val build_map: Proof.context -> (typ * typ -> term) -> typ * typ -> term
blanchet@55689
    85
  val build_rel: Proof.context -> (typ * typ -> term) -> typ * typ -> term
blanchet@55698
    86
  val flatten_type_args_of_bnf: bnf -> 'a -> 'a list -> 'a list
blanchet@55698
    87
  val map_flattened_map_args: Proof.context -> string -> (term list -> 'a list) -> term list ->
blanchet@55698
    88
    'a list
blanchet@55688
    89
blanchet@49990
    90
  val mk_witness: int list * term -> thm list -> nonemptiness_witness
traytel@50118
    91
  val minimize_wits: (''a list * 'b) list -> (''a list * 'b) list
blanchet@52974
    92
  val wits_of_bnf: bnf -> nonemptiness_witness list
blanchet@49990
    93
traytel@53772
    94
  val zip_axioms: 'a -> 'a -> 'a -> 'a list -> 'a -> 'a -> 'a list -> 'a -> 'a -> 'a list
blanchet@50471
    95
blanchet@57196
    96
  datatype inline_policy = Dont_Inline | Hardly_Inline | Smart_Inline | Do_Inline
blanchet@50553
    97
  datatype fact_policy = Dont_Note | Note_Some | Note_All
blanchet@50553
    98
blanchet@49990
    99
  val bnf_note_all: bool Config.T
traytel@54280
   100
  val bnf_timing: bool Config.T
traytel@50450
   101
  val user_policy: fact_policy -> Proof.context -> fact_policy
traytel@53857
   102
  val note_bnf_thms: fact_policy -> (binding -> binding) -> binding -> bnf -> Proof.context ->
traytel@53857
   103
    Proof.context
blanchet@49990
   104
blanchet@49990
   105
  val print_bnfs: Proof.context -> unit
traytel@57358
   106
  val prepare_def: inline_policy -> (Proof.context -> fact_policy) -> bool ->
traytel@57358
   107
    (binding -> binding) -> (Proof.context -> 'a -> typ) -> (Proof.context -> 'b -> term) ->
traytel@57358
   108
    typ list option -> binding -> binding -> binding list ->
traytel@55974
   109
    (((((binding * 'a) * 'b) * 'b list) * 'b) * 'b list) * 'b option -> Proof.context ->
traytel@55974
   110
    string * term list *
traytel@56539
   111
    ((Proof.context -> thm list -> tactic) option * term list list) *
traytel@55974
   112
    ((thm list -> thm list list) -> thm list list -> Proof.context -> bnf * local_theory) *
traytel@55974
   113
    local_theory * thm list
traytel@55974
   114
traytel@57358
   115
  val define_bnf_consts: inline_policy -> fact_policy -> bool -> typ list option ->
traytel@56183
   116
    binding -> binding -> binding list ->
traytel@56183
   117
    (((((binding * typ) * term) * term list) * term) * term list) * term option -> local_theory ->
traytel@56183
   118
      ((typ list * typ list * typ list * typ) *
traytel@56183
   119
       (term * term list * term * (int list * term) list * term) *
traytel@56183
   120
       (thm * thm list * thm * thm list * thm) *
traytel@56183
   121
       ((typ list -> typ list -> typ list -> term) *
traytel@56183
   122
        (typ list -> typ list -> term -> term) *
traytel@56183
   123
        (typ list -> typ list -> typ -> typ) *
traytel@56183
   124
        (typ list -> typ list -> typ list -> term) *
traytel@56183
   125
        (typ list -> typ list -> typ list -> term))) * local_theory
traytel@56183
   126
traytel@57358
   127
  val bnf_def: inline_policy -> (Proof.context -> fact_policy) -> bool -> (binding -> binding) ->
traytel@56539
   128
    (Proof.context -> tactic) list ->
traytel@56539
   129
    (Proof.context -> tactic) -> typ list option -> binding ->
blanchet@52904
   130
    binding -> binding list ->
traytel@55794
   131
    (((((binding * typ) * term) * term list) * term) * term list) * term option ->
blanchet@52974
   132
    local_theory -> bnf * local_theory
blanchet@49990
   133
end;
blanchet@49990
   134
blanchet@49990
   135
structure BNF_Def : BNF_DEF =
blanchet@49990
   136
struct
blanchet@49990
   137
blanchet@49990
   138
open BNF_Util
blanchet@50478
   139
open BNF_Tactics
blanchet@50299
   140
open BNF_Def_Tactics
blanchet@49990
   141
blanchet@55997
   142
val fundefcong_attrs = @{attributes [fundef_cong]};
blanchet@52902
   143
blanchet@49990
   144
type axioms = {
blanchet@54407
   145
  map_id0: thm,
blanchet@54424
   146
  map_comp0: thm,
blanchet@52898
   147
  map_cong0: thm,
blanchet@54426
   148
  set_map0: thm list,
blanchet@49990
   149
  bd_card_order: thm,
blanchet@49990
   150
  bd_cinfinite: thm,
blanchet@49990
   151
  set_bd: thm list,
traytel@56183
   152
  le_rel_OO: thm,
traytel@53030
   153
  rel_OO_Grp: thm
blanchet@49990
   154
};
blanchet@49990
   155
traytel@56183
   156
fun mk_axioms' ((((((((id, comp), cong), map), c_o), cinf), set_bd), le_rel_OO), rel) =
blanchet@54426
   157
  {map_id0 = id, map_comp0 = comp, map_cong0 = cong, set_map0 = map, bd_card_order = c_o,
traytel@56183
   158
   bd_cinfinite = cinf, set_bd = set_bd, le_rel_OO = le_rel_OO, rel_OO_Grp = rel};
blanchet@49990
   159
wenzelm@53067
   160
fun dest_cons [] = raise List.Empty
blanchet@49990
   161
  | dest_cons (x :: xs) = (x, xs);
blanchet@49990
   162
blanchet@49990
   163
fun mk_axioms n thms = thms
blanchet@49990
   164
  |> map the_single
blanchet@49990
   165
  |> dest_cons
blanchet@49990
   166
  ||>> dest_cons
blanchet@49990
   167
  ||>> dest_cons
blanchet@49990
   168
  ||>> chop n
blanchet@49990
   169
  ||>> dest_cons
blanchet@49990
   170
  ||>> dest_cons
blanchet@49990
   171
  ||>> chop n
blanchet@49990
   172
  ||>> dest_cons
blanchet@49990
   173
  ||> the_single
blanchet@49990
   174
  |> mk_axioms';
blanchet@49990
   175
traytel@56183
   176
fun zip_axioms mid mcomp mcong smap bdco bdinf sbd le_rel_OO rel =
traytel@56183
   177
  [mid, mcomp, mcong] @ smap @ [bdco, bdinf] @ sbd @ [le_rel_OO, rel];
blanchet@50475
   178
blanchet@54426
   179
fun dest_axioms {map_id0, map_comp0, map_cong0, set_map0, bd_card_order, bd_cinfinite, set_bd,
traytel@56183
   180
  le_rel_OO, rel_OO_Grp} =
traytel@56183
   181
  zip_axioms map_id0 map_comp0 map_cong0 set_map0 bd_card_order bd_cinfinite set_bd le_rel_OO
traytel@53030
   182
    rel_OO_Grp;
blanchet@49990
   183
blanchet@54426
   184
fun map_axioms f {map_id0, map_comp0, map_cong0, set_map0, bd_card_order, bd_cinfinite, set_bd,
traytel@56183
   185
  le_rel_OO, rel_OO_Grp} =
blanchet@54407
   186
  {map_id0 = f map_id0,
blanchet@54424
   187
    map_comp0 = f map_comp0,
blanchet@52898
   188
    map_cong0 = f map_cong0,
blanchet@54426
   189
    set_map0 = map f set_map0,
blanchet@50478
   190
    bd_card_order = f bd_card_order,
blanchet@50478
   191
    bd_cinfinite = f bd_cinfinite,
blanchet@50478
   192
    set_bd = map f set_bd,
traytel@56183
   193
    le_rel_OO = f le_rel_OO,
traytel@53030
   194
    rel_OO_Grp = f rel_OO_Grp};
blanchet@49990
   195
blanchet@49990
   196
val morph_axioms = map_axioms o Morphism.thm;
blanchet@49990
   197
blanchet@49990
   198
type defs = {
blanchet@49990
   199
  map_def: thm,
blanchet@49990
   200
  set_defs: thm list,
traytel@53030
   201
  rel_def: thm
blanchet@49990
   202
}
blanchet@49990
   203
traytel@53030
   204
fun mk_defs map sets rel = {map_def = map, set_defs = sets, rel_def = rel};
blanchet@49990
   205
traytel@53030
   206
fun map_defs f {map_def, set_defs, rel_def} =
traytel@53030
   207
  {map_def = f map_def, set_defs = map f set_defs, rel_def = f rel_def};
blanchet@49990
   208
blanchet@49990
   209
val morph_defs = map_defs o Morphism.thm;
blanchet@49990
   210
blanchet@49990
   211
type facts = {
blanchet@49990
   212
  bd_Card_order: thm,
blanchet@49990
   213
  bd_Cinfinite: thm,
blanchet@49990
   214
  bd_Cnotzero: thm,
blanchet@52903
   215
  collect_set_map: thm lazy,
traytel@53772
   216
  in_bd: thm lazy,
blanchet@49990
   217
  in_cong: thm lazy,
blanchet@49990
   218
  in_mono: thm lazy,
traytel@53030
   219
  in_rel: thm lazy,
blanchet@54425
   220
  map_comp: thm lazy,
blanchet@52899
   221
  map_cong: thm lazy,
blanchet@54422
   222
  map_id: thm lazy,
traytel@53856
   223
  map_transfer: thm lazy,
blanchet@50606
   224
  rel_eq: thm lazy,
blanchet@50552
   225
  rel_flip: thm lazy,
blanchet@54427
   226
  set_map: thm lazy list,
traytel@53030
   227
  rel_cong: thm lazy,
traytel@53030
   228
  rel_mono: thm lazy,
traytel@53053
   229
  rel_mono_strong: thm lazy,
traytel@53030
   230
  rel_Grp: thm lazy,
traytel@53030
   231
  rel_conversep: thm lazy,
traytel@53030
   232
  rel_OO: thm lazy
blanchet@49990
   233
};
blanchet@49990
   234
traytel@53772
   235
fun mk_facts bd_Card_order bd_Cinfinite bd_Cnotzero collect_set_map in_bd in_cong in_mono in_rel
traytel@56183
   236
    map_comp map_cong map_id map_transfer rel_eq rel_flip set_map rel_cong rel_mono
traytel@53856
   237
    rel_mono_strong rel_Grp rel_conversep rel_OO = {
blanchet@49990
   238
  bd_Card_order = bd_Card_order,
blanchet@49990
   239
  bd_Cinfinite = bd_Cinfinite,
blanchet@49990
   240
  bd_Cnotzero = bd_Cnotzero,
blanchet@52903
   241
  collect_set_map = collect_set_map,
traytel@53772
   242
  in_bd = in_bd,
blanchet@49990
   243
  in_cong = in_cong,
blanchet@49990
   244
  in_mono = in_mono,
traytel@53030
   245
  in_rel = in_rel,
blanchet@54425
   246
  map_comp = map_comp,
blanchet@52899
   247
  map_cong = map_cong,
blanchet@54422
   248
  map_id = map_id,
traytel@53856
   249
  map_transfer = map_transfer,
blanchet@50606
   250
  rel_eq = rel_eq,
blanchet@50552
   251
  rel_flip = rel_flip,
blanchet@54427
   252
  set_map = set_map,
traytel@53030
   253
  rel_cong = rel_cong,
traytel@53030
   254
  rel_mono = rel_mono,
traytel@53053
   255
  rel_mono_strong = rel_mono_strong,
traytel@53030
   256
  rel_Grp = rel_Grp,
traytel@53030
   257
  rel_conversep = rel_conversep,
traytel@53030
   258
  rel_OO = rel_OO};
blanchet@49990
   259
blanchet@49990
   260
fun map_facts f {
blanchet@49990
   261
  bd_Card_order,
blanchet@49990
   262
  bd_Cinfinite,
blanchet@49990
   263
  bd_Cnotzero,
blanchet@52903
   264
  collect_set_map,
traytel@53772
   265
  in_bd,
blanchet@49990
   266
  in_cong,
blanchet@49990
   267
  in_mono,
traytel@53030
   268
  in_rel,
blanchet@54425
   269
  map_comp,
blanchet@52899
   270
  map_cong,
blanchet@54422
   271
  map_id,
traytel@53856
   272
  map_transfer,
blanchet@50606
   273
  rel_eq,
blanchet@50552
   274
  rel_flip,
blanchet@54427
   275
  set_map,
traytel@53030
   276
  rel_cong,
traytel@53030
   277
  rel_mono,
traytel@53053
   278
  rel_mono_strong,
traytel@53030
   279
  rel_Grp,
traytel@53030
   280
  rel_conversep,
traytel@53030
   281
  rel_OO} =
blanchet@49990
   282
  {bd_Card_order = f bd_Card_order,
blanchet@49990
   283
    bd_Cinfinite = f bd_Cinfinite,
blanchet@49990
   284
    bd_Cnotzero = f bd_Cnotzero,
blanchet@52903
   285
    collect_set_map = Lazy.map f collect_set_map,
traytel@53772
   286
    in_bd = Lazy.map f in_bd,
blanchet@49990
   287
    in_cong = Lazy.map f in_cong,
blanchet@49990
   288
    in_mono = Lazy.map f in_mono,
traytel@53030
   289
    in_rel = Lazy.map f in_rel,
blanchet@54425
   290
    map_comp = Lazy.map f map_comp,
blanchet@52899
   291
    map_cong = Lazy.map f map_cong,
blanchet@54422
   292
    map_id = Lazy.map f map_id,
traytel@53856
   293
    map_transfer = Lazy.map f map_transfer,
blanchet@50606
   294
    rel_eq = Lazy.map f rel_eq,
blanchet@50552
   295
    rel_flip = Lazy.map f rel_flip,
blanchet@54427
   296
    set_map = map (Lazy.map f) set_map,
traytel@53030
   297
    rel_cong = Lazy.map f rel_cong,
traytel@53030
   298
    rel_mono = Lazy.map f rel_mono,
traytel@53053
   299
    rel_mono_strong = Lazy.map f rel_mono_strong,
traytel@53030
   300
    rel_Grp = Lazy.map f rel_Grp,
traytel@53030
   301
    rel_conversep = Lazy.map f rel_conversep,
traytel@53030
   302
    rel_OO = Lazy.map f rel_OO};
blanchet@49990
   303
blanchet@49990
   304
val morph_facts = map_facts o Morphism.thm;
blanchet@49990
   305
blanchet@49990
   306
type nonemptiness_witness = {
blanchet@49990
   307
  I: int list,
blanchet@49990
   308
  wit: term,
blanchet@49990
   309
  prop: thm list
blanchet@49990
   310
};
blanchet@49990
   311
blanchet@49990
   312
fun mk_witness (I, wit) prop = {I = I, wit = wit, prop = prop};
blanchet@49990
   313
fun map_witness f g {I, wit, prop} = {I = I, wit = f wit, prop = map g prop};
blanchet@49990
   314
fun morph_witness phi = map_witness (Morphism.term phi) (Morphism.thm phi);
blanchet@49990
   315
blanchet@52974
   316
datatype bnf = BNF of {
blanchet@49990
   317
  name: binding,
blanchet@49990
   318
  T: typ,
blanchet@49990
   319
  live: int,
panny@54398
   320
  lives: typ list, (*source type variables of map*)
panny@54398
   321
  lives': typ list, (*target type variables of map*)
blanchet@49990
   322
  dead: int,
panny@54398
   323
  deads: typ list,
blanchet@49990
   324
  map: term,
blanchet@49990
   325
  sets: term list,
blanchet@49990
   326
  bd: term,
blanchet@49990
   327
  axioms: axioms,
blanchet@49990
   328
  defs: defs,
blanchet@49990
   329
  facts: facts,
blanchet@49990
   330
  nwits: int,
blanchet@49990
   331
  wits: nonemptiness_witness list,
traytel@53030
   332
  rel: term
blanchet@49990
   333
};
blanchet@49990
   334
blanchet@49990
   335
(* getters *)
blanchet@49990
   336
blanchet@49990
   337
fun rep_bnf (BNF bnf) = bnf;
blanchet@49990
   338
val name_of_bnf = #name o rep_bnf;
blanchet@49990
   339
val T_of_bnf = #T o rep_bnf;
blanchet@49990
   340
fun mk_T_of_bnf Ds Ts bnf =
blanchet@49990
   341
  let val bnf_rep = rep_bnf bnf
blanchet@49990
   342
  in Term.typ_subst_atomic ((#deads bnf_rep ~~ Ds) @ (#lives bnf_rep ~~ Ts)) (#T bnf_rep) end;
blanchet@49990
   343
val live_of_bnf = #live o rep_bnf;
blanchet@49990
   344
val lives_of_bnf = #lives o rep_bnf;
blanchet@49990
   345
val dead_of_bnf = #dead o rep_bnf;
blanchet@49990
   346
val deads_of_bnf = #deads o rep_bnf;
blanchet@49990
   347
val axioms_of_bnf = #axioms o rep_bnf;
blanchet@49990
   348
val facts_of_bnf = #facts o rep_bnf;
blanchet@49990
   349
val nwits_of_bnf = #nwits o rep_bnf;
blanchet@49990
   350
val wits_of_bnf = #wits o rep_bnf;
blanchet@49990
   351
blanchet@54168
   352
fun flatten_type_args_of_bnf bnf dead_x xs =
blanchet@54168
   353
  let
blanchet@54168
   354
    val Type (_, Ts) = T_of_bnf bnf;
blanchet@54168
   355
    val lives = lives_of_bnf bnf;
blanchet@54168
   356
    val deads = deads_of_bnf bnf;
blanchet@54168
   357
  in
blanchet@56822
   358
    permute_like_unique (op =) (deads @ lives) Ts (replicate (length deads) dead_x @ xs)
blanchet@54168
   359
  end;
blanchet@54168
   360
blanchet@49990
   361
(*terms*)
blanchet@49990
   362
val map_of_bnf = #map o rep_bnf;
blanchet@49990
   363
val sets_of_bnf = #sets o rep_bnf;
blanchet@49990
   364
fun mk_map_of_bnf Ds Ts Us bnf =
blanchet@49990
   365
  let val bnf_rep = rep_bnf bnf;
blanchet@49990
   366
  in
blanchet@49990
   367
    Term.subst_atomic_types
blanchet@49990
   368
      ((#deads bnf_rep ~~ Ds) @ (#lives bnf_rep ~~ Ts) @ (#lives' bnf_rep ~~ Us)) (#map bnf_rep)
blanchet@49990
   369
  end;
blanchet@49990
   370
fun mk_sets_of_bnf Dss Tss bnf =
blanchet@49990
   371
  let val bnf_rep = rep_bnf bnf;
blanchet@49990
   372
  in
blanchet@49990
   373
    map2 (fn (Ds, Ts) => Term.subst_atomic_types
blanchet@49990
   374
      ((#deads bnf_rep ~~ Ds) @ (#lives bnf_rep ~~ Ts))) (Dss ~~ Tss) (#sets bnf_rep)
blanchet@49990
   375
  end;
blanchet@49990
   376
val bd_of_bnf = #bd o rep_bnf;
blanchet@49990
   377
fun mk_bd_of_bnf Ds Ts bnf =
blanchet@49990
   378
  let val bnf_rep = rep_bnf bnf;
blanchet@49990
   379
  in Term.subst_atomic_types ((#deads bnf_rep ~~ Ds) @ (#lives bnf_rep ~~ Ts)) (#bd bnf_rep) end;
blanchet@49990
   380
fun mk_wits_of_bnf Dss Tss bnf =
blanchet@49990
   381
  let
blanchet@49990
   382
    val bnf_rep = rep_bnf bnf;
blanchet@49990
   383
    val wits = map (fn x => (#I x, #wit x)) (#wits bnf_rep);
blanchet@49990
   384
  in
blanchet@49990
   385
    map2 (fn (Ds, Ts) => apsnd (Term.subst_atomic_types
blanchet@49990
   386
      ((#deads bnf_rep ~~ Ds) @ (#lives bnf_rep ~~ Ts)))) (Dss ~~ Tss) wits
blanchet@49990
   387
  end;
blanchet@50522
   388
val rel_of_bnf = #rel o rep_bnf;
blanchet@50522
   389
fun mk_rel_of_bnf Ds Ts Us bnf =
blanchet@50477
   390
  let val bnf_rep = rep_bnf bnf;
blanchet@50477
   391
  in
blanchet@50477
   392
    Term.subst_atomic_types
blanchet@50522
   393
      ((#deads bnf_rep ~~ Ds) @ (#lives bnf_rep ~~ Ts) @ (#lives' bnf_rep ~~ Us)) (#rel bnf_rep)
blanchet@50477
   394
  end;
blanchet@49990
   395
blanchet@49990
   396
(*thms*)
blanchet@49990
   397
val bd_card_order_of_bnf = #bd_card_order o #axioms o rep_bnf;
blanchet@49990
   398
val bd_cinfinite_of_bnf = #bd_cinfinite o #axioms o rep_bnf;
blanchet@49990
   399
val bd_Card_order_of_bnf = #bd_Card_order o #facts o rep_bnf;
blanchet@49990
   400
val bd_Cinfinite_of_bnf = #bd_Cinfinite o #facts o rep_bnf;
blanchet@49990
   401
val bd_Cnotzero_of_bnf = #bd_Cnotzero o #facts o rep_bnf;
blanchet@52903
   402
val collect_set_map_of_bnf = Lazy.force o #collect_set_map o #facts o rep_bnf;
traytel@53772
   403
val in_bd_of_bnf = Lazy.force o #in_bd o #facts o rep_bnf;
blanchet@49990
   404
val in_cong_of_bnf = Lazy.force o #in_cong o #facts o rep_bnf;
blanchet@49990
   405
val in_mono_of_bnf = Lazy.force o #in_mono o #facts o rep_bnf;
traytel@53030
   406
val in_rel_of_bnf = Lazy.force o #in_rel o #facts o rep_bnf;
blanchet@49990
   407
val map_def_of_bnf = #map_def o #defs o rep_bnf;
blanchet@54407
   408
val map_id0_of_bnf = #map_id0 o #axioms o rep_bnf;
blanchet@54422
   409
val map_id_of_bnf = Lazy.force o #map_id o #facts o rep_bnf;
blanchet@54424
   410
val map_comp0_of_bnf = #map_comp0 o #axioms o rep_bnf;
blanchet@54425
   411
val map_comp_of_bnf = Lazy.force o #map_comp o #facts o rep_bnf;
blanchet@52898
   412
val map_cong0_of_bnf = #map_cong0 o #axioms o rep_bnf;
blanchet@52899
   413
val map_cong_of_bnf = Lazy.force o #map_cong o #facts o rep_bnf;
traytel@53868
   414
val map_transfer_of_bnf = Lazy.force o #map_transfer o #facts o rep_bnf;
traytel@56183
   415
val le_rel_OO_of_bnf = #le_rel_OO o #axioms o rep_bnf;
blanchet@50522
   416
val rel_def_of_bnf = #rel_def o #defs o rep_bnf;
blanchet@50606
   417
val rel_eq_of_bnf = Lazy.force o #rel_eq o #facts o rep_bnf;
blanchet@50552
   418
val rel_flip_of_bnf = Lazy.force o #rel_flip o #facts o rep_bnf;
blanchet@49990
   419
val set_bd_of_bnf = #set_bd o #axioms o rep_bnf;
blanchet@49990
   420
val set_defs_of_bnf = #set_defs o #defs o rep_bnf;
blanchet@54426
   421
val set_map0_of_bnf = #set_map0 o #axioms o rep_bnf;
blanchet@54427
   422
val set_map_of_bnf = map Lazy.force o #set_map o #facts o rep_bnf;
traytel@53030
   423
val rel_cong_of_bnf = Lazy.force o #rel_cong o #facts o rep_bnf;
traytel@53030
   424
val rel_mono_of_bnf = Lazy.force o #rel_mono o #facts o rep_bnf;
traytel@53053
   425
val rel_mono_strong_of_bnf = Lazy.force o #rel_mono_strong o #facts o rep_bnf;
traytel@53030
   426
val rel_Grp_of_bnf = Lazy.force o #rel_Grp o #facts o rep_bnf;
traytel@53030
   427
val rel_conversep_of_bnf = Lazy.force o #rel_conversep o #facts o rep_bnf;
traytel@53030
   428
val rel_OO_of_bnf = Lazy.force o #rel_OO o #facts o rep_bnf;
traytel@53030
   429
val rel_OO_Grp_of_bnf = #rel_OO_Grp o #axioms o rep_bnf;
blanchet@49990
   430
val wit_thms_of_bnf = maps #prop o wits_of_bnf;
blanchet@49990
   431
val wit_thmss_of_bnf = map #prop o wits_of_bnf;
blanchet@49990
   432
traytel@53030
   433
fun mk_bnf name T live lives lives' dead deads map sets bd axioms defs facts wits rel =
blanchet@49990
   434
  BNF {name = name, T = T,
blanchet@49990
   435
       live = live, lives = lives, lives' = lives', dead = dead, deads = deads,
blanchet@49990
   436
       map = map, sets = sets, bd = bd,
blanchet@49990
   437
       axioms = axioms, defs = defs, facts = facts,
traytel@53030
   438
       nwits = length wits, wits = wits, rel = rel};
blanchet@49990
   439
traytel@57358
   440
fun map_bnf f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16
traytel@57358
   441
  (BNF {name = name, T = T, live = live, lives = lives, lives' = lives',
blanchet@49990
   442
  dead = dead, deads = deads, map = map, sets = sets, bd = bd,
blanchet@49990
   443
  axioms = axioms, defs = defs, facts = facts,
traytel@53030
   444
  nwits = nwits, wits = wits, rel = rel}) =
traytel@57358
   445
  BNF {name = f1 name, T = f2 T,
traytel@57358
   446
       live = f3 live, lives = f4 lives, lives' = f5 lives', dead = f6 dead, deads = f7 deads,
traytel@57358
   447
       map = f8 map, sets = f9 sets, bd = f10 bd,
traytel@57358
   448
       axioms = f11 axioms, defs = f12 defs, facts = f13 facts,
traytel@57358
   449
       nwits = f14 nwits, wits = f15 wits, rel = f16 rel};
traytel@57358
   450
traytel@57358
   451
fun morph_bnf phi =
traytel@57358
   452
  let
traytel@57358
   453
    val Tphi = Morphism.typ phi;
traytel@57358
   454
    val tphi = Morphism.term phi;
traytel@57358
   455
  in
traytel@57358
   456
    map_bnf (Morphism.binding phi) Tphi I (map Tphi) (map Tphi) I (map Tphi) tphi (map tphi) tphi
traytel@57358
   457
      (morph_axioms phi) (morph_defs phi) (morph_facts phi) I (map (morph_witness phi)) tphi
traytel@57358
   458
  end;
traytel@57358
   459
traytel@57358
   460
fun morph_bnf_defs phi = map_bnf I I I I I I I I I I I (morph_defs phi) I I I I;
blanchet@49990
   461
blanchet@49990
   462
structure Data = Generic_Data
blanchet@49990
   463
(
blanchet@52974
   464
  type T = bnf Symtab.table;
blanchet@49990
   465
  val empty = Symtab.empty;
blanchet@49990
   466
  val extend = I;
blanchet@56736
   467
  fun merge data : T = Symtab.merge (K true) data;
blanchet@49990
   468
);
blanchet@49990
   469
traytel@54263
   470
fun bnf_of ctxt =
traytel@54263
   471
  Symtab.lookup (Data.get (Context.Proof ctxt))
wenzelm@56082
   472
  #> Option.map (morph_bnf (Morphism.transfer_morphism (Proof_Context.theory_of ctxt)));
blanchet@49990
   473
blanchet@49990
   474
blanchet@49990
   475
(* Utilities *)
blanchet@49990
   476
blanchet@49990
   477
fun normalize_set insts instA set =
blanchet@49990
   478
  let
blanchet@49990
   479
    val (T, T') = dest_funT (fastype_of set);
blanchet@49990
   480
    val A = fst (Term.dest_TVar (HOLogic.dest_setT T'));
blanchet@49990
   481
    val params = Term.add_tvar_namesT T [];
blanchet@49990
   482
  in Term.subst_TVars ((A :: params) ~~ (instA :: insts)) set end;
blanchet@49990
   483
blanchet@50522
   484
fun normalize_rel ctxt instTs instA instB rel =
blanchet@50477
   485
  let
blanchet@50477
   486
    val thy = Proof_Context.theory_of ctxt;
blanchet@50477
   487
    val tyenv =
blanchet@50522
   488
      Sign.typ_match thy (fastype_of rel, Library.foldr (op -->) (instTs, mk_pred2T instA instB))
blanchet@50478
   489
        Vartab.empty;
blanchet@50522
   490
  in Envir.subst_term (tyenv, Vartab.empty) rel end
blanchet@50468
   491
  handle Type.TYPE_MATCH => error "Bad relator";
blanchet@49990
   492
blanchet@49990
   493
fun normalize_wit insts CA As wit =
blanchet@49990
   494
  let
blanchet@49990
   495
    fun strip_param (Ts, T as Type (@{type_name fun}, [T1, T2])) =
blanchet@49990
   496
        if Type.raw_instance (CA, T) then (Ts, T) else strip_param (T1 :: Ts, T2)
blanchet@49990
   497
      | strip_param x = x;
blanchet@49990
   498
    val (Ts, T) = strip_param ([], fastype_of wit);
blanchet@49990
   499
    val subst = Term.add_tvar_namesT T [] ~~ insts;
blanchet@49990
   500
    fun find y = find_index (fn x => x = y) As;
blanchet@49990
   501
  in
blanchet@49990
   502
    (map (find o Term.typ_subst_TVars subst) (rev Ts), Term.subst_TVars subst wit)
blanchet@49990
   503
  end;
blanchet@49990
   504
blanchet@49990
   505
fun minimize_wits wits =
blanchet@49990
   506
 let
blanchet@49990
   507
   fun minimize done [] = done
traytel@50118
   508
     | minimize done ((I, wit) :: todo) =
blanchet@49990
   509
       if exists (fn (J, _) => subset (op =) (J, I)) (done @ todo)
blanchet@49990
   510
       then minimize done todo
blanchet@49990
   511
       else minimize ((I, wit) :: done) todo;
blanchet@49990
   512
 in minimize [] wits end;
blanchet@49990
   513
blanchet@55688
   514
fun mk_map live Ts Us t =
blanchet@55688
   515
  let val (Type (_, Ts0), Type (_, Us0)) = strip_typeN (live + 1) (fastype_of t) |>> List.last in
blanchet@55688
   516
    Term.subst_atomic_types (Ts0 @ Us0 ~~ Ts @ Us) t
blanchet@55688
   517
  end;
blanchet@55688
   518
blanchet@55688
   519
fun mk_rel live Ts Us t =
blanchet@55688
   520
  let val [Type (_, Ts0), Type (_, Us0)] = binder_types (snd (strip_typeN live (fastype_of t))) in
blanchet@55688
   521
    Term.subst_atomic_types (Ts0 @ Us0 ~~ Ts @ Us) t
blanchet@55688
   522
  end;
blanchet@55688
   523
blanchet@55689
   524
fun build_map_or_rel mk const of_bnf dest ctxt build_simple =
blanchet@55688
   525
  let
blanchet@55688
   526
    fun build (TU as (T, U)) =
blanchet@55688
   527
      if T = U then
blanchet@55688
   528
        const T
blanchet@55688
   529
      else
blanchet@55688
   530
        (case TU of
blanchet@55688
   531
          (Type (s, Ts), Type (s', Us)) =>
blanchet@55688
   532
          if s = s' then
blanchet@55688
   533
            let
blanchet@55689
   534
              val bnf = the (bnf_of ctxt s);
blanchet@55688
   535
              val live = live_of_bnf bnf;
blanchet@55688
   536
              val mapx = mk live Ts Us (of_bnf bnf);
blanchet@55688
   537
              val TUs' = map dest (fst (strip_typeN live (fastype_of mapx)));
blanchet@55688
   538
            in Term.list_comb (mapx, map build TUs') end
blanchet@55688
   539
          else
blanchet@55688
   540
            build_simple TU
blanchet@55688
   541
        | _ => build_simple TU);
blanchet@55688
   542
  in build end;
blanchet@55688
   543
blanchet@55688
   544
val build_map = build_map_or_rel mk_map HOLogic.id_const map_of_bnf dest_funT;
blanchet@55688
   545
val build_rel = build_map_or_rel mk_rel HOLogic.eq_const rel_of_bnf dest_pred2T;
blanchet@49990
   546
blanchet@55698
   547
fun map_flattened_map_args ctxt s map_args fs =
blanchet@55698
   548
  let
blanchet@55698
   549
    val flat_fs = flatten_type_args_of_bnf (the (bnf_of ctxt s)) Term.dummy fs;
blanchet@55698
   550
    val flat_fs' = map_args flat_fs;
blanchet@55698
   551
  in
blanchet@56822
   552
    permute_like_unique (op aconv) flat_fs fs flat_fs'
blanchet@55698
   553
  end;
blanchet@55698
   554
blanchet@49990
   555
blanchet@49990
   556
(* Names *)
blanchet@49990
   557
blanchet@49990
   558
val mapN = "map";
blanchet@49990
   559
val setN = "set";
blanchet@49990
   560
fun mk_setN i = setN ^ nonzero_string_of_int i;
blanchet@49990
   561
val bdN = "bd";
blanchet@49990
   562
val witN = "wit";
blanchet@49990
   563
fun mk_witN i = witN ^ nonzero_string_of_int i;
blanchet@50522
   564
val relN = "rel";
blanchet@49990
   565
blanchet@49990
   566
val bd_card_orderN = "bd_card_order";
blanchet@49990
   567
val bd_cinfiniteN = "bd_cinfinite";
blanchet@49990
   568
val bd_Card_orderN = "bd_Card_order";
blanchet@49990
   569
val bd_CinfiniteN = "bd_Cinfinite";
blanchet@49990
   570
val bd_CnotzeroN = "bd_Cnotzero";
blanchet@52903
   571
val collect_set_mapN = "collect_set_map";
blanchet@49990
   572
val in_bdN = "in_bd";
blanchet@49990
   573
val in_monoN = "in_mono";
traytel@53030
   574
val in_relN = "in_rel";
blanchet@54407
   575
val map_id0N = "map_id0";
blanchet@54422
   576
val map_idN = "map_id";
blanchet@54424
   577
val map_comp0N = "map_comp0";
blanchet@54425
   578
val map_compN = "map_comp";
blanchet@52898
   579
val map_cong0N = "map_cong0";
blanchet@52899
   580
val map_congN = "map_cong";
traytel@53856
   581
val map_transferN = "map_transfer";
blanchet@50606
   582
val rel_eqN = "rel_eq";
blanchet@50552
   583
val rel_flipN = "rel_flip";
blanchet@54426
   584
val set_map0N = "set_map0";
blanchet@54427
   585
val set_mapN = "set_map";
blanchet@50552
   586
val set_bdN = "set_bd";
traytel@53030
   587
val rel_GrpN = "rel_Grp";
traytel@53030
   588
val rel_conversepN = "rel_conversep";
traytel@53030
   589
val rel_monoN = "rel_mono"
traytel@53053
   590
val rel_mono_strongN = "rel_mono_strong"
blanchet@55993
   591
val rel_comppN = "rel_compp";
blanchet@55993
   592
val rel_compp_GrpN = "rel_compp_Grp";
blanchet@49990
   593
blanchet@57196
   594
datatype inline_policy = Dont_Inline | Hardly_Inline | Smart_Inline | Do_Inline;
blanchet@49990
   595
blanchet@50553
   596
datatype fact_policy = Dont_Note | Note_Some | Note_All;
blanchet@49990
   597
blanchet@49990
   598
val bnf_note_all = Attrib.setup_config_bool @{binding bnf_note_all} (K false);
traytel@54280
   599
val bnf_timing = Attrib.setup_config_bool @{binding bnf_timing} (K false);
blanchet@49990
   600
blanchet@50553
   601
fun user_policy policy ctxt = if Config.get ctxt bnf_note_all then Note_All else policy;
blanchet@49990
   602
blanchet@57196
   603
val smart_max_inline_term_size = 25; (*FUDGE*)
blanchet@49990
   604
traytel@55182
   605
fun note_bnf_thms fact_policy qualify' bnf_b bnf =
traytel@53857
   606
  let
traytel@53857
   607
    val axioms = axioms_of_bnf bnf;
traytel@53857
   608
    val facts = facts_of_bnf bnf;
traytel@53857
   609
    val wits = wits_of_bnf bnf;
traytel@55182
   610
    val qualify =
traytel@55182
   611
      let val (_, qs, _) = Binding.dest bnf_b;
traytel@55182
   612
      in fold_rev (fn (s, mand) => Binding.qualify mand s) qs #> qualify' end;
traytel@53857
   613
  in
traytel@53857
   614
    (if fact_policy = Note_All then
traytel@53857
   615
      let
traytel@53857
   616
        val witNs = if length wits = 1 then [witN] else map mk_witN (1 upto length wits);
traytel@53857
   617
        val notes =
traytel@53857
   618
          [(bd_card_orderN, [#bd_card_order axioms]),
traytel@53857
   619
            (bd_cinfiniteN, [#bd_cinfinite axioms]),
traytel@53857
   620
            (bd_Card_orderN, [#bd_Card_order facts]),
traytel@53857
   621
            (bd_CinfiniteN, [#bd_Cinfinite facts]),
traytel@53857
   622
            (bd_CnotzeroN, [#bd_Cnotzero facts]),
traytel@53857
   623
            (collect_set_mapN, [Lazy.force (#collect_set_map facts)]),
traytel@53857
   624
            (in_bdN, [Lazy.force (#in_bd facts)]),
traytel@53857
   625
            (in_monoN, [Lazy.force (#in_mono facts)]),
traytel@53857
   626
            (in_relN, [Lazy.force (#in_rel facts)]),
blanchet@54424
   627
            (map_comp0N, [#map_comp0 axioms]),
blanchet@54407
   628
            (map_id0N, [#map_id0 axioms]),
traytel@53857
   629
            (map_transferN, [Lazy.force (#map_transfer facts)]),
traytel@53857
   630
            (rel_mono_strongN, [Lazy.force (#rel_mono_strong facts)]),
blanchet@54426
   631
            (set_map0N, #set_map0 axioms),
traytel@53857
   632
            (set_bdN, #set_bd axioms)] @
traytel@53857
   633
            (witNs ~~ wit_thmss_of_bnf bnf)
traytel@53857
   634
            |> map (fn (thmN, thms) =>
blanchet@54402
   635
              ((qualify (Binding.qualify true (Binding.name_of bnf_b) (Binding.name thmN)), []),
traytel@53857
   636
              [(thms, [])]));
traytel@53857
   637
        in
traytel@53857
   638
          Local_Theory.notes notes #> snd
traytel@53857
   639
        end
traytel@53857
   640
      else
traytel@53857
   641
        I)
traytel@53857
   642
    #> (if fact_policy <> Dont_Note then
traytel@53857
   643
        let
traytel@53857
   644
          val notes =
blanchet@54425
   645
            [(map_compN, [Lazy.force (#map_comp facts)], []),
traytel@53857
   646
            (map_cong0N, [#map_cong0 axioms], []),
blanchet@55997
   647
            (map_congN, [Lazy.force (#map_cong facts)], fundefcong_attrs),
blanchet@54422
   648
            (map_idN, [Lazy.force (#map_id facts)], []),
blanchet@55993
   649
            (rel_comppN, [Lazy.force (#rel_OO facts)], []),
blanchet@55993
   650
            (rel_compp_GrpN, no_refl [#rel_OO_Grp axioms], []),
blanchet@55993
   651
            (rel_conversepN, [Lazy.force (#rel_conversep facts)], []),
traytel@53857
   652
            (rel_eqN, [Lazy.force (#rel_eq facts)], []),
traytel@53857
   653
            (rel_flipN, [Lazy.force (#rel_flip facts)], []),
traytel@53857
   654
            (rel_GrpN, [Lazy.force (#rel_Grp facts)], []),
traytel@53857
   655
            (rel_monoN, [Lazy.force (#rel_mono facts)], []),
blanchet@55993
   656
            (set_mapN, map Lazy.force (#set_map facts), [])]
traytel@53857
   657
            |> filter_out (null o #2)
traytel@53857
   658
            |> map (fn (thmN, thms, attrs) =>
blanchet@54402
   659
              ((qualify (Binding.qualify true (Binding.name_of bnf_b) (Binding.name thmN)),
traytel@53857
   660
                attrs), [(thms, [])]));
traytel@53857
   661
        in
traytel@53857
   662
          Local_Theory.notes notes #> snd
traytel@53857
   663
        end
traytel@53857
   664
      else
traytel@53857
   665
        I)
traytel@53857
   666
  end;
traytel@53857
   667
blanchet@49990
   668
blanchet@49990
   669
(* Define new BNFs *)
blanchet@49990
   670
traytel@57358
   671
fun define_bnf_consts const_policy fact_policy internal Ds_opt map_b rel_b set_bs
traytel@56183
   672
  ((((((bnf_b, T_rhs), map_rhs), set_rhss), bd_rhs), wit_rhss), rel_rhs_opt) no_defs_lthy =
blanchet@49990
   673
  let
traytel@56183
   674
    val live = length set_rhss;
blanchet@54402
   675
blanchet@55608
   676
    val def_qualify = Binding.conceal o Binding.qualify false (Binding.name_of bnf_b);
blanchet@54402
   677
blanchet@55863
   678
    fun mk_prefix_binding pre = Binding.prefix_name (pre ^ "_") bnf_b;
traytel@50449
   679
blanchet@50478
   680
    fun maybe_define user_specified (b, rhs) lthy =
blanchet@49990
   681
      let
blanchet@49990
   682
        val inline =
blanchet@50553
   683
          (user_specified orelse fact_policy = Dont_Note) andalso
blanchet@49990
   684
          (case const_policy of
blanchet@49990
   685
            Dont_Inline => false
blanchet@49990
   686
          | Hardly_Inline => Term.is_Free rhs orelse Term.is_Const rhs
blanchet@57196
   687
          | Smart_Inline => Term.size_of_term rhs <= smart_max_inline_term_size
blanchet@49990
   688
          | Do_Inline => true)
blanchet@49990
   689
      in
blanchet@49990
   690
        if inline then
blanchet@50475
   691
          ((rhs, Drule.reflexive_thm), lthy)
blanchet@49990
   692
        else
blanchet@49990
   693
          let val b = b () in
traytel@57358
   694
            apfst (apsnd snd)
traytel@57358
   695
              ((if internal then Local_Theory.define_internal else Local_Theory.define)
traytel@57358
   696
                ((b, NoSyn), ((Thm.def_binding b, []), rhs)) lthy)
blanchet@49990
   697
          end
blanchet@49990
   698
      end;
blanchet@49990
   699
blanchet@50474
   700
    fun maybe_restore lthy_old lthy =
blanchet@50474
   701
      lthy |> not (pointer_eq (lthy_old, lthy)) ? Local_Theory.restore;
blanchet@50474
   702
blanchet@52895
   703
    val map_bind_def =
blanchet@55863
   704
      (fn () => def_qualify (if Binding.is_empty map_b then mk_prefix_binding mapN else map_b),
blanchet@54402
   705
         map_rhs);
blanchet@50474
   706
    val set_binds_defs =
blanchet@50474
   707
      let
blanchet@52894
   708
        fun set_name i get_b =
blanchet@52894
   709
          (case try (nth set_bs) (i - 1) of
blanchet@52894
   710
            SOME b => if Binding.is_empty b then get_b else K b
blanchet@54402
   711
          | NONE => get_b) #> def_qualify;
blanchet@55863
   712
        val bs = if live = 1 then [set_name 1 (fn () => mk_prefix_binding setN)]
blanchet@55863
   713
          else map (fn i => set_name i (fn () => mk_prefix_binding (mk_setN i))) (1 upto live);
blanchet@52894
   714
      in bs ~~ set_rhss end;
blanchet@55863
   715
    val bd_bind_def = (fn () => def_qualify (mk_prefix_binding bdN), bd_rhs);
blanchet@50474
   716
traytel@55641
   717
    val ((((bnf_map_term, raw_map_def),
blanchet@49990
   718
      (bnf_set_terms, raw_set_defs)),
traytel@55641
   719
      (bnf_bd_term, raw_bd_def)), (lthy, lthy_old)) =
blanchet@49990
   720
        no_defs_lthy
blanchet@50478
   721
        |> maybe_define true map_bind_def
blanchet@50478
   722
        ||>> apfst split_list o fold_map (maybe_define true) set_binds_defs
blanchet@50478
   723
        ||>> maybe_define true bd_bind_def
blanchet@49990
   724
        ||> `(maybe_restore no_defs_lthy);
blanchet@49990
   725
blanchet@50474
   726
    val phi = Proof_Context.export_morphism lthy_old lthy;
blanchet@49990
   727
traytel@56183
   728
blanchet@49990
   729
    val bnf_map_def = Morphism.thm phi raw_map_def;
blanchet@49990
   730
    val bnf_set_defs = map (Morphism.thm phi) raw_set_defs;
blanchet@49990
   731
    val bnf_bd_def = Morphism.thm phi raw_bd_def;
blanchet@49990
   732
blanchet@49990
   733
    val bnf_map = Morphism.term phi bnf_map_term;
blanchet@49990
   734
blanchet@49990
   735
    (*TODO: handle errors*)
blanchet@49990
   736
    (*simple shape analysis of a map function*)
traytel@56183
   737
    val ((alphas, betas), (Calpha, _)) =
traytel@50410
   738
      fastype_of bnf_map
traytel@50410
   739
      |> strip_typeN live
traytel@50410
   740
      |>> map_split dest_funT
traytel@50410
   741
      ||> dest_funT
traytel@50410
   742
      handle TYPE _ => error "Bad map function";
blanchet@49990
   743
traytel@56183
   744
    val Calpha_params = map TVar (Term.add_tvarsT Calpha []);
blanchet@49990
   745
traytel@55799
   746
    val bnf_T = Morphism.typ phi T_rhs;
traytel@55799
   747
    val bad_args = Term.add_tfreesT bnf_T [];
traytel@55799
   748
    val _ = if null bad_args then () else error ("Locally fixed type arguments " ^
traytel@55799
   749
      commas_quote (map (Syntax.string_of_typ no_defs_lthy o TFree) bad_args));
traytel@55799
   750
traytel@56183
   751
    val bnf_sets =
traytel@56183
   752
      map2 (normalize_set Calpha_params) alphas (map (Morphism.term phi) bnf_set_terms);
blanchet@49990
   753
    val bnf_bd =
traytel@56183
   754
      Term.subst_TVars (Term.add_tvar_namesT bnf_T [] ~~ Calpha_params)
traytel@56183
   755
        (Morphism.term phi bnf_bd_term);
blanchet@49990
   756
blanchet@49990
   757
    (*TODO: assert Ds = (TVars of bnf_map) \ (alphas @ betas) as sets*)
blanchet@49990
   758
    val deads = (case Ds_opt of
blanchet@49990
   759
      NONE => subtract (op =) (alphas @ betas) (map TVar (Term.add_tvars bnf_map []))
blanchet@49990
   760
    | SOME Ds => map (Morphism.typ phi) Ds);
blanchet@49990
   761
blanchet@49990
   762
    (*TODO: further checks of type of bnf_map*)
blanchet@49990
   763
    (*TODO: check types of bnf_sets*)
blanchet@49990
   764
    (*TODO: check type of bnf_bd*)
blanchet@50522
   765
    (*TODO: check type of bnf_rel*)
blanchet@49990
   766
traytel@56183
   767
    fun mk_bnf_map Ds As' Bs' =
traytel@56183
   768
      Term.subst_atomic_types ((deads ~~ Ds) @ (alphas ~~ As') @ (betas ~~ Bs')) bnf_map;
traytel@56183
   769
    fun mk_bnf_t Ds As' = Term.subst_atomic_types ((deads ~~ Ds) @ (alphas ~~ As'));
traytel@56183
   770
    fun mk_bnf_T Ds As' = Term.typ_subst_atomic ((deads ~~ Ds) @ (alphas ~~ As'));
traytel@56183
   771
traytel@56183
   772
    val (((As, Bs), Ds), names_lthy) = lthy
blanchet@49990
   773
      |> mk_TFrees live
blanchet@49990
   774
      ||>> mk_TFrees live
traytel@56183
   775
      ||>> mk_TFrees (length deads);
traytel@56183
   776
    val RTs = map2 (curry HOLogic.mk_prodT) As Bs;
traytel@56183
   777
    val pred2RTs = map2 mk_pred2T As Bs;
traytel@56183
   778
    val (Rs, Rs') = names_lthy |> mk_Frees' "R" pred2RTs |> fst
traytel@56183
   779
    val CA = mk_bnf_T Ds As Calpha;
traytel@56183
   780
    val CR = mk_bnf_T Ds RTs Calpha;
traytel@56183
   781
    val setRs =
traytel@56183
   782
      map3 (fn R => fn T => fn U =>
traytel@56183
   783
          HOLogic.Collect_const (HOLogic.mk_prodT (T, U)) $ HOLogic.mk_split R) Rs As Bs;
traytel@56183
   784
traytel@56183
   785
    (*Grp (in (Collect (split R1) .. Collect (split Rn))) (map fst .. fst)^--1 OO
traytel@56183
   786
      Grp (in (Collect (split R1) .. Collect (split Rn))) (map snd .. snd)*)
traytel@56183
   787
    val OO_Grp =
traytel@56183
   788
      let
traytel@56183
   789
        val map1 = Term.list_comb (mk_bnf_map Ds RTs As, map fst_const RTs);
traytel@56183
   790
        val map2 = Term.list_comb (mk_bnf_map Ds RTs Bs, map snd_const RTs);
traytel@56183
   791
        val bnf_in = mk_in setRs (map (mk_bnf_t Ds RTs) bnf_sets) CR;
traytel@56183
   792
      in
traytel@56183
   793
        mk_rel_compp (mk_conversep (mk_Grp bnf_in map1), mk_Grp bnf_in map2)
traytel@56183
   794
        |> fold_rev Term.absfree Rs'
traytel@56183
   795
      end;
traytel@56183
   796
traytel@56183
   797
    val rel_rhs = the_default OO_Grp rel_rhs_opt;
traytel@56183
   798
traytel@56183
   799
    val rel_bind_def =
traytel@56183
   800
      (fn () => def_qualify (if Binding.is_empty rel_b then mk_prefix_binding relN else rel_b),
traytel@56183
   801
         rel_rhs);
traytel@56183
   802
traytel@56183
   803
    val wit_rhss =
traytel@56183
   804
      if null wit_rhss then
traytel@56183
   805
        [fold_rev Term.absdummy As (Term.list_comb (mk_bnf_map Ds As As,
traytel@56183
   806
          map2 (fn T => fn i => Term.absdummy T (Bound i)) As (live downto 1)) $
traytel@56183
   807
          Const (@{const_name undefined}, CA))]
traytel@56183
   808
      else wit_rhss;
traytel@56183
   809
    val nwits = length wit_rhss;
traytel@56183
   810
    val wit_binds_defs =
traytel@56183
   811
      let
traytel@56183
   812
        val bs = if nwits = 1 then [fn () => def_qualify (mk_prefix_binding witN)]
traytel@56183
   813
          else map (fn i => fn () => def_qualify (mk_prefix_binding (mk_witN i))) (1 upto nwits);
traytel@56183
   814
      in bs ~~ wit_rhss end;
traytel@56183
   815
traytel@56183
   816
    val (((bnf_rel_term, raw_rel_def), (bnf_wit_terms, raw_wit_defs)), (lthy, lthy_old)) =
traytel@56183
   817
      lthy
traytel@56183
   818
      |> maybe_define (is_some rel_rhs_opt) rel_bind_def
traytel@56183
   819
      ||>> apfst split_list o fold_map (maybe_define (not (null wit_rhss))) wit_binds_defs
traytel@56183
   820
      ||> `(maybe_restore lthy);
traytel@56183
   821
traytel@56183
   822
    val phi = Proof_Context.export_morphism lthy_old lthy;
traytel@56183
   823
    val bnf_rel_def = Morphism.thm phi raw_rel_def;
traytel@56183
   824
    val bnf_rel = Morphism.term phi bnf_rel_term;
traytel@56183
   825
    fun mk_bnf_rel Ds As' Bs' =
traytel@56183
   826
      normalize_rel lthy (map2 mk_pred2T As' Bs') (mk_bnf_T Ds As' Calpha) (mk_bnf_T Ds Bs' Calpha)
traytel@56183
   827
        bnf_rel;
traytel@56183
   828
traytel@56183
   829
    val bnf_wit_defs = map (Morphism.thm phi) raw_wit_defs;
traytel@56183
   830
    val bnf_wits =
traytel@56183
   831
      map (normalize_wit Calpha_params Calpha alphas o Morphism.term phi) bnf_wit_terms;
traytel@56183
   832
traytel@56183
   833
    fun mk_OO_Grp Ds' As' Bs' =
traytel@56183
   834
      Term.subst_atomic_types ((Ds ~~ Ds') @ (As ~~ As') @ (Bs ~~ Bs')) OO_Grp;
traytel@56183
   835
  in
traytel@56183
   836
    (((alphas, betas, deads, Calpha),
traytel@56183
   837
     (bnf_map, bnf_sets, bnf_bd, bnf_wits, bnf_rel),
traytel@56183
   838
     (bnf_map_def, bnf_set_defs, bnf_bd_def, bnf_wit_defs, bnf_rel_def),
traytel@56183
   839
     (mk_bnf_map, mk_bnf_t, mk_bnf_T, mk_bnf_rel, mk_OO_Grp)), lthy)
traytel@56183
   840
  end;
traytel@56183
   841
traytel@57358
   842
fun prepare_def const_policy mk_fact_policy internal qualify prep_typ prep_term Ds_opt map_b rel_b
traytel@57358
   843
  set_bs ((((((raw_bnf_b, raw_bnf_T), raw_map), raw_sets), raw_bd), raw_wits), raw_rel_opt)
traytel@56183
   844
  no_defs_lthy =
traytel@56183
   845
  let
traytel@56183
   846
    val fact_policy = mk_fact_policy no_defs_lthy;
traytel@56183
   847
    val bnf_b = qualify raw_bnf_b;
traytel@56183
   848
    val live = length raw_sets;
traytel@56183
   849
traytel@56183
   850
    val T_rhs = prep_typ no_defs_lthy raw_bnf_T;
traytel@56183
   851
    val map_rhs = prep_term no_defs_lthy raw_map;
traytel@56183
   852
    val set_rhss = map (prep_term no_defs_lthy) raw_sets;
traytel@56183
   853
    val bd_rhs = prep_term no_defs_lthy raw_bd;
traytel@56183
   854
    val wit_rhss = map (prep_term no_defs_lthy) raw_wits;
traytel@56183
   855
    val rel_rhs_opt = Option.map (prep_term no_defs_lthy) raw_rel_opt;
traytel@56183
   856
traytel@56183
   857
    fun err T =
traytel@56183
   858
      error ("Trying to register the type " ^ quote (Syntax.string_of_typ no_defs_lthy T) ^
traytel@56183
   859
        " as unnamed BNF");
traytel@56183
   860
traytel@56183
   861
    val (bnf_b, key) =
traytel@56183
   862
      if Binding.eq_name (bnf_b, Binding.empty) then
traytel@56183
   863
        (case T_rhs of
traytel@56183
   864
          Type (C, Ts) => if forall (can dest_TFree) Ts
traytel@56183
   865
            then (Binding.qualified_name C, C) else err T_rhs
traytel@56183
   866
        | T => err T)
traytel@56183
   867
      else (bnf_b, Local_Theory.full_name no_defs_lthy bnf_b);
traytel@56183
   868
traytel@56183
   869
    val (((alphas, betas, deads, Calpha),
traytel@56183
   870
     (bnf_map, bnf_sets, bnf_bd, bnf_wits, bnf_rel),
traytel@56183
   871
     (bnf_map_def, bnf_set_defs, bnf_bd_def, bnf_wit_defs, bnf_rel_def),
traytel@56183
   872
     (mk_bnf_map_Ds, mk_bnf_t_Ds, mk_bnf_T_Ds, _, mk_OO_Grp)), lthy) =
traytel@57358
   873
       define_bnf_consts const_policy fact_policy internal Ds_opt map_b rel_b set_bs
traytel@56183
   874
         ((((((bnf_b, T_rhs), map_rhs), set_rhss), bd_rhs), wit_rhss), rel_rhs_opt) no_defs_lthy;
traytel@56183
   875
traytel@56183
   876
    val dead = length deads;
traytel@56183
   877
traytel@56183
   878
    val ((((((As', Bs'), Cs), Ds), B1Ts), B2Ts), (Ts, T)) = lthy
traytel@56183
   879
      |> mk_TFrees live
traytel@56183
   880
      ||>> mk_TFrees live
blanchet@49990
   881
      ||>> mk_TFrees live
blanchet@49990
   882
      ||>> mk_TFrees dead
blanchet@49990
   883
      ||>> mk_TFrees live
blanchet@49990
   884
      ||>> mk_TFrees live
blanchet@49990
   885
      ||> fst o mk_TFrees 1
blanchet@49990
   886
      ||> the_single
blanchet@49990
   887
      ||> `(replicate live);
blanchet@49990
   888
traytel@56183
   889
    val mk_bnf_map = mk_bnf_map_Ds Ds;
traytel@56183
   890
    val mk_bnf_t = mk_bnf_t_Ds Ds;
traytel@56183
   891
    val mk_bnf_T = mk_bnf_T_Ds Ds;
blanchet@50468
   892
traytel@53030
   893
    val pred2RTs = map2 mk_pred2T As' Bs';
traytel@53030
   894
    val pred2RTsAsCs = map2 mk_pred2T As' Cs;
traytel@53030
   895
    val pred2RTsBsCs = map2 mk_pred2T Bs' Cs;
traytel@53030
   896
    val pred2RT's = map2 mk_pred2T Bs' As';
traytel@53030
   897
    val self_pred2RTs = map2 mk_pred2T As' As';
traytel@53856
   898
    val transfer_domRTs = map2 mk_pred2T As' B1Ts;
traytel@53856
   899
    val transfer_ranRTs = map2 mk_pred2T Bs' B2Ts;
blanchet@49990
   900
traytel@56183
   901
    val CA' = mk_bnf_T As' Calpha;
traytel@56183
   902
    val CB' = mk_bnf_T Bs' Calpha;
traytel@56183
   903
    val CC' = mk_bnf_T Cs Calpha;
traytel@56183
   904
    val CB1 = mk_bnf_T B1Ts Calpha;
traytel@56183
   905
    val CB2 = mk_bnf_T B2Ts Calpha;
blanchet@50468
   906
blanchet@49990
   907
    val bnf_map_AsAs = mk_bnf_map As' As';
blanchet@49990
   908
    val bnf_map_AsBs = mk_bnf_map As' Bs';
blanchet@49990
   909
    val bnf_map_AsCs = mk_bnf_map As' Cs;
blanchet@49990
   910
    val bnf_map_BsCs = mk_bnf_map Bs' Cs;
blanchet@49990
   911
    val bnf_sets_As = map (mk_bnf_t As') bnf_sets;
blanchet@49990
   912
    val bnf_sets_Bs = map (mk_bnf_t Bs') bnf_sets;
blanchet@49990
   913
    val bnf_bd_As = mk_bnf_t As' bnf_bd;
traytel@56183
   914
    fun mk_bnf_rel RTs CA CB = normalize_rel lthy RTs CA CB bnf_rel;
blanchet@49990
   915
blanchet@50610
   916
    val pre_names_lthy = lthy;
traytel@56183
   917
    val (((((((((((((((fs, gs), hs), x), y), zs), ys), As),
traytel@56183
   918
      As_copy), bs), Rs), Rs_copy), Ss),
traytel@53856
   919
      transfer_domRs), transfer_ranRs), names_lthy) = pre_names_lthy
traytel@54060
   920
      |> mk_Frees "f" (map2 (curry op -->) As' Bs')
traytel@54060
   921
      ||>> mk_Frees "g" (map2 (curry op -->) Bs' Cs)
traytel@54060
   922
      ||>> mk_Frees "h" (map2 (curry op -->) As' Ts)
traytel@53031
   923
      ||>> yield_singleton (mk_Frees "x") CA'
traytel@53031
   924
      ||>> yield_singleton (mk_Frees "y") CB'
blanchet@49990
   925
      ||>> mk_Frees "z" As'
traytel@53053
   926
      ||>> mk_Frees "y" Bs'
blanchet@49990
   927
      ||>> mk_Frees "A" (map HOLogic.mk_setT As')
blanchet@49990
   928
      ||>> mk_Frees "A" (map HOLogic.mk_setT As')
blanchet@49990
   929
      ||>> mk_Frees "b" As'
traytel@56183
   930
      ||>> mk_Frees "R" pred2RTs
traytel@53030
   931
      ||>> mk_Frees "R" pred2RTs
traytel@53856
   932
      ||>> mk_Frees "S" pred2RTsBsCs
traytel@53856
   933
      ||>> mk_Frees "R" transfer_domRTs
traytel@53856
   934
      ||>> mk_Frees "S" transfer_ranRTs;
blanchet@49990
   935
blanchet@52899
   936
    val fs_copy = map2 (retype_free o fastype_of) fs gs;
blanchet@52899
   937
    val x_copy = retype_free CA' y;
blanchet@52899
   938
traytel@53030
   939
    val rel = mk_bnf_rel pred2RTs CA' CB';
traytel@53856
   940
    val relAsAs = mk_bnf_rel self_pred2RTs CA' CA';
traytel@55641
   941
    val bnf_wit_As = map (apsnd (mk_bnf_t As')) bnf_wits;
traytel@55641
   942
blanchet@54407
   943
    val map_id0_goal =
blanchet@52899
   944
      let val bnf_map_app_id = Term.list_comb (bnf_map_AsAs, map HOLogic.id_const As') in
blanchet@52899
   945
        mk_Trueprop_eq (bnf_map_app_id, HOLogic.id_const CA')
blanchet@49990
   946
      end;
blanchet@49990
   947
blanchet@54424
   948
    val map_comp0_goal =
blanchet@49990
   949
      let
blanchet@50033
   950
        val bnf_map_app_comp = Term.list_comb (bnf_map_AsCs, map2 (curry HOLogic.mk_comp) gs fs);
blanchet@49990
   951
        val comp_bnf_map_app = HOLogic.mk_comp
blanchet@52899
   952
          (Term.list_comb (bnf_map_BsCs, gs), Term.list_comb (bnf_map_AsBs, fs));
blanchet@49990
   953
      in
blanchet@50138
   954
        fold_rev Logic.all (fs @ gs) (mk_Trueprop_eq (bnf_map_app_comp, comp_bnf_map_app))
blanchet@49990
   955
      end;
blanchet@49990
   956
blanchet@52899
   957
    fun mk_map_cong_prem x z set f f_copy =
blanchet@52899
   958
      Logic.all z (Logic.mk_implies
blanchet@52899
   959
        (HOLogic.mk_Trueprop (HOLogic.mk_mem (z, set $ x)),
blanchet@52899
   960
        mk_Trueprop_eq (f $ z, f_copy $ z)));
blanchet@52899
   961
blanchet@52898
   962
    val map_cong0_goal =
blanchet@49990
   963
      let
blanchet@52899
   964
        val prems = map4 (mk_map_cong_prem x) zs bnf_sets_As fs fs_copy;
blanchet@52899
   965
        val eq = mk_Trueprop_eq (Term.list_comb (bnf_map_AsBs, fs) $ x,
blanchet@49990
   966
          Term.list_comb (bnf_map_AsBs, fs_copy) $ x);
blanchet@49990
   967
      in
blanchet@52899
   968
        fold_rev Logic.all (x :: fs @ fs_copy) (Logic.list_implies (prems, eq))
blanchet@49990
   969
      end;
blanchet@49990
   970
blanchet@54426
   971
    val set_map0s_goal =
blanchet@49990
   972
      let
blanchet@49990
   973
        fun mk_goal setA setB f =
blanchet@49990
   974
          let
blanchet@49990
   975
            val set_comp_map =
blanchet@49990
   976
              HOLogic.mk_comp (setB, Term.list_comb (bnf_map_AsBs, fs));
blanchet@49990
   977
            val image_comp_set = HOLogic.mk_comp (mk_image f, setA);
blanchet@49990
   978
          in
blanchet@50138
   979
            fold_rev Logic.all fs (mk_Trueprop_eq (set_comp_map, image_comp_set))
blanchet@49990
   980
          end;
blanchet@49990
   981
      in
blanchet@49990
   982
        map3 mk_goal bnf_sets_As bnf_sets_Bs fs
blanchet@49990
   983
      end;
blanchet@49990
   984
blanchet@50473
   985
    val card_order_bd_goal = HOLogic.mk_Trueprop (mk_card_order bnf_bd_As);
blanchet@49990
   986
blanchet@50473
   987
    val cinfinite_bd_goal = HOLogic.mk_Trueprop (mk_cinfinite bnf_bd_As);
blanchet@49990
   988
blanchet@50473
   989
    val set_bds_goal =
blanchet@49990
   990
      let
blanchet@49990
   991
        fun mk_goal set =
blanchet@49990
   992
          Logic.all x (HOLogic.mk_Trueprop (mk_ordLeq (mk_card_of (set $ x)) bnf_bd_As));
blanchet@49990
   993
      in
blanchet@49990
   994
        map mk_goal bnf_sets_As
blanchet@49990
   995
      end;
blanchet@49990
   996
traytel@56183
   997
    val relAsCs = mk_bnf_rel pred2RTsAsCs CA' CC';
traytel@56183
   998
    val relBsCs = mk_bnf_rel pred2RTsBsCs CB' CC';
traytel@56183
   999
    val rel_OO_lhs = Term.list_comb (relAsCs, map2 (curry mk_rel_compp) Rs Ss);
traytel@56183
  1000
    val rel_OO_rhs = mk_rel_compp (Term.list_comb (rel, Rs), Term.list_comb (relBsCs, Ss));
traytel@56183
  1001
    val le_rel_OO_goal =
traytel@56183
  1002
      fold_rev Logic.all (Rs @ Ss) (HOLogic.mk_Trueprop (mk_leq rel_OO_rhs rel_OO_lhs));
blanchet@49990
  1003
traytel@56183
  1004
    val rel_OO_Grp_goal = fold_rev Logic.all Rs (mk_Trueprop_eq (Term.list_comb (rel, Rs),
traytel@56183
  1005
      Term.list_comb (mk_OO_Grp Ds As' Bs', Rs)));
blanchet@50468
  1006
blanchet@54426
  1007
    val goals = zip_axioms map_id0_goal map_comp0_goal map_cong0_goal set_map0s_goal
traytel@56183
  1008
      card_order_bd_goal cinfinite_bd_goal set_bds_goal le_rel_OO_goal rel_OO_Grp_goal;
blanchet@49990
  1009
blanchet@49990
  1010
    fun mk_wit_goals (I, wit) =
blanchet@49990
  1011
      let
blanchet@49990
  1012
        val xs = map (nth bs) I;
blanchet@49990
  1013
        fun wit_goal i =
blanchet@49990
  1014
          let
blanchet@49990
  1015
            val z = nth zs i;
blanchet@49990
  1016
            val set_wit = nth bnf_sets_As i $ Term.list_comb (wit, xs);
blanchet@49990
  1017
            val concl = HOLogic.mk_Trueprop
blanchet@49990
  1018
              (if member (op =) I i then HOLogic.mk_eq (z, nth bs i)
blanchet@49990
  1019
              else @{term False});
blanchet@49990
  1020
          in
blanchet@49990
  1021
            fold_rev Logic.all (z :: xs)
blanchet@49990
  1022
              (Logic.mk_implies (HOLogic.mk_Trueprop (HOLogic.mk_mem (z, set_wit)), concl))
blanchet@49990
  1023
          end;
blanchet@49990
  1024
      in
blanchet@49990
  1025
        map wit_goal (0 upto live - 1)
blanchet@49990
  1026
      end;
blanchet@49990
  1027
traytel@56539
  1028
    fun triv_wit_tac ctxt = mk_trivial_wit_tac ctxt bnf_wit_defs;
blanchet@49990
  1029
traytel@55641
  1030
    val wit_goalss =
blanchet@56263
  1031
      (if null raw_wits then SOME triv_wit_tac else NONE, map mk_wit_goals bnf_wit_As);
traytel@55641
  1032
traytel@55641
  1033
    fun after_qed mk_wit_thms thms lthy =
blanchet@49990
  1034
      let
traytel@55641
  1035
        val (axioms, nontriv_wit_thms) = apfst (mk_axioms live) (chop (length goals) thms);
blanchet@49990
  1036
traytel@50124
  1037
        val bd_Card_order = #bd_card_order axioms RS @{thm conjunct2[OF card_order_on_Card_order]};
blanchet@49990
  1038
        val bd_Cinfinite = @{thm conjI} OF [#bd_cinfinite axioms, bd_Card_order];
blanchet@49990
  1039
        val bd_Cnotzero = bd_Cinfinite RS @{thm Cinfinite_Cnotzero};
blanchet@49990
  1040
blanchet@52903
  1041
        fun mk_collect_set_map () =
blanchet@49990
  1042
          let
traytel@56183
  1043
            val defT = mk_bnf_T Ts Calpha --> HOLogic.mk_setT T;
blanchet@49990
  1044
            val collect_map = HOLogic.mk_comp
blanchet@49990
  1045
              (mk_collect (map (mk_bnf_t Ts) bnf_sets) defT,
blanchet@49990
  1046
              Term.list_comb (mk_bnf_map As' Ts, hs));
blanchet@49990
  1047
            val image_collect = mk_collect
blanchet@49990
  1048
              (map2 (fn h => fn set => HOLogic.mk_comp (mk_image h, set)) hs bnf_sets_As)
blanchet@49990
  1049
              defT;
blanchet@49990
  1050
            (*collect {set1 ... setm} o map f1 ... fm = collect {f1` o set1 ... fm` o setm}*)
blanchet@50138
  1051
            val goal = fold_rev Logic.all hs (mk_Trueprop_eq (collect_map, image_collect));
blanchet@49990
  1052
          in
blanchet@54426
  1053
            Goal.prove_sorry lthy [] [] goal (K (mk_collect_set_map_tac (#set_map0 axioms)))
traytel@50124
  1054
            |> Thm.close_derivation
blanchet@49990
  1055
          end;
blanchet@49990
  1056
blanchet@52903
  1057
        val collect_set_map = Lazy.lazy mk_collect_set_map;
blanchet@49990
  1058
blanchet@49990
  1059
        fun mk_in_mono () =
blanchet@49990
  1060
          let
traytel@53030
  1061
            val prems_mono = map2 (HOLogic.mk_Trueprop oo mk_leq) As As_copy;
blanchet@50473
  1062
            val in_mono_goal =
blanchet@49990
  1063
              fold_rev Logic.all (As @ As_copy)
blanchet@49990
  1064
                (Logic.list_implies (prems_mono, HOLogic.mk_Trueprop
traytel@53030
  1065
                  (mk_leq (mk_in As bnf_sets_As CA') (mk_in As_copy bnf_sets_As CA'))));
blanchet@49990
  1066
          in
wenzelm@52688
  1067
            Goal.prove_sorry lthy [] [] in_mono_goal (K (mk_in_mono_tac live))
traytel@50124
  1068
            |> Thm.close_derivation
blanchet@49990
  1069
          end;
blanchet@49990
  1070
blanchet@50553
  1071
        val in_mono = Lazy.lazy mk_in_mono;
blanchet@49990
  1072
blanchet@49990
  1073
        fun mk_in_cong () =
blanchet@49990
  1074
          let
blanchet@52899
  1075
            val prems_cong = map2 (curry mk_Trueprop_eq) As As_copy;
blanchet@50473
  1076
            val in_cong_goal =
blanchet@49990
  1077
              fold_rev Logic.all (As @ As_copy)
blanchet@52899
  1078
                (Logic.list_implies (prems_cong,
blanchet@52899
  1079
                  mk_Trueprop_eq (mk_in As bnf_sets_As CA', mk_in As_copy bnf_sets_As CA')));
blanchet@49990
  1080
          in
wenzelm@52935
  1081
            Goal.prove_sorry lthy [] [] in_cong_goal
wenzelm@52935
  1082
              (K ((TRY o hyp_subst_tac lthy THEN' rtac refl) 1))
traytel@50124
  1083
            |> Thm.close_derivation
blanchet@49990
  1084
          end;
blanchet@49990
  1085
blanchet@50553
  1086
        val in_cong = Lazy.lazy mk_in_cong;
blanchet@49990
  1087
blanchet@54422
  1088
        val map_id = Lazy.lazy (fn () => mk_map_id (#map_id0 axioms));
blanchet@54425
  1089
        val map_comp = Lazy.lazy (fn () => mk_map_comp (#map_comp0 axioms));
blanchet@52899
  1090
blanchet@52899
  1091
        fun mk_map_cong () =
blanchet@52899
  1092
          let
blanchet@52899
  1093
            val prem0 = mk_Trueprop_eq (x, x_copy);
blanchet@52899
  1094
            val prems = map4 (mk_map_cong_prem x_copy) zs bnf_sets_As fs fs_copy;
blanchet@52899
  1095
            val eq = mk_Trueprop_eq (Term.list_comb (bnf_map_AsBs, fs) $ x,
blanchet@52899
  1096
              Term.list_comb (bnf_map_AsBs, fs_copy) $ x_copy);
blanchet@52899
  1097
            val goal = fold_rev Logic.all (x :: x_copy :: fs @ fs_copy)
blanchet@52899
  1098
              (Logic.list_implies (prem0 :: prems, eq));
blanchet@52899
  1099
          in
wenzelm@52935
  1100
            Goal.prove_sorry lthy [] [] goal (fn _ => mk_map_cong_tac lthy (#map_cong0 axioms))
blanchet@52899
  1101
            |> Thm.close_derivation
blanchet@52899
  1102
          end;
blanchet@52899
  1103
blanchet@52899
  1104
        val map_cong = Lazy.lazy mk_map_cong;
blanchet@49990
  1105
blanchet@54427
  1106
        val set_map = map (fn thm => Lazy.lazy (fn () => mk_set_map thm)) (#set_map0 axioms);
blanchet@49990
  1107
traytel@55641
  1108
        val wit_thms =
traytel@55641
  1109
          if null nontriv_wit_thms then mk_wit_thms (map Lazy.force set_map) else nontriv_wit_thms;
traytel@55641
  1110
traytel@53772
  1111
        fun mk_in_bd () =
traytel@53772
  1112
          let
traytel@53950
  1113
            val bdT = fst (dest_relT (fastype_of bnf_bd_As));
traytel@53950
  1114
            val bdTs = replicate live bdT;
traytel@56183
  1115
            val bd_bnfT = mk_bnf_T bdTs Calpha;
traytel@53950
  1116
            val surj_imp_ordLeq_inst = (if live = 0 then TrueI else
traytel@53950
  1117
              let
traytel@53950
  1118
                val ranTs = map (fn AT => mk_sumT (AT, HOLogic.unitT)) As';
traytel@53950
  1119
                val funTs = map (fn T => bdT --> T) ranTs;
traytel@56183
  1120
                val ran_bnfT = mk_bnf_T ranTs Calpha;
traytel@53950
  1121
                val (revTs, Ts) = `rev (bd_bnfT :: funTs);
traytel@53950
  1122
                val cTs = map (SOME o certifyT lthy) [ran_bnfT, Library.foldr1 HOLogic.mk_prodT Ts];
traytel@53950
  1123
                val tinst = fold (fn T => fn t => HOLogic.mk_split (Term.absdummy T t)) (tl revTs)
traytel@53950
  1124
                  (Term.absdummy (hd revTs) (Term.list_comb (mk_bnf_map bdTs ranTs,
traytel@53950
  1125
                    map Bound (live - 1 downto 0)) $ Bound live));
traytel@53950
  1126
                val cts = [NONE, SOME (certify lthy tinst)];
traytel@53950
  1127
              in
traytel@53950
  1128
                Drule.instantiate' cTs cts @{thm surj_imp_ordLeq}
traytel@53950
  1129
              end);
traytel@53772
  1130
            val bd = mk_cexp
traytel@53772
  1131
              (if live = 0 then ctwo
traytel@53772
  1132
                else mk_csum (Library.foldr1 (uncurry mk_csum) (map mk_card_of As)) ctwo)
traytel@53950
  1133
              (mk_csum bnf_bd_As (mk_card_of (HOLogic.mk_UNIV bd_bnfT)));
traytel@53772
  1134
            val in_bd_goal =
traytel@53772
  1135
              fold_rev Logic.all As
traytel@53772
  1136
                (HOLogic.mk_Trueprop (mk_ordLeq (mk_card_of (mk_in As bnf_sets_As CA')) bd));
traytel@53772
  1137
          in
traytel@53772
  1138
            Goal.prove_sorry lthy [] [] in_bd_goal
traytel@56539
  1139
              (fn {context = ctxt, prems = _} => mk_in_bd_tac ctxt live surj_imp_ordLeq_inst
blanchet@54425
  1140
                (Lazy.force map_comp) (Lazy.force map_id) (#map_cong0 axioms)
blanchet@54427
  1141
                (map Lazy.force set_map) (#set_bd axioms) (#bd_card_order axioms)
traytel@53772
  1142
                bd_Card_order bd_Cinfinite bd_Cnotzero)
traytel@53772
  1143
            |> Thm.close_derivation
traytel@53772
  1144
          end;
traytel@53772
  1145
traytel@53772
  1146
        val in_bd = Lazy.lazy mk_in_bd;
traytel@53772
  1147
traytel@54698
  1148
        val rel_OO_Grp = #rel_OO_Grp axioms;
traytel@54698
  1149
        val rel_OO_Grps = no_refl [rel_OO_Grp];
blanchet@50468
  1150
traytel@53030
  1151
        fun mk_rel_Grp () =
blanchet@49990
  1152
          let
traytel@53030
  1153
            val lhs = Term.list_comb (rel, map2 mk_Grp As fs);
traytel@53030
  1154
            val rhs = mk_Grp (mk_in As bnf_sets_As CA') (Term.list_comb (bnf_map_AsBs, fs));
blanchet@50138
  1155
            val goal = fold_rev Logic.all (As @ fs) (mk_Trueprop_eq (lhs, rhs));
blanchet@49990
  1156
          in
wenzelm@52688
  1157
            Goal.prove_sorry lthy [] [] goal
traytel@56539
  1158
              (fn {context = ctxt, prems = _} => mk_rel_Grp_tac ctxt rel_OO_Grps (#map_id0 axioms)
traytel@56539
  1159
                (#map_cong0 axioms) (Lazy.force map_id) (Lazy.force map_comp)
traytel@56539
  1160
                (map Lazy.force set_map))
traytel@50124
  1161
            |> Thm.close_derivation
blanchet@49990
  1162
          end;
blanchet@49990
  1163
traytel@53030
  1164
        val rel_Grp = Lazy.lazy mk_rel_Grp;
blanchet@49990
  1165
traytel@53030
  1166
        fun mk_rel_prems f = map2 (HOLogic.mk_Trueprop oo f) Rs Rs_copy
traytel@53030
  1167
        fun mk_rel_concl f = HOLogic.mk_Trueprop
traytel@53030
  1168
          (f (Term.list_comb (rel, Rs), Term.list_comb (rel, Rs_copy)));
blanchet@49990
  1169
traytel@53030
  1170
        fun mk_rel_mono () =
blanchet@49990
  1171
          let
traytel@53030
  1172
            val mono_prems = mk_rel_prems mk_leq;
traytel@53030
  1173
            val mono_concl = mk_rel_concl (uncurry mk_leq);
blanchet@49990
  1174
          in
wenzelm@52688
  1175
            Goal.prove_sorry lthy [] []
blanchet@49990
  1176
              (fold_rev Logic.all (Rs @ Rs_copy) (Logic.list_implies (mono_prems, mono_concl)))
traytel@53981
  1177
              (K (mk_rel_mono_tac rel_OO_Grps (Lazy.force in_mono)))
traytel@50124
  1178
            |> Thm.close_derivation
blanchet@49990
  1179
          end;
blanchet@49990
  1180
traytel@53030
  1181
        fun mk_rel_cong () =
blanchet@49990
  1182
          let
traytel@53030
  1183
            val cong_prems = mk_rel_prems (curry HOLogic.mk_eq);
traytel@53030
  1184
            val cong_concl = mk_rel_concl HOLogic.mk_eq;
blanchet@49990
  1185
          in
wenzelm@52688
  1186
            Goal.prove_sorry lthy [] []
blanchet@49990
  1187
              (fold_rev Logic.all (Rs @ Rs_copy) (Logic.list_implies (cong_prems, cong_concl)))
wenzelm@52935
  1188
              (fn _ => (TRY o hyp_subst_tac lthy THEN' rtac refl) 1)
traytel@50124
  1189
            |> Thm.close_derivation
blanchet@49990
  1190
          end;
blanchet@49990
  1191
traytel@53030
  1192
        val rel_mono = Lazy.lazy mk_rel_mono;
traytel@53030
  1193
        val rel_cong = Lazy.lazy mk_rel_cong;
blanchet@49990
  1194
traytel@53030
  1195
        fun mk_rel_eq () =
traytel@53856
  1196
          Goal.prove_sorry lthy [] []
traytel@53856
  1197
            (mk_Trueprop_eq (Term.list_comb (relAsAs, map HOLogic.eq_const As'),
traytel@53856
  1198
              HOLogic.eq_const CA'))
blanchet@54407
  1199
            (K (mk_rel_eq_tac live (Lazy.force rel_Grp) (Lazy.force rel_cong) (#map_id0 axioms)))
traytel@53856
  1200
          |> Thm.close_derivation;
blanchet@49990
  1201
traytel@53030
  1202
        val rel_eq = Lazy.lazy mk_rel_eq;
blanchet@49990
  1203
traytel@53030
  1204
        fun mk_rel_conversep () =
blanchet@49990
  1205
          let
traytel@53030
  1206
            val relBsAs = mk_bnf_rel pred2RT's CB' CA';
traytel@53030
  1207
            val lhs = Term.list_comb (relBsAs, map mk_conversep Rs);
traytel@53030
  1208
            val rhs = mk_conversep (Term.list_comb (rel, Rs));
traytel@53030
  1209
            val le_goal = fold_rev Logic.all Rs (HOLogic.mk_Trueprop (mk_leq lhs rhs));
wenzelm@52688
  1210
            val le_thm = Goal.prove_sorry lthy [] [] le_goal
traytel@56539
  1211
              (fn {context = ctxt, prems = _} => mk_rel_conversep_le_tac ctxt rel_OO_Grps
traytel@56539
  1212
                (Lazy.force rel_eq) (#map_cong0 axioms) (Lazy.force map_comp)
traytel@56539
  1213
                (map Lazy.force set_map))
traytel@50124
  1214
              |> Thm.close_derivation
blanchet@50138
  1215
            val goal = fold_rev Logic.all Rs (mk_Trueprop_eq (lhs, rhs));
blanchet@49990
  1216
          in
traytel@53030
  1217
            Goal.prove_sorry lthy [] [] goal
traytel@53030
  1218
              (K (mk_rel_conversep_tac le_thm (Lazy.force rel_mono)))
traytel@50124
  1219
            |> Thm.close_derivation
blanchet@49990
  1220
          end;
blanchet@49990
  1221
traytel@53030
  1222
        val rel_conversep = Lazy.lazy mk_rel_conversep;
blanchet@49990
  1223
traytel@53030
  1224
        fun mk_rel_OO () =
traytel@56183
  1225
          Goal.prove_sorry lthy [] []
traytel@56183
  1226
            (fold_rev Logic.all (Rs @ Ss) (HOLogic.mk_Trueprop (mk_leq rel_OO_lhs rel_OO_rhs)))
traytel@56539
  1227
            (fn {context = ctxt, prems = _} => mk_rel_OO_le_tac ctxt rel_OO_Grps (Lazy.force rel_eq)
traytel@56539
  1228
              (#map_cong0 axioms) (Lazy.force map_comp) (map Lazy.force set_map))
traytel@56183
  1229
          |> Thm.close_derivation
traytel@56183
  1230
          |> (fn thm => @{thm antisym} OF [thm, #le_rel_OO axioms]);
blanchet@49990
  1231
traytel@53030
  1232
        val rel_OO = Lazy.lazy mk_rel_OO;
blanchet@49990
  1233
traytel@54698
  1234
        fun mk_in_rel () = trans OF [rel_OO_Grp, @{thm OO_Grp_alt}] RS @{thm predicate2_eqD};
blanchet@49990
  1235
traytel@53030
  1236
        val in_rel = Lazy.lazy mk_in_rel;
blanchet@49990
  1237
blanchet@50552
  1238
        fun mk_rel_flip () =
blanchet@50552
  1239
          let
traytel@53030
  1240
            val rel_conversep_thm = Lazy.force rel_conversep;
traytel@53030
  1241
            val cts = map (SOME o certify lthy) Rs;
traytel@53030
  1242
            val rel_conversep_thm' = cterm_instantiate_pos cts rel_conversep_thm;
blanchet@50552
  1243
          in
traytel@53054
  1244
            unfold_thms lthy @{thms conversep_iff} (rel_conversep_thm' RS @{thm predicate2_eqD})
blanchet@50610
  1245
            |> singleton (Proof_Context.export names_lthy pre_names_lthy)
blanchet@50552
  1246
          end;
blanchet@50552
  1247
blanchet@50553
  1248
        val rel_flip = Lazy.lazy mk_rel_flip;
blanchet@50552
  1249
traytel@53053
  1250
        fun mk_rel_mono_strong () =
traytel@53053
  1251
          let
traytel@53053
  1252
            fun mk_prem setA setB R S a b =
traytel@53053
  1253
              HOLogic.mk_Trueprop
traytel@53053
  1254
                (mk_Ball (setA $ x) (Term.absfree (dest_Free a)
traytel@53053
  1255
                  (mk_Ball (setB $ y) (Term.absfree (dest_Free b)
traytel@53053
  1256
                    (HOLogic.mk_imp (R $ a $ b, S $ a $ b))))));
traytel@53053
  1257
            val prems = HOLogic.mk_Trueprop (Term.list_comb (rel, Rs) $ x $ y) :: 
traytel@53053
  1258
              map6 mk_prem bnf_sets_As bnf_sets_Bs Rs Rs_copy zs ys;
traytel@53053
  1259
            val concl = HOLogic.mk_Trueprop (Term.list_comb (rel, Rs_copy) $ x $ y);
traytel@53053
  1260
          in
traytel@53053
  1261
            Goal.prove_sorry lthy [] []
traytel@53053
  1262
              (fold_rev Logic.all (x :: y :: Rs @ Rs_copy) (Logic.list_implies (prems, concl)))
traytel@56539
  1263
              (fn {context = ctxt, prems = _} => mk_rel_mono_strong_tac ctxt (Lazy.force in_rel)
traytel@56539
  1264
                (map Lazy.force set_map))
traytel@53053
  1265
            |> Thm.close_derivation
traytel@53053
  1266
          end;
traytel@53053
  1267
traytel@53053
  1268
        val rel_mono_strong = Lazy.lazy mk_rel_mono_strong;
traytel@53053
  1269
traytel@53856
  1270
        fun mk_map_transfer () =
traytel@53856
  1271
          let
blanchet@57287
  1272
            val rels = map2 mk_rel_fun transfer_domRs transfer_ranRs;
blanchet@57287
  1273
            val rel = mk_rel_fun
traytel@53862
  1274
              (Term.list_comb (mk_bnf_rel transfer_domRTs CA' CB1, transfer_domRs))
traytel@53862
  1275
              (Term.list_comb (mk_bnf_rel transfer_ranRTs CB' CB2, transfer_ranRs));
traytel@53862
  1276
            val concl = HOLogic.mk_Trueprop
blanchet@57287
  1277
              (fold_rev mk_rel_fun rels rel $ bnf_map_AsBs $ mk_bnf_map B1Ts B2Ts);
traytel@53856
  1278
          in
traytel@53856
  1279
            Goal.prove_sorry lthy [] []
traytel@53862
  1280
              (fold_rev Logic.all (transfer_domRs @ transfer_ranRs) concl)
traytel@56539
  1281
              (fn {context = ctxt, prems = _} => mk_map_transfer_tac ctxt (Lazy.force rel_mono)
traytel@56539
  1282
                (Lazy.force in_rel) (map Lazy.force set_map) (#map_cong0 axioms)
traytel@56539
  1283
                (Lazy.force map_comp))
traytel@53856
  1284
            |> Thm.close_derivation
traytel@53856
  1285
          end;
traytel@53856
  1286
traytel@53856
  1287
        val map_transfer = Lazy.lazy mk_map_transfer;
traytel@53856
  1288
traytel@53030
  1289
        val defs = mk_defs bnf_map_def bnf_set_defs bnf_rel_def;
blanchet@49990
  1290
traytel@53772
  1291
        val facts = mk_facts bd_Card_order bd_Cinfinite bd_Cnotzero collect_set_map in_bd in_cong
traytel@56183
  1292
          in_mono in_rel map_comp map_cong map_id map_transfer rel_eq rel_flip set_map
traytel@53053
  1293
          rel_cong rel_mono rel_mono_strong rel_Grp rel_conversep rel_OO;
blanchet@49990
  1294
blanchet@49990
  1295
        val wits = map2 mk_witness bnf_wits wit_thms;
blanchet@49990
  1296
blanchet@50522
  1297
        val bnf_rel =
blanchet@50522
  1298
          Term.subst_atomic_types ((Ds ~~ deads) @ (As' ~~ alphas) @ (Bs' ~~ betas)) rel;
blanchet@49990
  1299
traytel@56183
  1300
        val bnf = mk_bnf bnf_b Calpha live alphas betas dead deads bnf_map bnf_sets bnf_bd axioms
traytel@56183
  1301
          defs facts wits bnf_rel;
blanchet@49990
  1302
      in
blanchet@54402
  1303
        (bnf, lthy |> note_bnf_thms fact_policy qualify bnf_b bnf)
blanchet@49990
  1304
      end;
blanchet@50474
  1305
blanchet@50474
  1306
    val one_step_defs =
traytel@53030
  1307
      no_reflexive (bnf_map_def :: bnf_bd_def :: bnf_set_defs @ bnf_wit_defs @ [bnf_rel_def]);
blanchet@49990
  1308
  in
blanchet@50474
  1309
    (key, goals, wit_goalss, after_qed, lthy, one_step_defs)
blanchet@49990
  1310
  end;
blanchet@49990
  1311
blanchet@57688
  1312
structure BNF_Interpretation = Interpretation
blanchet@57688
  1313
(
blanchet@57688
  1314
  type T = bnf;
blanchet@57688
  1315
  val eq: T * T -> bool = op = o pairself T_of_bnf;
blanchet@57688
  1316
);
blanchet@57688
  1317
blanchet@57718
  1318
fun with_repaired_path f bnf thy =
blanchet@57718
  1319
  let
blanchet@57718
  1320
    val qualifiers =
blanchet@57718
  1321
      (case Binding.dest (name_of_bnf bnf) of
blanchet@57718
  1322
        (* arbitrarily use "Fun" as prefix for "fun"*)
blanchet@57718
  1323
        (_, [], @{type_name fun}) => [(Context.theory_name @{theory Fun}, false)]
blanchet@57718
  1324
      | (_, qs, _) => qs)
blanchet@57718
  1325
  in
blanchet@57718
  1326
    thy
blanchet@57718
  1327
    |> Sign.root_path
blanchet@57718
  1328
    |> fold (uncurry (fn true => Sign.mandatory_path | false => Sign.add_path) o swap) qualifiers
kuncar@57864
  1329
    |> (fn thy => f (morph_bnf (Morphism.transfer_morphism thy) bnf) thy)
blanchet@57718
  1330
    |> Sign.restore_naming thy
blanchet@57718
  1331
  end;
blanchet@57718
  1332
blanchet@57718
  1333
val bnf_interpretation = BNF_Interpretation.interpretation o with_repaired_path;
blanchet@57688
  1334
blanchet@57688
  1335
fun register_bnf key bnf =
blanchet@57688
  1336
  Local_Theory.declaration {syntax = false, pervasive = true}
blanchet@57688
  1337
    (fn phi => Data.map (Symtab.update (key, morph_bnf phi bnf)))
blanchet@57688
  1338
  #> Local_Theory.background_theory (BNF_Interpretation.data bnf);
traytel@50449
  1339
traytel@57358
  1340
fun bnf_def const_policy fact_policy internal qualify tacs wit_tac Ds map_b rel_b set_bs =
traytel@55641
  1341
  (fn (_, goals, (triv_tac_opt, wit_goalss), after_qed, lthy, one_step_defs) =>
blanchet@49990
  1342
  let
traytel@56539
  1343
    fun mk_wits_tac ctxt set_maps =
traytel@56539
  1344
      TRYALL Goal.conjunction_tac THEN
traytel@55641
  1345
      (case triv_tac_opt of
traytel@56539
  1346
        SOME tac => tac ctxt set_maps
traytel@56539
  1347
      | NONE => unfold_thms_tac ctxt one_step_defs THEN wit_tac ctxt);
traytel@55641
  1348
    val wit_goals = map Logic.mk_conjunction_balanced wit_goalss;
traytel@55641
  1349
    fun mk_wit_thms set_maps =
traytel@56539
  1350
      Goal.prove_sorry lthy [] [] (Logic.mk_conjunction_balanced wit_goals)
traytel@56539
  1351
        (fn {context = ctxt, prems = _} => mk_wits_tac ctxt set_maps)
traytel@55641
  1352
        |> Conjunction.elim_balanced (length wit_goals)
traytel@55641
  1353
        |> map2 (Conjunction.elim_balanced o length) wit_goalss
traytel@55641
  1354
        |> map (map (Thm.close_derivation o Thm.forall_elim_vars 0));
blanchet@49990
  1355
  in
wenzelm@52688
  1356
    map2 (Thm.close_derivation oo Goal.prove_sorry lthy [] [])
traytel@56539
  1357
      goals (map (fn tac => fn {context = ctxt, prems = _} =>
traytel@56539
  1358
        unfold_thms_tac ctxt one_step_defs THEN tac ctxt) tacs)
traytel@55641
  1359
    |> (fn thms => after_qed mk_wit_thms (map single thms) lthy)
traytel@57358
  1360
  end) oo prepare_def const_policy fact_policy internal qualify (K I) (K I) Ds map_b rel_b set_bs;
blanchet@49990
  1361
traytel@55641
  1362
val bnf_cmd = (fn (key, goals, (triv_tac_opt, wit_goalss), after_qed, lthy, defs) =>
traytel@55641
  1363
  let
traytel@55641
  1364
    val wit_goals = map Logic.mk_conjunction_balanced wit_goalss;
traytel@55641
  1365
    fun mk_triv_wit_thms tac set_maps =
traytel@55641
  1366
      Goal.prove_sorry lthy [] [] (Logic.mk_conjunction_balanced wit_goals)
traytel@56539
  1367
        (fn {context = ctxt, prems = _} => TRYALL Goal.conjunction_tac THEN tac ctxt set_maps)
traytel@55641
  1368
        |> Conjunction.elim_balanced (length wit_goals)
traytel@55641
  1369
        |> map2 (Conjunction.elim_balanced o length) wit_goalss
traytel@55641
  1370
        |> map (map (Thm.close_derivation o Thm.forall_elim_vars 0));
traytel@55641
  1371
    val (mk_wit_thms, nontriv_wit_goals) = 
traytel@55641
  1372
      (case triv_tac_opt of
traytel@55641
  1373
        NONE => (fn _ => [], map (map (rpair [])) wit_goalss)
traytel@55641
  1374
      | SOME tac => (mk_triv_wit_thms tac, []));
traytel@55641
  1375
  in
traytel@55641
  1376
    Proof.unfolding ([[(defs, [])]])
blanchet@57688
  1377
      (Proof.theorem NONE (uncurry (register_bnf key) oo after_qed mk_wit_thms)
traytel@55641
  1378
        (map (single o rpair []) goals @ nontriv_wit_goals) lthy)
traytel@57358
  1379
  end) oo prepare_def Do_Inline (user_policy Note_Some) false I Syntax.read_typ Syntax.read_term
traytel@57358
  1380
    NONE Binding.empty Binding.empty [];
blanchet@49990
  1381
blanchet@49990
  1382
fun print_bnfs ctxt =
blanchet@49990
  1383
  let
blanchet@49990
  1384
    fun pretty_set sets i = Pretty.block
blanchet@49990
  1385
      [Pretty.str (mk_setN (i + 1) ^ ":"), Pretty.brk 1,
blanchet@49990
  1386
          Pretty.quote (Syntax.pretty_term ctxt (nth sets i))];
blanchet@49990
  1387
blanchet@49990
  1388
    fun pretty_bnf (key, BNF {T = T, map = map, sets = sets, bd = bd,
blanchet@49990
  1389
      live = live, lives = lives, dead = dead, deads = deads, ...}) =
blanchet@49990
  1390
      Pretty.big_list
blanchet@49990
  1391
        (Pretty.string_of (Pretty.block [Pretty.str key, Pretty.str ":", Pretty.brk 1,
blanchet@49990
  1392
          Pretty.quote (Syntax.pretty_typ ctxt T)]))
blanchet@49990
  1393
        ([Pretty.block [Pretty.str "live:", Pretty.brk 1, Pretty.str (string_of_int live),
blanchet@49990
  1394
            Pretty.brk 3, Pretty.list "[" "]" (List.map (Syntax.pretty_typ ctxt) lives)],
blanchet@49990
  1395
          Pretty.block [Pretty.str "dead:", Pretty.brk 1, Pretty.str (string_of_int dead),
blanchet@49990
  1396
            Pretty.brk 3, Pretty.list "[" "]" (List.map (Syntax.pretty_typ ctxt) deads)],
blanchet@49990
  1397
          Pretty.block [Pretty.str (mapN ^ ":"), Pretty.brk 1,
blanchet@49990
  1398
            Pretty.quote (Syntax.pretty_term ctxt map)]] @
blanchet@49990
  1399
          List.map (pretty_set sets) (0 upto length sets - 1) @
blanchet@49990
  1400
          [Pretty.block [Pretty.str (bdN ^ ":"), Pretty.brk 1,
blanchet@49990
  1401
            Pretty.quote (Syntax.pretty_term ctxt bd)]]);
blanchet@49990
  1402
  in
blanchet@49990
  1403
    Pretty.big_list "BNFs:" (map pretty_bnf (Symtab.dest (Data.get (Context.Proof ctxt))))
blanchet@49990
  1404
    |> Pretty.writeln
blanchet@49990
  1405
  end;
blanchet@49990
  1406
blanchet@49990
  1407
val _ =
blanchet@52973
  1408
  Outer_Syntax.improper_command @{command_spec "print_bnfs"}
blanchet@54426
  1409
    "print all bounded natural functors"
blanchet@49990
  1410
    (Scan.succeed (Toplevel.keep (print_bnfs o Toplevel.context_of)));
blanchet@49990
  1411
blanchet@49990
  1412
val _ =
blanchet@52973
  1413
  Outer_Syntax.local_theory_to_proof @{command_spec "bnf"}
blanchet@54426
  1414
    "register a type as a bounded natural functor"
traytel@55794
  1415
    (parse_opt_binding_colon -- Parse.typ --|
traytel@55794
  1416
       (Parse.reserved "map" -- @{keyword ":"}) -- Parse.term --
traytel@55794
  1417
       (Scan.option ((Parse.reserved "sets" -- @{keyword ":"}) |--
traytel@55794
  1418
         Scan.repeat1 (Scan.unless (Parse.reserved "bd") Parse.term)) >> the_default []) --|
traytel@55794
  1419
       (Parse.reserved "bd" -- @{keyword ":"}) -- Parse.term --
traytel@55794
  1420
       (Scan.option ((Parse.reserved "wits" -- @{keyword ":"}) |--
traytel@55794
  1421
         Scan.repeat1 (Scan.unless (Parse.reserved "rel") Parse.term)) >> the_default []) --
traytel@55794
  1422
       Scan.option ((Parse.reserved "rel" -- @{keyword ":"}) |-- Parse.term)
blanchet@52973
  1423
       >> bnf_cmd);
blanchet@49990
  1424
kuncar@57864
  1425
val _ = Context.>> (Context.map_theory BNF_Interpretation.init);
kuncar@57864
  1426
blanchet@49990
  1427
end;