src/HOL/Codatatype/Tools/bnf_def.ML
author blanchet
Thu, 20 Sep 2012 02:42:48 +0200
changeset 50475 4dd451a075ce
parent 50474 3f8e2b5249ec
child 50476 de07eecb2664
permissions -rw-r--r--
fixed infinite loop with trivial rel_O_Gr + tuning
blanchet@49990
     1
(*  Title:      HOL/Codatatype/Tools/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@49990
    11
  type BNF
blanchet@49990
    12
  type nonemptiness_witness = {I: int list, wit: term, prop: thm list}
blanchet@49990
    13
blanchet@49990
    14
  val bnf_of: Proof.context -> string -> BNF option
traytel@50449
    15
  val register_bnf: string -> (BNF * local_theory) -> (BNF * local_theory)
traytel@50449
    16
blanchet@49990
    17
  val name_of_bnf: BNF -> binding
blanchet@49990
    18
  val T_of_bnf: BNF -> typ
blanchet@49990
    19
  val live_of_bnf: BNF -> int
blanchet@49990
    20
  val lives_of_bnf: BNF -> typ list
blanchet@49990
    21
  val dead_of_bnf: BNF -> int
blanchet@49990
    22
  val deads_of_bnf: BNF -> typ list
blanchet@49990
    23
  val nwits_of_bnf: BNF -> int
blanchet@49990
    24
blanchet@49990
    25
  val mapN: string
blanchet@49990
    26
  val setN: string
blanchet@49990
    27
  val relN: string
blanchet@49990
    28
  val predN: string
blanchet@49990
    29
  val mk_setN: int -> string
blanchet@49990
    30
blanchet@50229
    31
  val map_of_bnf: BNF -> term
blanchet@50229
    32
blanchet@49990
    33
  val mk_T_of_bnf: typ list -> typ list -> BNF -> typ
blanchet@49990
    34
  val mk_bd_of_bnf: typ list -> typ list -> BNF -> term
blanchet@49990
    35
  val mk_map_of_bnf: typ list -> typ list -> typ list -> BNF -> term
blanchet@49990
    36
  val mk_pred_of_bnf: typ list -> typ list -> typ list -> BNF -> term
blanchet@49990
    37
  val mk_rel_of_bnf: typ list -> typ list -> typ list -> BNF -> term
blanchet@49990
    38
  val mk_sets_of_bnf: typ list list -> typ list list -> BNF -> term list
blanchet@49990
    39
  val mk_wits_of_bnf: typ list list -> typ list list -> BNF -> (int list * term) list
blanchet@49990
    40
blanchet@49990
    41
  val bd_Card_order_of_bnf: BNF -> thm
blanchet@49990
    42
  val bd_Cinfinite_of_bnf: BNF -> thm
blanchet@49990
    43
  val bd_Cnotzero_of_bnf: BNF -> thm
blanchet@49990
    44
  val bd_card_order_of_bnf: BNF -> thm
blanchet@49990
    45
  val bd_cinfinite_of_bnf: BNF -> thm
blanchet@49990
    46
  val collect_set_natural_of_bnf: BNF -> thm
blanchet@49990
    47
  val in_bd_of_bnf: BNF -> thm
blanchet@49990
    48
  val in_cong_of_bnf: BNF -> thm
blanchet@49990
    49
  val in_mono_of_bnf: BNF -> thm
blanchet@49990
    50
  val in_rel_of_bnf: BNF -> thm
blanchet@49990
    51
  val map_comp'_of_bnf: BNF -> thm
blanchet@49990
    52
  val map_comp_of_bnf: BNF -> thm
blanchet@49990
    53
  val map_cong_of_bnf: BNF -> thm
blanchet@49990
    54
  val map_def_of_bnf: BNF -> thm
blanchet@49990
    55
  val map_id'_of_bnf: BNF -> thm
blanchet@49990
    56
  val map_id_of_bnf: BNF -> thm
blanchet@49990
    57
  val map_wppull_of_bnf: BNF -> thm
blanchet@49990
    58
  val map_wpull_of_bnf: BNF -> thm
blanchet@49990
    59
  val pred_def_of_bnf: BNF -> thm
blanchet@50467
    60
  val rel_def_of_bnf: BNF -> thm
blanchet@49990
    61
  val rel_Gr_of_bnf: BNF -> thm
blanchet@49990
    62
  val rel_Id_of_bnf: BNF -> thm
blanchet@49990
    63
  val rel_O_of_bnf: BNF -> thm
blanchet@50467
    64
  val rel_O_Gr_of_bnf: BNF -> thm
blanchet@49990
    65
  val rel_cong_of_bnf: BNF -> thm
blanchet@49990
    66
  val rel_converse_of_bnf: BNF -> thm
blanchet@49990
    67
  val rel_mono_of_bnf: BNF -> thm
blanchet@49990
    68
  val set_bd_of_bnf: BNF -> thm list
blanchet@49990
    69
  val set_defs_of_bnf: BNF -> thm list
blanchet@49990
    70
  val set_natural'_of_bnf: BNF -> thm list
blanchet@49990
    71
  val set_natural_of_bnf: BNF -> thm list
blanchet@49990
    72
  val sets_of_bnf: BNF -> term list
blanchet@49990
    73
  val wit_thms_of_bnf: BNF -> thm list
blanchet@49990
    74
  val wit_thmss_of_bnf: BNF -> thm list list
blanchet@49990
    75
blanchet@49990
    76
  val mk_witness: int list * term -> thm list -> nonemptiness_witness
traytel@50118
    77
  val minimize_wits: (''a list * 'b) list -> (''a list * 'b) list
blanchet@49990
    78
  val wits_of_bnf: BNF -> nonemptiness_witness list
blanchet@49990
    79
blanchet@50475
    80
  val no_reflexive: thm list -> thm list
blanchet@50475
    81
  val zip_axioms: 'a -> 'a -> 'a -> 'a list -> 'a -> 'a -> 'a list -> 'a -> 'a -> 'a -> 'a list
blanchet@50471
    82
blanchet@49990
    83
  datatype const_policy = Dont_Inline | Hardly_Inline | Smart_Inline | Do_Inline
blanchet@49990
    84
  datatype fact_policy =
blanchet@49990
    85
    Derive_Some_Facts | Derive_All_Facts | Derive_All_Facts_Note_Most | Note_All_Facts_and_Axioms
blanchet@49990
    86
  val bnf_note_all: bool Config.T
traytel@50450
    87
  val user_policy: fact_policy -> Proof.context -> fact_policy
blanchet@49990
    88
blanchet@49990
    89
  val print_bnfs: Proof.context -> unit
blanchet@50033
    90
  val bnf_def: const_policy -> (Proof.context -> fact_policy) -> (binding -> binding) ->
blanchet@49990
    91
    ({prems: thm list, context: Proof.context} -> tactic) list ->
blanchet@49990
    92
    ({prems: thm list, context: Proof.context} -> tactic) -> typ list option ->
blanchet@50474
    93
    ((((binding * term) * term list) * term) * term list) * term option -> local_theory ->
blanchet@49990
    94
    BNF * local_theory
blanchet@49990
    95
end;
blanchet@49990
    96
blanchet@49990
    97
structure BNF_Def : BNF_DEF =
blanchet@49990
    98
struct
blanchet@49990
    99
blanchet@49990
   100
open BNF_Util
blanchet@50299
   101
open BNF_Def_Tactics
blanchet@49990
   102
blanchet@49990
   103
type axioms = {
blanchet@49990
   104
  map_id: thm,
blanchet@49990
   105
  map_comp: thm,
blanchet@49990
   106
  map_cong: thm,
blanchet@49990
   107
  set_natural: thm list,
blanchet@49990
   108
  bd_card_order: thm,
blanchet@49990
   109
  bd_cinfinite: thm,
blanchet@49990
   110
  set_bd: thm list,
blanchet@49990
   111
  in_bd: thm,
blanchet@50468
   112
  map_wpull: thm,
blanchet@50468
   113
  rel_O_Gr: thm
blanchet@49990
   114
};
blanchet@49990
   115
blanchet@50468
   116
fun mk_axioms' (((((((((id, comp), cong), nat), c_o), cinf), set_bd), in_bd), wpull), rel) =
blanchet@49990
   117
  {map_id = id, map_comp = comp, map_cong = cong, set_natural = nat, bd_card_order = c_o,
blanchet@50468
   118
   bd_cinfinite = cinf, set_bd = set_bd, in_bd = in_bd, map_wpull = wpull, rel_O_Gr = rel};
blanchet@49990
   119
blanchet@49990
   120
fun dest_cons [] = raise Empty
blanchet@49990
   121
  | dest_cons (x :: xs) = (x, xs);
blanchet@49990
   122
blanchet@49990
   123
fun mk_axioms n thms = thms
blanchet@49990
   124
  |> map the_single
blanchet@49990
   125
  |> dest_cons
blanchet@49990
   126
  ||>> dest_cons
blanchet@49990
   127
  ||>> dest_cons
blanchet@49990
   128
  ||>> chop n
blanchet@49990
   129
  ||>> dest_cons
blanchet@49990
   130
  ||>> dest_cons
blanchet@49990
   131
  ||>> chop n
blanchet@49990
   132
  ||>> dest_cons
blanchet@50468
   133
  ||>> dest_cons
blanchet@49990
   134
  ||> the_single
blanchet@49990
   135
  |> mk_axioms';
blanchet@49990
   136
blanchet@50475
   137
fun zip_axioms mid mcomp mcong snat bdco bdinf sbd inbd wpull rel =
blanchet@50475
   138
  [mid, mcomp, mcong] @ snat @ [bdco, bdinf] @ sbd @ [inbd, wpull, rel];
blanchet@50475
   139
blanchet@50468
   140
fun dest_axioms {map_id, map_comp, map_cong, set_natural, bd_card_order, bd_cinfinite, set_bd,
blanchet@50468
   141
  in_bd, map_wpull, rel_O_Gr} =
blanchet@50475
   142
  zip_axioms map_id map_comp map_cong set_natural bd_card_order bd_cinfinite set_bd in_bd map_wpull
blanchet@50475
   143
  rel_O_Gr;
blanchet@49990
   144
blanchet@49990
   145
fun map_axioms f
blanchet@49990
   146
  {map_id = map_id, map_comp = map_comp, map_cong = map_cong, set_natural = set_natural,
blanchet@50468
   147
   bd_card_order = bd_card_order, bd_cinfinite = bd_cinfinite, set_bd = set_bd, in_bd = in_bd,
blanchet@50468
   148
   map_wpull = map_wpull, rel_O_Gr} =
blanchet@49990
   149
  {map_id = f map_id,
blanchet@49990
   150
   map_comp = f map_comp,
blanchet@49990
   151
   map_cong = f map_cong,
blanchet@49990
   152
   set_natural = map f set_natural,
blanchet@49990
   153
   bd_card_order = f bd_card_order,
blanchet@49990
   154
   bd_cinfinite = f bd_cinfinite,
blanchet@49990
   155
   set_bd = map f set_bd,
blanchet@49990
   156
   in_bd = f in_bd,
blanchet@50468
   157
   map_wpull = f map_wpull,
blanchet@50468
   158
   rel_O_Gr = f rel_O_Gr};
blanchet@49990
   159
blanchet@49990
   160
val morph_axioms = map_axioms o Morphism.thm;
blanchet@49990
   161
blanchet@49990
   162
type defs = {
blanchet@49990
   163
  map_def: thm,
blanchet@49990
   164
  set_defs: thm list,
blanchet@49990
   165
  rel_def: thm,
blanchet@49990
   166
  pred_def: thm
blanchet@49990
   167
}
blanchet@49990
   168
blanchet@49990
   169
fun mk_defs map sets rel pred = {map_def = map, set_defs = sets, rel_def = rel, pred_def = pred};
blanchet@49990
   170
blanchet@49990
   171
fun map_defs f {map_def = map, set_defs = sets, rel_def = rel, pred_def = pred} =
blanchet@49990
   172
  {map_def = f map, set_defs = List.map f sets, rel_def = f rel, pred_def = f pred};
blanchet@49990
   173
blanchet@49990
   174
val morph_defs = map_defs o Morphism.thm;
blanchet@49990
   175
blanchet@49990
   176
type facts = {
blanchet@49990
   177
  bd_Card_order: thm,
blanchet@49990
   178
  bd_Cinfinite: thm,
blanchet@49990
   179
  bd_Cnotzero: thm,
blanchet@49990
   180
  collect_set_natural: thm lazy,
blanchet@49990
   181
  in_cong: thm lazy,
blanchet@49990
   182
  in_mono: thm lazy,
blanchet@49990
   183
  in_rel: thm lazy,
blanchet@49990
   184
  map_comp': thm lazy,
blanchet@49990
   185
  map_id': thm lazy,
blanchet@49990
   186
  map_wppull: thm lazy,
blanchet@49990
   187
  rel_cong: thm lazy,
blanchet@49990
   188
  rel_mono: thm lazy,
blanchet@49990
   189
  rel_Id: thm lazy,
blanchet@49990
   190
  rel_Gr: thm lazy,
blanchet@49990
   191
  rel_converse: thm lazy,
blanchet@49990
   192
  rel_O: thm lazy,
blanchet@49990
   193
  set_natural': thm lazy list
blanchet@49990
   194
};
blanchet@49990
   195
blanchet@49990
   196
fun mk_facts bd_Card_order bd_Cinfinite bd_Cnotzero
blanchet@49990
   197
    collect_set_natural in_cong in_mono in_rel map_comp' map_id' map_wppull
blanchet@50468
   198
    rel_cong rel_mono rel_Id rel_Gr rel_converse rel_O set_natural' = {
blanchet@49990
   199
  bd_Card_order = bd_Card_order,
blanchet@49990
   200
  bd_Cinfinite = bd_Cinfinite,
blanchet@49990
   201
  bd_Cnotzero = bd_Cnotzero,
blanchet@49990
   202
  collect_set_natural = collect_set_natural,
blanchet@49990
   203
  in_cong = in_cong,
blanchet@49990
   204
  in_mono = in_mono,
blanchet@49990
   205
  in_rel = in_rel,
blanchet@49990
   206
  map_comp' = map_comp',
blanchet@49990
   207
  map_id' = map_id',
blanchet@49990
   208
  map_wppull = map_wppull,
blanchet@49990
   209
  rel_cong = rel_cong,
blanchet@49990
   210
  rel_mono = rel_mono,
blanchet@49990
   211
  rel_Id = rel_Id,
blanchet@49990
   212
  rel_Gr = rel_Gr,
blanchet@49990
   213
  rel_converse = rel_converse,
blanchet@49990
   214
  rel_O = rel_O,
blanchet@49990
   215
  set_natural' = set_natural'};
blanchet@49990
   216
blanchet@49990
   217
fun map_facts f {
blanchet@49990
   218
  bd_Card_order,
blanchet@49990
   219
  bd_Cinfinite,
blanchet@49990
   220
  bd_Cnotzero,
blanchet@49990
   221
  collect_set_natural,
blanchet@49990
   222
  in_cong,
blanchet@49990
   223
  in_mono,
blanchet@49990
   224
  in_rel,
blanchet@49990
   225
  map_comp',
blanchet@49990
   226
  map_id',
blanchet@49990
   227
  map_wppull,
blanchet@49990
   228
  rel_cong,
blanchet@49990
   229
  rel_mono,
blanchet@49990
   230
  rel_Id,
blanchet@49990
   231
  rel_Gr,
blanchet@49990
   232
  rel_converse,
blanchet@49990
   233
  rel_O,
blanchet@49990
   234
  set_natural'} =
blanchet@49990
   235
  {bd_Card_order = f bd_Card_order,
blanchet@49990
   236
    bd_Cinfinite = f bd_Cinfinite,
blanchet@49990
   237
    bd_Cnotzero = f bd_Cnotzero,
blanchet@49990
   238
    collect_set_natural = Lazy.map f collect_set_natural,
blanchet@49990
   239
    in_cong = Lazy.map f in_cong,
blanchet@49990
   240
    in_mono = Lazy.map f in_mono,
blanchet@49990
   241
    in_rel = Lazy.map f in_rel,
blanchet@49990
   242
    map_comp' = Lazy.map f map_comp',
blanchet@49990
   243
    map_id' = Lazy.map f map_id',
blanchet@49990
   244
    map_wppull = Lazy.map f map_wppull,
blanchet@49990
   245
    rel_cong = Lazy.map f rel_cong,
blanchet@49990
   246
    rel_mono = Lazy.map f rel_mono,
blanchet@49990
   247
    rel_Id = Lazy.map f rel_Id,
blanchet@49990
   248
    rel_Gr = Lazy.map f rel_Gr,
blanchet@49990
   249
    rel_converse = Lazy.map f rel_converse,
blanchet@49990
   250
    rel_O = Lazy.map f rel_O,
blanchet@49990
   251
    set_natural' = map (Lazy.map f) set_natural'};
blanchet@49990
   252
blanchet@49990
   253
val morph_facts = map_facts o Morphism.thm;
blanchet@49990
   254
blanchet@49990
   255
type nonemptiness_witness = {
blanchet@49990
   256
  I: int list,
blanchet@49990
   257
  wit: term,
blanchet@49990
   258
  prop: thm list
blanchet@49990
   259
};
blanchet@49990
   260
blanchet@49990
   261
fun mk_witness (I, wit) prop = {I = I, wit = wit, prop = prop};
blanchet@49990
   262
fun map_witness f g {I, wit, prop} = {I = I, wit = f wit, prop = map g prop};
blanchet@49990
   263
fun morph_witness phi = map_witness (Morphism.term phi) (Morphism.thm phi);
blanchet@49990
   264
blanchet@49990
   265
datatype BNF = BNF of {
blanchet@49990
   266
  name: binding,
blanchet@49990
   267
  T: typ,
blanchet@49990
   268
  live: int,
blanchet@49990
   269
  lives: typ list, (*source type variables of map, only for composition*)
blanchet@49990
   270
  lives': typ list, (*target type variables of map, only for composition*)
blanchet@49990
   271
  dead: int,
blanchet@49990
   272
  deads: typ list, (*only for composition*)
blanchet@49990
   273
  map: term,
blanchet@49990
   274
  sets: term list,
blanchet@49990
   275
  bd: term,
blanchet@49990
   276
  axioms: axioms,
blanchet@49990
   277
  defs: defs,
blanchet@49990
   278
  facts: facts,
blanchet@49990
   279
  nwits: int,
blanchet@49990
   280
  wits: nonemptiness_witness list,
blanchet@49990
   281
  rel: term,
blanchet@49990
   282
  pred: term
blanchet@49990
   283
};
blanchet@49990
   284
blanchet@49990
   285
(* getters *)
blanchet@49990
   286
blanchet@49990
   287
fun rep_bnf (BNF bnf) = bnf;
blanchet@49990
   288
val name_of_bnf = #name o rep_bnf;
blanchet@49990
   289
val T_of_bnf = #T o rep_bnf;
blanchet@49990
   290
fun mk_T_of_bnf Ds Ts bnf =
blanchet@49990
   291
  let val bnf_rep = rep_bnf bnf
blanchet@49990
   292
  in Term.typ_subst_atomic ((#deads bnf_rep ~~ Ds) @ (#lives bnf_rep ~~ Ts)) (#T bnf_rep) end;
blanchet@49990
   293
val live_of_bnf = #live o rep_bnf;
blanchet@49990
   294
val lives_of_bnf = #lives o rep_bnf;
blanchet@49990
   295
val dead_of_bnf = #dead o rep_bnf;
blanchet@49990
   296
val deads_of_bnf = #deads o rep_bnf;
blanchet@49990
   297
val axioms_of_bnf = #axioms o rep_bnf;
blanchet@49990
   298
val facts_of_bnf = #facts o rep_bnf;
blanchet@49990
   299
val nwits_of_bnf = #nwits o rep_bnf;
blanchet@49990
   300
val wits_of_bnf = #wits o rep_bnf;
blanchet@49990
   301
blanchet@49990
   302
(*terms*)
blanchet@49990
   303
val map_of_bnf = #map o rep_bnf;
blanchet@49990
   304
val sets_of_bnf = #sets o rep_bnf;
blanchet@49990
   305
fun mk_map_of_bnf Ds Ts Us bnf =
blanchet@49990
   306
  let val bnf_rep = rep_bnf bnf;
blanchet@49990
   307
  in
blanchet@49990
   308
    Term.subst_atomic_types
blanchet@49990
   309
      ((#deads bnf_rep ~~ Ds) @ (#lives bnf_rep ~~ Ts) @ (#lives' bnf_rep ~~ Us)) (#map bnf_rep)
blanchet@49990
   310
  end;
blanchet@49990
   311
fun mk_sets_of_bnf Dss Tss bnf =
blanchet@49990
   312
  let val bnf_rep = rep_bnf bnf;
blanchet@49990
   313
  in
blanchet@49990
   314
    map2 (fn (Ds, Ts) => Term.subst_atomic_types
blanchet@49990
   315
      ((#deads bnf_rep ~~ Ds) @ (#lives bnf_rep ~~ Ts))) (Dss ~~ Tss) (#sets bnf_rep)
blanchet@49990
   316
  end;
blanchet@49990
   317
val bd_of_bnf = #bd o rep_bnf;
blanchet@49990
   318
fun mk_bd_of_bnf Ds Ts bnf =
blanchet@49990
   319
  let val bnf_rep = rep_bnf bnf;
blanchet@49990
   320
  in Term.subst_atomic_types ((#deads bnf_rep ~~ Ds) @ (#lives bnf_rep ~~ Ts)) (#bd bnf_rep) end;
blanchet@49990
   321
fun mk_wits_of_bnf Dss Tss bnf =
blanchet@49990
   322
  let
blanchet@49990
   323
    val bnf_rep = rep_bnf bnf;
blanchet@49990
   324
    val wits = map (fn x => (#I x, #wit x)) (#wits bnf_rep);
blanchet@49990
   325
  in
blanchet@49990
   326
    map2 (fn (Ds, Ts) => apsnd (Term.subst_atomic_types
blanchet@49990
   327
      ((#deads bnf_rep ~~ Ds) @ (#lives bnf_rep ~~ Ts)))) (Dss ~~ Tss) wits
blanchet@49990
   328
  end;
blanchet@49990
   329
val rel_of_bnf = #rel o rep_bnf;
blanchet@49990
   330
fun mk_rel_of_bnf Ds Ts Us bnf =
blanchet@49990
   331
  let val bnf_rep = rep_bnf bnf;
blanchet@49990
   332
  in
blanchet@49990
   333
    Term.subst_atomic_types
blanchet@49990
   334
      ((#deads bnf_rep ~~ Ds) @ (#lives bnf_rep ~~ Ts) @ (#lives' bnf_rep ~~ Us)) (#rel bnf_rep)
blanchet@49990
   335
  end;
blanchet@49990
   336
val pred_of_bnf = #pred o rep_bnf;
blanchet@49990
   337
fun mk_pred_of_bnf Ds Ts Us bnf =
blanchet@49990
   338
  let val bnf_rep = rep_bnf bnf;
blanchet@49990
   339
  in
blanchet@49990
   340
    Term.subst_atomic_types
blanchet@49990
   341
      ((#deads bnf_rep ~~ Ds) @ (#lives bnf_rep ~~ Ts) @ (#lives' bnf_rep ~~ Us)) (#pred bnf_rep)
blanchet@49990
   342
  end;
blanchet@49990
   343
blanchet@49990
   344
(*thms*)
blanchet@49990
   345
val bd_card_order_of_bnf = #bd_card_order o #axioms o rep_bnf;
blanchet@49990
   346
val bd_cinfinite_of_bnf = #bd_cinfinite o #axioms o rep_bnf;
blanchet@49990
   347
val bd_Card_order_of_bnf = #bd_Card_order o #facts o rep_bnf;
blanchet@49990
   348
val bd_Cinfinite_of_bnf = #bd_Cinfinite o #facts o rep_bnf;
blanchet@49990
   349
val bd_Cnotzero_of_bnf = #bd_Cnotzero o #facts o rep_bnf;
blanchet@49990
   350
val collect_set_natural_of_bnf = Lazy.force o #collect_set_natural o #facts o rep_bnf;
blanchet@49990
   351
val in_bd_of_bnf = #in_bd o #axioms o rep_bnf;
blanchet@49990
   352
val in_cong_of_bnf = Lazy.force o #in_cong o #facts o rep_bnf;
blanchet@49990
   353
val in_mono_of_bnf = Lazy.force o #in_mono o #facts o rep_bnf;
blanchet@49990
   354
val in_rel_of_bnf = Lazy.force o #in_rel o #facts o rep_bnf;
blanchet@49990
   355
val map_def_of_bnf = #map_def o #defs o rep_bnf;
blanchet@49990
   356
val map_id_of_bnf = #map_id o #axioms o rep_bnf;
blanchet@49990
   357
val map_id'_of_bnf = Lazy.force o #map_id' o #facts o rep_bnf;
blanchet@49990
   358
val map_comp_of_bnf = #map_comp o #axioms o rep_bnf;
blanchet@49990
   359
val map_comp'_of_bnf = Lazy.force o #map_comp' o #facts o rep_bnf;
blanchet@49990
   360
val map_cong_of_bnf = #map_cong o #axioms o rep_bnf;
blanchet@49990
   361
val map_wppull_of_bnf = Lazy.force o #map_wppull o #facts o rep_bnf;
blanchet@49990
   362
val map_wpull_of_bnf = #map_wpull o #axioms o rep_bnf;
blanchet@49990
   363
val pred_def_of_bnf = #pred_def o #defs o rep_bnf;
blanchet@49990
   364
val rel_cong_of_bnf = Lazy.force o #rel_cong o #facts o rep_bnf;
blanchet@49990
   365
val rel_mono_of_bnf = Lazy.force o #rel_mono o #facts o rep_bnf;
blanchet@49990
   366
val rel_def_of_bnf = #rel_def o #defs o rep_bnf;
blanchet@49990
   367
val rel_Id_of_bnf = Lazy.force o #rel_Id o #facts o rep_bnf;
blanchet@49990
   368
val rel_Gr_of_bnf = Lazy.force o #rel_Gr o #facts o rep_bnf;
blanchet@49990
   369
val rel_converse_of_bnf = Lazy.force o #rel_converse o #facts o rep_bnf;
blanchet@49990
   370
val rel_O_of_bnf = Lazy.force o #rel_O o #facts o rep_bnf;
blanchet@50468
   371
val rel_O_Gr_of_bnf = #rel_O_Gr o #axioms o rep_bnf;
blanchet@49990
   372
val set_bd_of_bnf = #set_bd o #axioms o rep_bnf;
blanchet@49990
   373
val set_defs_of_bnf = #set_defs o #defs o rep_bnf;
blanchet@49990
   374
val set_natural_of_bnf = #set_natural o #axioms o rep_bnf;
blanchet@49990
   375
val set_natural'_of_bnf = map Lazy.force o #set_natural' o #facts o rep_bnf;
blanchet@49990
   376
val wit_thms_of_bnf = maps #prop o wits_of_bnf;
blanchet@49990
   377
val wit_thmss_of_bnf = map #prop o wits_of_bnf;
blanchet@49990
   378
blanchet@49990
   379
fun mk_bnf name T live lives lives' dead deads map sets bd axioms defs facts wits rel pred =
blanchet@49990
   380
  BNF {name = name, T = T,
blanchet@49990
   381
       live = live, lives = lives, lives' = lives', dead = dead, deads = deads,
blanchet@49990
   382
       map = map, sets = sets, bd = bd,
blanchet@49990
   383
       axioms = axioms, defs = defs, facts = facts,
blanchet@49990
   384
       nwits = length wits, wits = wits, rel = rel, pred = pred};
blanchet@49990
   385
blanchet@49990
   386
fun morph_bnf phi (BNF {name = name, T = T, live = live, lives = lives, lives' = lives',
blanchet@49990
   387
  dead = dead, deads = deads, map = map, sets = sets, bd = bd,
blanchet@49990
   388
  axioms = axioms, defs = defs, facts = facts,
blanchet@49990
   389
  nwits = nwits, wits = wits, rel = rel, pred = pred}) =
blanchet@49990
   390
  BNF {name = Morphism.binding phi name, T = Morphism.typ phi T,
blanchet@49990
   391
    live = live, lives = List.map (Morphism.typ phi) lives,
blanchet@49990
   392
    lives' = List.map (Morphism.typ phi) lives',
blanchet@49990
   393
    dead = dead, deads = List.map (Morphism.typ phi) deads,
blanchet@49990
   394
    map = Morphism.term phi map, sets = List.map (Morphism.term phi) sets,
blanchet@49990
   395
    bd = Morphism.term phi bd,
blanchet@49990
   396
    axioms = morph_axioms phi axioms,
blanchet@49990
   397
    defs = morph_defs phi defs,
blanchet@49990
   398
    facts = morph_facts phi facts,
blanchet@49990
   399
    nwits = nwits,
blanchet@49990
   400
    wits = List.map (morph_witness phi) wits,
blanchet@49990
   401
    rel = Morphism.term phi rel, pred = Morphism.term phi pred};
blanchet@49990
   402
blanchet@49990
   403
fun eq_bnf (BNF {T = T1, live = live1, dead = dead1, ...},
blanchet@49990
   404
  BNF {T = T2, live = live2, dead = dead2, ...}) =
blanchet@49990
   405
  Type.could_unify (T1, T2) andalso live1 = live2 andalso dead1 = dead2;
blanchet@49990
   406
blanchet@49990
   407
structure Data = Generic_Data
blanchet@49990
   408
(
blanchet@49990
   409
  type T = BNF Symtab.table;
blanchet@49990
   410
  val empty = Symtab.empty;
blanchet@49990
   411
  val extend = I;
blanchet@49990
   412
  val merge = Symtab.merge (eq_bnf);
blanchet@49990
   413
);
blanchet@49990
   414
blanchet@49990
   415
val bnf_of = Symtab.lookup o Data.get o Context.Proof;
blanchet@49990
   416
blanchet@49990
   417
blanchet@49990
   418
blanchet@49990
   419
(* Utilities *)
blanchet@49990
   420
blanchet@49990
   421
fun normalize_set insts instA set =
blanchet@49990
   422
  let
blanchet@49990
   423
    val (T, T') = dest_funT (fastype_of set);
blanchet@49990
   424
    val A = fst (Term.dest_TVar (HOLogic.dest_setT T'));
blanchet@49990
   425
    val params = Term.add_tvar_namesT T [];
blanchet@49990
   426
  in Term.subst_TVars ((A :: params) ~~ (instA :: insts)) set end;
blanchet@49990
   427
blanchet@49990
   428
fun normalize_rel ctxt instTs instA instB rel =
blanchet@49990
   429
  let
blanchet@49990
   430
    val thy = Proof_Context.theory_of ctxt;
blanchet@49990
   431
    val tyenv =
blanchet@49990
   432
      Sign.typ_match thy (fastype_of rel, Library.foldr (op -->) (instTs, mk_relT (instA, instB)))
blanchet@49990
   433
        Vartab.empty;
blanchet@50468
   434
  in Envir.subst_term (tyenv, Vartab.empty) rel end
blanchet@50468
   435
  handle Type.TYPE_MATCH => error "Bad relator";
blanchet@49990
   436
blanchet@49990
   437
fun normalize_pred ctxt instTs instA instB pred =
blanchet@49990
   438
  let
blanchet@49990
   439
    val thy = Proof_Context.theory_of ctxt;
blanchet@49990
   440
    val tyenv =
blanchet@49990
   441
      Sign.typ_match thy (fastype_of pred,
blanchet@49990
   442
        Library.foldr (op -->) (instTs, instA --> instB --> HOLogic.boolT)) Vartab.empty;
blanchet@50468
   443
  in Envir.subst_term (tyenv, Vartab.empty) pred end
blanchet@50468
   444
  handle Type.TYPE_MATCH => error "Bad predicator";
blanchet@49990
   445
blanchet@49990
   446
fun normalize_wit insts CA As wit =
blanchet@49990
   447
  let
blanchet@49990
   448
    fun strip_param (Ts, T as Type (@{type_name fun}, [T1, T2])) =
blanchet@49990
   449
        if Type.raw_instance (CA, T) then (Ts, T) else strip_param (T1 :: Ts, T2)
blanchet@49990
   450
      | strip_param x = x;
blanchet@49990
   451
    val (Ts, T) = strip_param ([], fastype_of wit);
blanchet@49990
   452
    val subst = Term.add_tvar_namesT T [] ~~ insts;
blanchet@49990
   453
    fun find y = find_index (fn x => x = y) As;
blanchet@49990
   454
  in
blanchet@49990
   455
    (map (find o Term.typ_subst_TVars subst) (rev Ts), Term.subst_TVars subst wit)
blanchet@49990
   456
  end;
blanchet@49990
   457
blanchet@49990
   458
fun minimize_wits wits =
blanchet@49990
   459
 let
blanchet@49990
   460
   fun minimize done [] = done
traytel@50118
   461
     | minimize done ((I, wit) :: todo) =
blanchet@49990
   462
       if exists (fn (J, _) => subset (op =) (J, I)) (done @ todo)
blanchet@49990
   463
       then minimize done todo
blanchet@49990
   464
       else minimize ((I, wit) :: done) todo;
blanchet@49990
   465
 in minimize [] wits end;
blanchet@49990
   466
blanchet@49990
   467
fun unfold_defs_tac lthy defs mk_tac context = Local_Defs.unfold_tac lthy defs THEN mk_tac context;
blanchet@49990
   468
blanchet@49990
   469
blanchet@49990
   470
blanchet@49990
   471
(* Names *)
blanchet@49990
   472
blanchet@49990
   473
val mapN = "map";
blanchet@49990
   474
val setN = "set";
blanchet@49990
   475
fun mk_setN i = setN ^ nonzero_string_of_int i;
blanchet@49990
   476
val bdN = "bd";
blanchet@49990
   477
val witN = "wit";
blanchet@49990
   478
fun mk_witN i = witN ^ nonzero_string_of_int i;
blanchet@49990
   479
val relN = "rel";
blanchet@49990
   480
val predN = "pred";
blanchet@49990
   481
blanchet@49990
   482
val bd_card_orderN = "bd_card_order";
blanchet@49990
   483
val bd_cinfiniteN = "bd_cinfinite";
blanchet@49990
   484
val bd_Card_orderN = "bd_Card_order";
blanchet@49990
   485
val bd_CinfiniteN = "bd_Cinfinite";
blanchet@49990
   486
val bd_CnotzeroN = "bd_Cnotzero";
blanchet@49990
   487
val collect_set_naturalN = "collect_set_natural";
blanchet@49990
   488
val in_bdN = "in_bd";
blanchet@49990
   489
val in_monoN = "in_mono";
blanchet@49990
   490
val in_relN = "in_rel";
blanchet@49990
   491
val map_idN = "map_id";
blanchet@49990
   492
val map_id'N = "map_id'";
blanchet@49990
   493
val map_compN = "map_comp";
blanchet@49990
   494
val map_comp'N = "map_comp'";
blanchet@49990
   495
val map_congN = "map_cong";
blanchet@49990
   496
val map_wpullN = "map_wpull";
blanchet@49990
   497
val rel_IdN = "rel_Id";
blanchet@49990
   498
val rel_GrN = "rel_Gr";
blanchet@49990
   499
val rel_converseN = "rel_converse";
traytel@50401
   500
val rel_monoN = "rel_mono"
blanchet@49990
   501
val rel_ON = "rel_comp";
blanchet@50467
   502
val rel_O_GrN = "rel_comp_Gr";
blanchet@49990
   503
val set_naturalN = "set_natural";
blanchet@49990
   504
val set_natural'N = "set_natural'";
blanchet@49990
   505
val set_bdN = "set_bd";
blanchet@49990
   506
blanchet@49990
   507
datatype const_policy = Dont_Inline | Hardly_Inline | Smart_Inline | Do_Inline;
blanchet@49990
   508
blanchet@49990
   509
datatype fact_policy =
blanchet@49990
   510
  Derive_Some_Facts | Derive_All_Facts | Derive_All_Facts_Note_Most | Note_All_Facts_and_Axioms;
blanchet@49990
   511
blanchet@49990
   512
val bnf_note_all = Attrib.setup_config_bool @{binding bnf_note_all} (K false);
blanchet@49990
   513
traytel@50450
   514
fun user_policy policy ctxt =
traytel@50450
   515
  if Config.get ctxt bnf_note_all then Note_All_Facts_and_Axioms else policy;
blanchet@49990
   516
blanchet@49990
   517
val smart_max_inline_size = 25; (*FUDGE*)
blanchet@49990
   518
blanchet@50475
   519
fun is_refl thm =
blanchet@50475
   520
  op aconv (HOLogic.dest_eq (HOLogic.dest_Trueprop (Thm.prop_of thm)))
blanchet@50475
   521
  handle TERM _ => false;
blanchet@49990
   522
blanchet@50475
   523
val no_refl = filter_out is_refl;
blanchet@50475
   524
val no_reflexive = filter_out Thm.is_reflexive;
blanchet@49990
   525
blanchet@49990
   526
blanchet@49990
   527
(* Define new BNFs *)
blanchet@49990
   528
blanchet@50034
   529
fun prepare_def const_policy mk_fact_policy qualify prep_term Ds_opt
blanchet@50474
   530
  (((((raw_b, raw_map), raw_sets), raw_bd_Abs), raw_wits), raw_rel_opt) no_defs_lthy =
blanchet@49990
   531
  let
blanchet@49990
   532
    val fact_policy = mk_fact_policy no_defs_lthy;
blanchet@49990
   533
    val b = qualify raw_b;
blanchet@49990
   534
    val live = length raw_sets;
blanchet@49990
   535
    val nwits = length raw_wits;
blanchet@49990
   536
blanchet@49990
   537
    val map_rhs = prep_term no_defs_lthy raw_map;
blanchet@49990
   538
    val set_rhss = map (prep_term no_defs_lthy) raw_sets;
blanchet@49990
   539
    val (bd_rhsT, bd_rhs) = (case prep_term no_defs_lthy raw_bd_Abs of
blanchet@49990
   540
      Abs (_, T, t) => (T, t)
blanchet@49990
   541
    | _ => error "Bad bound constant");
blanchet@49990
   542
    val wit_rhss = map (prep_term no_defs_lthy) raw_wits;
blanchet@49990
   543
traytel@50449
   544
    fun err T =
traytel@50449
   545
      error ("Trying to register the type " ^ quote (Syntax.string_of_typ no_defs_lthy T) ^
traytel@50449
   546
        " as unnamed BNF");
traytel@50449
   547
traytel@50449
   548
    val (b, key) =
traytel@50449
   549
      if Binding.eq_name (b, Binding.empty) then
traytel@50449
   550
        (case bd_rhsT of
traytel@50449
   551
          Type (C, Ts) => if forall (is_some o try dest_TFree) Ts
traytel@50449
   552
            then (Binding.qualified_name C, C) else err bd_rhsT
traytel@50449
   553
        | T => err T)
traytel@50449
   554
      else (b, Local_Theory.full_name no_defs_lthy b);
traytel@50449
   555
blanchet@50471
   556
    fun maybe_define (b, rhs) lthy =
blanchet@49990
   557
      let
blanchet@49990
   558
        val inline =
blanchet@49990
   559
          (case const_policy of
blanchet@49990
   560
            Dont_Inline => false
blanchet@49990
   561
          | Hardly_Inline => Term.is_Free rhs orelse Term.is_Const rhs
blanchet@49990
   562
          | Smart_Inline => Term.size_of_term rhs <= smart_max_inline_size
blanchet@49990
   563
          | Do_Inline => true)
blanchet@49990
   564
      in
blanchet@49990
   565
        if inline then
blanchet@50475
   566
          ((rhs, Drule.reflexive_thm), lthy)
blanchet@49990
   567
        else
blanchet@49990
   568
          let val b = b () in
blanchet@49990
   569
            apfst (apsnd snd) (Local_Theory.define ((b, NoSyn), ((Thm.def_binding b, []), rhs))
blanchet@49990
   570
              lthy)
blanchet@49990
   571
          end
blanchet@49990
   572
      end;
blanchet@49990
   573
blanchet@50474
   574
    fun maybe_restore lthy_old lthy =
blanchet@50474
   575
      lthy |> not (pointer_eq (lthy_old, lthy)) ? Local_Theory.restore;
blanchet@50474
   576
blanchet@50474
   577
    val map_bind_def = (fn () => Binding.suffix_name ("_" ^ mapN) b, map_rhs);
blanchet@50474
   578
    val set_binds_defs =
blanchet@50474
   579
      let
blanchet@50474
   580
        val bs = if live = 1 then [fn () => Binding.suffix_name ("_" ^ setN) b]
blanchet@50474
   581
          else map (fn i => fn () => Binding.suffix_name ("_" ^ mk_setN i) b) (1 upto live)
blanchet@50474
   582
      in map2 pair bs set_rhss end;
blanchet@50474
   583
    val bd_bind_def = (fn () => Binding.suffix_name ("_" ^ bdN) b, bd_rhs);
blanchet@50474
   584
    val wit_binds_defs =
blanchet@50474
   585
      let
blanchet@50474
   586
        val bs = if nwits = 1 then [fn () => Binding.suffix_name ("_" ^ witN) b]
blanchet@50474
   587
          else map (fn i => fn () => Binding.suffix_name ("_" ^ mk_witN i) b) (1 upto nwits);
blanchet@50474
   588
      in map2 pair bs wit_rhss end;
blanchet@50474
   589
blanchet@50474
   590
    val (((((bnf_map_term, raw_map_def),
blanchet@49990
   591
      (bnf_set_terms, raw_set_defs)),
blanchet@49990
   592
      (bnf_bd_term, raw_bd_def)),
blanchet@50474
   593
      (bnf_wit_terms, raw_wit_defs)), (lthy, lthy_old)) =
blanchet@49990
   594
        no_defs_lthy
blanchet@50471
   595
        |> maybe_define map_bind_def
blanchet@50471
   596
        ||>> apfst split_list o fold_map maybe_define set_binds_defs
blanchet@50471
   597
        ||>> maybe_define bd_bind_def
blanchet@50471
   598
        ||>> apfst split_list o fold_map maybe_define wit_binds_defs
blanchet@49990
   599
        ||> `(maybe_restore no_defs_lthy);
blanchet@49990
   600
blanchet@50474
   601
    val phi = Proof_Context.export_morphism lthy_old lthy;
blanchet@49990
   602
blanchet@49990
   603
    val bnf_map_def = Morphism.thm phi raw_map_def;
blanchet@49990
   604
    val bnf_set_defs = map (Morphism.thm phi) raw_set_defs;
blanchet@49990
   605
    val bnf_bd_def = Morphism.thm phi raw_bd_def;
blanchet@49990
   606
    val bnf_wit_defs = map (Morphism.thm phi) raw_wit_defs;
blanchet@49990
   607
blanchet@49990
   608
    val bnf_map = Morphism.term phi bnf_map_term;
blanchet@49990
   609
blanchet@49990
   610
    (*TODO: handle errors*)
blanchet@49990
   611
    (*simple shape analysis of a map function*)
traytel@50410
   612
    val ((alphas, betas), (CA, _)) =
traytel@50410
   613
      fastype_of bnf_map
traytel@50410
   614
      |> strip_typeN live
traytel@50410
   615
      |>> map_split dest_funT
traytel@50410
   616
      ||> dest_funT
traytel@50410
   617
      handle TYPE _ => error "Bad map function";
blanchet@49990
   618
blanchet@49990
   619
    val CA_params = map TVar (Term.add_tvarsT CA []);
blanchet@49990
   620
blanchet@49990
   621
    val bnf_sets = map2 (normalize_set CA_params) alphas (map (Morphism.term phi) bnf_set_terms);
blanchet@49990
   622
    val bdT = Morphism.typ phi bd_rhsT;
blanchet@49990
   623
    val bnf_bd =
blanchet@49990
   624
      Term.subst_TVars (Term.add_tvar_namesT bdT [] ~~ CA_params) (Morphism.term phi bnf_bd_term);
blanchet@49990
   625
    val bnf_wits = map (normalize_wit CA_params CA alphas o Morphism.term phi) bnf_wit_terms;
blanchet@49990
   626
blanchet@49990
   627
    (*TODO: assert Ds = (TVars of bnf_map) \ (alphas @ betas) as sets*)
blanchet@49990
   628
    val deads = (case Ds_opt of
blanchet@49990
   629
      NONE => subtract (op =) (alphas @ betas) (map TVar (Term.add_tvars bnf_map []))
blanchet@49990
   630
    | SOME Ds => map (Morphism.typ phi) Ds);
blanchet@49990
   631
    val dead = length deads;
blanchet@49990
   632
blanchet@49990
   633
    (*TODO: further checks of type of bnf_map*)
blanchet@49990
   634
    (*TODO: check types of bnf_sets*)
blanchet@49990
   635
    (*TODO: check type of bnf_bd*)
blanchet@50469
   636
    (*TODO: check type of bnf_rel*)
blanchet@49990
   637
blanchet@49990
   638
    val ((((((((((As', Bs'), Cs), Ds), B1Ts), B2Ts), domTs), ranTs), ranTs'), ranTs''),
blanchet@50474
   639
      (Ts, T)) = lthy
blanchet@49990
   640
      |> mk_TFrees live
blanchet@49990
   641
      ||>> mk_TFrees live
blanchet@49990
   642
      ||>> mk_TFrees live
blanchet@49990
   643
      ||>> mk_TFrees dead
blanchet@49990
   644
      ||>> mk_TFrees live
blanchet@49990
   645
      ||>> mk_TFrees live
blanchet@49990
   646
      ||>> mk_TFrees live
blanchet@49990
   647
      ||>> mk_TFrees live
blanchet@49990
   648
      ||>> mk_TFrees live
blanchet@49990
   649
      ||>> mk_TFrees live
blanchet@49990
   650
      ||> fst o mk_TFrees 1
blanchet@49990
   651
      ||> the_single
blanchet@49990
   652
      ||> `(replicate live);
blanchet@49990
   653
blanchet@49990
   654
    fun mk_bnf_map As' Bs' =
blanchet@49990
   655
      Term.subst_atomic_types ((deads ~~ Ds) @ (alphas ~~ As') @ (betas ~~ Bs')) bnf_map;
blanchet@50468
   656
    fun mk_bnf_t As' = Term.subst_atomic_types ((deads ~~ Ds) @ (alphas ~~ As'));
blanchet@50468
   657
    fun mk_bnf_T As' = Term.typ_subst_atomic ((deads ~~ Ds) @ (alphas ~~ As'));
blanchet@50468
   658
blanchet@49990
   659
    val (setRTs, RTs) = map_split (`HOLogic.mk_setT o HOLogic.mk_prodT) (As' ~~ Bs');
blanchet@49990
   660
    val setRTsAsCs = map (HOLogic.mk_setT o HOLogic.mk_prodT) (As' ~~ Cs);
blanchet@49990
   661
    val setRTsBsCs = map (HOLogic.mk_setT o HOLogic.mk_prodT) (Bs' ~~ Cs);
blanchet@49990
   662
    val setRT's = map (HOLogic.mk_setT o HOLogic.mk_prodT) (Bs' ~~ As');
blanchet@49990
   663
    val self_setRTs = map (HOLogic.mk_setT o HOLogic.mk_prodT) (As' ~~ As');
blanchet@49990
   664
    val QTs = map2 (fn T => fn U => T --> U --> HOLogic.boolT) As' Bs';
blanchet@49990
   665
blanchet@50468
   666
    val CA' = mk_bnf_T As' CA;
blanchet@50468
   667
    val CB' = mk_bnf_T Bs' CA;
blanchet@50468
   668
    val CC' = mk_bnf_T Cs CA;
blanchet@50468
   669
    val CRs' = mk_bnf_T RTs CA;
blanchet@50468
   670
blanchet@49990
   671
    val bnf_map_AsAs = mk_bnf_map As' As';
blanchet@49990
   672
    val bnf_map_AsBs = mk_bnf_map As' Bs';
blanchet@49990
   673
    val bnf_map_AsCs = mk_bnf_map As' Cs;
blanchet@49990
   674
    val bnf_map_BsCs = mk_bnf_map Bs' Cs;
blanchet@49990
   675
    val bnf_sets_As = map (mk_bnf_t As') bnf_sets;
blanchet@49990
   676
    val bnf_sets_Bs = map (mk_bnf_t Bs') bnf_sets;
blanchet@49990
   677
    val bnf_bd_As = mk_bnf_t As' bnf_bd;
blanchet@49990
   678
    val bnf_wit_As = map (apsnd (mk_bnf_t As')) bnf_wits;
blanchet@49990
   679
blanchet@49990
   680
    val ((((((((((((((((((((((((fs, fs_copy), gs), hs), (x, x')), (y, y')), (z, z')), zs), As),
blanchet@50474
   681
      As_copy), Xs), B1s), B2s), f1s), f2s), e1s), e2s), p1s), p2s), bs), (Rs, Rs')), Rs_copy), Ss),
blanchet@50474
   682
      (Qs, Qs')), _) = lthy
blanchet@49990
   683
      |> mk_Frees "f" (map2 (curry (op -->)) As' Bs')
blanchet@49990
   684
      ||>> mk_Frees "f" (map2 (curry (op -->)) As' Bs')
blanchet@49990
   685
      ||>> mk_Frees "g" (map2 (curry (op -->)) Bs' Cs)
blanchet@49990
   686
      ||>> mk_Frees "h" (map2 (curry (op -->)) As' Ts)
blanchet@49990
   687
      ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "x") CA'
blanchet@49990
   688
      ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "y") CB'
blanchet@49990
   689
      ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "z") CRs'
blanchet@49990
   690
      ||>> mk_Frees "z" As'
blanchet@49990
   691
      ||>> mk_Frees "A" (map HOLogic.mk_setT As')
blanchet@49990
   692
      ||>> mk_Frees "A" (map HOLogic.mk_setT As')
blanchet@49990
   693
      ||>> mk_Frees "A" (map HOLogic.mk_setT domTs)
blanchet@49990
   694
      ||>> mk_Frees "B1" (map HOLogic.mk_setT B1Ts)
blanchet@49990
   695
      ||>> mk_Frees "B2" (map HOLogic.mk_setT B2Ts)
blanchet@49990
   696
      ||>> mk_Frees "f1" (map2 (curry (op -->)) B1Ts ranTs)
blanchet@49990
   697
      ||>> mk_Frees "f2" (map2 (curry (op -->)) B2Ts ranTs)
blanchet@49990
   698
      ||>> mk_Frees "e1" (map2 (curry (op -->)) B1Ts ranTs')
blanchet@49990
   699
      ||>> mk_Frees "e2" (map2 (curry (op -->)) B2Ts ranTs'')
blanchet@49990
   700
      ||>> mk_Frees "p1" (map2 (curry (op -->)) domTs B1Ts)
blanchet@49990
   701
      ||>> mk_Frees "p2" (map2 (curry (op -->)) domTs B2Ts)
blanchet@49990
   702
      ||>> mk_Frees "b" As'
blanchet@50474
   703
      ||>> mk_Frees' "R" setRTs
blanchet@49990
   704
      ||>> mk_Frees "R" setRTs
blanchet@49990
   705
      ||>> mk_Frees "S" setRTsBsCs
blanchet@49990
   706
      ||>> mk_Frees' "Q" QTs;
blanchet@49990
   707
blanchet@50474
   708
    (*Gr (in R1 .. Rn) (map fst .. fst)^-1 O Gr (in R1 .. Rn) (map snd .. snd)*)
blanchet@50474
   709
    val rel_O_Gr_rhs =
blanchet@50474
   710
      let
blanchet@50474
   711
        val map1 = Term.list_comb (mk_bnf_map RTs As', map fst_const RTs);
blanchet@50474
   712
        val map2 = Term.list_comb (mk_bnf_map RTs Bs', map snd_const RTs);
blanchet@50474
   713
        val bnf_in = mk_in (map Free Rs') (map (mk_bnf_t RTs) bnf_sets) CRs';
blanchet@50474
   714
      in
blanchet@50474
   715
        mk_rel_comp (mk_converse (mk_Gr bnf_in map1), mk_Gr bnf_in map2)
blanchet@50474
   716
      end;
blanchet@50474
   717
blanchet@50474
   718
    val rel_rhs = (case raw_rel_opt of
blanchet@50474
   719
        NONE => fold_rev Term.absfree Rs' rel_O_Gr_rhs
blanchet@50474
   720
      | SOME raw_rel => prep_term no_defs_lthy raw_rel);
blanchet@50474
   721
blanchet@50474
   722
    val rel_bind_def = (fn () => Binding.suffix_name ("_" ^ relN) b, rel_rhs);
blanchet@50474
   723
blanchet@50474
   724
    val ((bnf_rel_term, raw_rel_def), (lthy, lthy_old)) =
blanchet@50474
   725
      lthy
blanchet@50474
   726
      |> maybe_define rel_bind_def
blanchet@50474
   727
      ||> `(maybe_restore lthy);
blanchet@50474
   728
blanchet@50474
   729
    val phi = Proof_Context.export_morphism lthy_old lthy;
blanchet@50474
   730
    val bnf_rel_def = Morphism.thm phi raw_rel_def;
blanchet@50474
   731
    val bnf_rel = Morphism.term phi bnf_rel_term;
blanchet@50474
   732
blanchet@50474
   733
    fun mk_bnf_rel setRTs CA' CB' = normalize_rel lthy setRTs CA' CB' bnf_rel;
blanchet@50474
   734
blanchet@50474
   735
    val relAsBs = mk_bnf_rel setRTs CA' CB';
blanchet@50474
   736
blanchet@50469
   737
    val pred_rhs = fold absfree (y' :: x' :: rev Qs') (HOLogic.mk_mem (HOLogic.mk_prod (x, y),
blanchet@50469
   738
      Term.list_comb (relAsBs, map3 (fn Q => fn T => fn U =>
blanchet@50469
   739
        HOLogic.Collect_const (HOLogic.mk_prodT (T, U)) $ HOLogic.mk_split Q)
blanchet@50469
   740
        Qs As' Bs')));
blanchet@50469
   741
    val pred_bind_def = (fn () => Binding.suffix_name ("_" ^ predN) b, pred_rhs);
blanchet@50469
   742
blanchet@50474
   743
    val ((bnf_pred_term, raw_pred_def), (lthy, lthy_old)) =
blanchet@50474
   744
      lthy
blanchet@50471
   745
      |> maybe_define pred_bind_def
blanchet@50474
   746
      ||> `(maybe_restore lthy);
blanchet@50469
   747
blanchet@50474
   748
    val phi = Proof_Context.export_morphism lthy_old lthy;
blanchet@50475
   749
    val bnf_pred_def = Morphism.thm phi raw_pred_def;
blanchet@50471
   750
    val bnf_pred = Morphism.term phi bnf_pred_term;
blanchet@50471
   751
blanchet@50474
   752
    fun mk_bnf_pred QTs CA' CB' = normalize_pred lthy QTs CA' CB' bnf_pred;
blanchet@50471
   753
blanchet@50471
   754
    val pred = mk_bnf_pred QTs CA' CB';
blanchet@50471
   755
blanchet@50475
   756
    val _ = case no_reflexive (raw_map_def :: raw_set_defs @ [raw_bd_def] @
blanchet@50474
   757
        raw_wit_defs @ [raw_rel_def, raw_pred_def]) of
blanchet@50474
   758
        [] => ()
blanchet@50474
   759
      | defs => Proof_Display.print_consts true lthy_old (K false)
blanchet@50474
   760
          (map (dest_Free o fst o Logic.dest_equals o prop_of) defs);
blanchet@50474
   761
blanchet@50473
   762
    val map_id_goal =
blanchet@49990
   763
      let
blanchet@50033
   764
        val bnf_map_app_id = Term.list_comb (bnf_map_AsAs, map HOLogic.id_const As');
blanchet@49990
   765
      in
blanchet@49990
   766
        HOLogic.mk_Trueprop
blanchet@49990
   767
          (HOLogic.mk_eq (bnf_map_app_id, HOLogic.id_const CA'))
blanchet@49990
   768
      end;
blanchet@49990
   769
blanchet@50473
   770
    val map_comp_goal =
blanchet@49990
   771
      let
blanchet@50033
   772
        val bnf_map_app_comp = Term.list_comb (bnf_map_AsCs, map2 (curry HOLogic.mk_comp) gs fs);
blanchet@49990
   773
        val comp_bnf_map_app = HOLogic.mk_comp
blanchet@49990
   774
          (Term.list_comb (bnf_map_BsCs, gs),
blanchet@49990
   775
           Term.list_comb (bnf_map_AsBs, fs));
blanchet@49990
   776
      in
blanchet@50138
   777
        fold_rev Logic.all (fs @ gs) (mk_Trueprop_eq (bnf_map_app_comp, comp_bnf_map_app))
blanchet@49990
   778
      end;
blanchet@49990
   779
blanchet@50473
   780
    val map_cong_goal =
blanchet@49990
   781
      let
blanchet@49990
   782
        fun mk_prem z set f f_copy =
blanchet@49990
   783
          Logic.all z (Logic.mk_implies
blanchet@49990
   784
            (HOLogic.mk_Trueprop (HOLogic.mk_mem (z, set $ x)),
blanchet@50138
   785
            mk_Trueprop_eq (f $ z, f_copy $ z)));
blanchet@49990
   786
        val prems = map4 mk_prem zs bnf_sets_As fs fs_copy;
blanchet@49990
   787
        val eq = HOLogic.mk_eq (Term.list_comb (bnf_map_AsBs, fs) $ x,
blanchet@49990
   788
          Term.list_comb (bnf_map_AsBs, fs_copy) $ x);
blanchet@49990
   789
      in
blanchet@49990
   790
        fold_rev Logic.all (x :: fs @ fs_copy)
blanchet@49990
   791
          (Logic.list_implies (prems, HOLogic.mk_Trueprop eq))
blanchet@49990
   792
      end;
blanchet@49990
   793
blanchet@50473
   794
    val set_naturals_goal =
blanchet@49990
   795
      let
blanchet@49990
   796
        fun mk_goal setA setB f =
blanchet@49990
   797
          let
blanchet@49990
   798
            val set_comp_map =
blanchet@49990
   799
              HOLogic.mk_comp (setB, Term.list_comb (bnf_map_AsBs, fs));
blanchet@49990
   800
            val image_comp_set = HOLogic.mk_comp (mk_image f, setA);
blanchet@49990
   801
          in
blanchet@50138
   802
            fold_rev Logic.all fs (mk_Trueprop_eq (set_comp_map, image_comp_set))
blanchet@49990
   803
          end;
blanchet@49990
   804
      in
blanchet@49990
   805
        map3 mk_goal bnf_sets_As bnf_sets_Bs fs
blanchet@49990
   806
      end;
blanchet@49990
   807
blanchet@50473
   808
    val card_order_bd_goal = HOLogic.mk_Trueprop (mk_card_order bnf_bd_As);
blanchet@49990
   809
blanchet@50473
   810
    val cinfinite_bd_goal = HOLogic.mk_Trueprop (mk_cinfinite bnf_bd_As);
blanchet@49990
   811
blanchet@50473
   812
    val set_bds_goal =
blanchet@49990
   813
      let
blanchet@49990
   814
        fun mk_goal set =
blanchet@49990
   815
          Logic.all x (HOLogic.mk_Trueprop (mk_ordLeq (mk_card_of (set $ x)) bnf_bd_As));
blanchet@49990
   816
      in
blanchet@49990
   817
        map mk_goal bnf_sets_As
blanchet@49990
   818
      end;
blanchet@49990
   819
blanchet@50473
   820
    val in_bd_goal =
blanchet@49990
   821
      let
blanchet@49990
   822
        val bd = mk_cexp
blanchet@49990
   823
          (if live = 0 then ctwo
blanchet@49990
   824
            else mk_csum (Library.foldr1 (uncurry mk_csum) (map mk_card_of As)) ctwo)
blanchet@49990
   825
          bnf_bd_As;
blanchet@49990
   826
      in
blanchet@49990
   827
        fold_rev Logic.all As
blanchet@49990
   828
          (HOLogic.mk_Trueprop (mk_ordLeq (mk_card_of (mk_in As bnf_sets_As CA')) bd))
blanchet@49990
   829
      end;
blanchet@49990
   830
blanchet@50473
   831
    val map_wpull_goal =
blanchet@49990
   832
      let
blanchet@49990
   833
        val prems = map HOLogic.mk_Trueprop
blanchet@49990
   834
          (map8 mk_wpull Xs B1s B2s f1s f2s (replicate live NONE) p1s p2s);
blanchet@49990
   835
        val CX = mk_bnf_T domTs CA;
blanchet@49990
   836
        val CB1 = mk_bnf_T B1Ts CA;
blanchet@49990
   837
        val CB2 = mk_bnf_T B2Ts CA;
blanchet@49990
   838
        val bnf_sets_CX = map2 (normalize_set (map (mk_bnf_T domTs) CA_params)) domTs bnf_sets;
blanchet@49990
   839
        val bnf_sets_CB1 = map2 (normalize_set (map (mk_bnf_T B1Ts) CA_params)) B1Ts bnf_sets;
blanchet@49990
   840
        val bnf_sets_CB2 = map2 (normalize_set (map (mk_bnf_T B2Ts) CA_params)) B2Ts bnf_sets;
blanchet@49990
   841
        val bnf_map_app_f1 = Term.list_comb (mk_bnf_map B1Ts ranTs, f1s);
blanchet@49990
   842
        val bnf_map_app_f2 = Term.list_comb (mk_bnf_map B2Ts ranTs, f2s);
blanchet@49990
   843
        val bnf_map_app_p1 = Term.list_comb (mk_bnf_map domTs B1Ts, p1s);
blanchet@49990
   844
        val bnf_map_app_p2 = Term.list_comb (mk_bnf_map domTs B2Ts, p2s);
blanchet@49990
   845
blanchet@49990
   846
        val map_wpull = mk_wpull (mk_in Xs bnf_sets_CX CX)
blanchet@49990
   847
          (mk_in B1s bnf_sets_CB1 CB1) (mk_in B2s bnf_sets_CB2 CB2)
blanchet@49990
   848
          bnf_map_app_f1 bnf_map_app_f2 NONE bnf_map_app_p1 bnf_map_app_p2;
blanchet@49990
   849
      in
blanchet@49990
   850
        fold_rev Logic.all (Xs @ B1s @ B2s @ f1s @ f2s @ p1s @ p2s)
blanchet@49990
   851
          (Logic.list_implies (prems, HOLogic.mk_Trueprop map_wpull))
blanchet@49990
   852
      end;
blanchet@49990
   853
blanchet@50473
   854
    val rel_O_Gr_goal =
blanchet@50474
   855
      fold_rev Logic.all Rs (mk_Trueprop_eq (Term.list_comb (relAsBs, Rs), rel_O_Gr_rhs));
blanchet@50468
   856
blanchet@50475
   857
    val goals = zip_axioms map_id_goal map_comp_goal map_cong_goal set_naturals_goal
blanchet@50473
   858
      card_order_bd_goal cinfinite_bd_goal set_bds_goal in_bd_goal map_wpull_goal rel_O_Gr_goal;
blanchet@49990
   859
blanchet@49990
   860
    fun mk_wit_goals (I, wit) =
blanchet@49990
   861
      let
blanchet@49990
   862
        val xs = map (nth bs) I;
blanchet@49990
   863
        fun wit_goal i =
blanchet@49990
   864
          let
blanchet@49990
   865
            val z = nth zs i;
blanchet@49990
   866
            val set_wit = nth bnf_sets_As i $ Term.list_comb (wit, xs);
blanchet@49990
   867
            val concl = HOLogic.mk_Trueprop
blanchet@49990
   868
              (if member (op =) I i then HOLogic.mk_eq (z, nth bs i)
blanchet@49990
   869
              else @{term False});
blanchet@49990
   870
          in
blanchet@49990
   871
            fold_rev Logic.all (z :: xs)
blanchet@49990
   872
              (Logic.mk_implies (HOLogic.mk_Trueprop (HOLogic.mk_mem (z, set_wit)), concl))
blanchet@49990
   873
          end;
blanchet@49990
   874
      in
blanchet@49990
   875
        map wit_goal (0 upto live - 1)
blanchet@49990
   876
      end;
blanchet@49990
   877
blanchet@49990
   878
    val wit_goalss = map mk_wit_goals bnf_wit_As;
blanchet@49990
   879
blanchet@49990
   880
    fun after_qed thms lthy =
blanchet@49990
   881
      let
blanchet@49990
   882
        val (axioms, wit_thms) = apfst (mk_axioms live) (chop (length goals) thms);
blanchet@49990
   883
traytel@50124
   884
        val bd_Card_order = #bd_card_order axioms RS @{thm conjunct2[OF card_order_on_Card_order]};
blanchet@49990
   885
        val bd_Cinfinite = @{thm conjI} OF [#bd_cinfinite axioms, bd_Card_order];
blanchet@49990
   886
        val bd_Cnotzero = bd_Cinfinite RS @{thm Cinfinite_Cnotzero};
blanchet@49990
   887
blanchet@49990
   888
        fun mk_lazy f = if fact_policy <> Derive_Some_Facts then Lazy.value (f ()) else Lazy.lazy f;
blanchet@49990
   889
blanchet@49990
   890
        fun mk_collect_set_natural () =
blanchet@49990
   891
          let
blanchet@49990
   892
            val defT = mk_bnf_T Ts CA --> HOLogic.mk_setT T;
blanchet@49990
   893
            val collect_map = HOLogic.mk_comp
blanchet@49990
   894
              (mk_collect (map (mk_bnf_t Ts) bnf_sets) defT,
blanchet@49990
   895
              Term.list_comb (mk_bnf_map As' Ts, hs));
blanchet@49990
   896
            val image_collect = mk_collect
blanchet@49990
   897
              (map2 (fn h => fn set => HOLogic.mk_comp (mk_image h, set)) hs bnf_sets_As)
blanchet@49990
   898
              defT;
blanchet@49990
   899
            (*collect {set1 ... setm} o map f1 ... fm = collect {f1` o set1 ... fm` o setm}*)
blanchet@50138
   900
            val goal = fold_rev Logic.all hs (mk_Trueprop_eq (collect_map, image_collect));
blanchet@49990
   901
          in
blanchet@49990
   902
            Skip_Proof.prove lthy [] [] goal
blanchet@49990
   903
              (fn {context = ctxt, ...} => mk_collect_set_natural_tac ctxt (#set_natural axioms))
traytel@50124
   904
            |> Thm.close_derivation
blanchet@49990
   905
          end;
blanchet@49990
   906
blanchet@49990
   907
        val collect_set_natural = mk_lazy mk_collect_set_natural;
blanchet@49990
   908
blanchet@49990
   909
        fun mk_in_mono () =
blanchet@49990
   910
          let
blanchet@49990
   911
            val prems_mono = map2 (HOLogic.mk_Trueprop oo mk_subset) As As_copy;
blanchet@50473
   912
            val in_mono_goal =
blanchet@49990
   913
              fold_rev Logic.all (As @ As_copy)
blanchet@49990
   914
                (Logic.list_implies (prems_mono, HOLogic.mk_Trueprop
blanchet@49990
   915
                  (mk_subset (mk_in As bnf_sets_As CA') (mk_in As_copy bnf_sets_As CA'))));
blanchet@49990
   916
          in
blanchet@50473
   917
            Skip_Proof.prove lthy [] [] in_mono_goal (K (mk_in_mono_tac live))
traytel@50124
   918
            |> Thm.close_derivation
blanchet@49990
   919
          end;
blanchet@49990
   920
blanchet@49990
   921
        val in_mono = mk_lazy mk_in_mono;
blanchet@49990
   922
blanchet@49990
   923
        fun mk_in_cong () =
blanchet@49990
   924
          let
blanchet@49990
   925
            val prems_cong = map2 (HOLogic.mk_Trueprop oo curry HOLogic.mk_eq) As As_copy;
blanchet@50473
   926
            val in_cong_goal =
blanchet@49990
   927
              fold_rev Logic.all (As @ As_copy)
blanchet@49990
   928
                (Logic.list_implies (prems_cong, HOLogic.mk_Trueprop
blanchet@49990
   929
                  (HOLogic.mk_eq (mk_in As bnf_sets_As CA', mk_in As_copy bnf_sets_As CA'))));
blanchet@49990
   930
          in
blanchet@50473
   931
            Skip_Proof.prove lthy [] [] in_cong_goal (K ((TRY o hyp_subst_tac THEN' rtac refl) 1))
traytel@50124
   932
            |> Thm.close_derivation
blanchet@49990
   933
          end;
blanchet@49990
   934
blanchet@49990
   935
        val in_cong = mk_lazy mk_in_cong;
blanchet@49990
   936
blanchet@49990
   937
        val map_id' = mk_lazy (fn () => mk_id' (#map_id axioms));
blanchet@49990
   938
        val map_comp' = mk_lazy (fn () => mk_comp' (#map_comp axioms));
blanchet@49990
   939
blanchet@49990
   940
        val set_natural' =
blanchet@49990
   941
          map (fn thm => mk_lazy (fn () => mk_set_natural' thm)) (#set_natural axioms);
blanchet@49990
   942
blanchet@49990
   943
        fun mk_map_wppull () =
blanchet@49990
   944
          let
blanchet@49990
   945
            val prems = if live = 0 then [] else
blanchet@49990
   946
              [HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
blanchet@49990
   947
                (map8 mk_wpull Xs B1s B2s f1s f2s (map SOME (e1s ~~ e2s)) p1s p2s))];
blanchet@49990
   948
            val CX = mk_bnf_T domTs CA;
blanchet@49990
   949
            val CB1 = mk_bnf_T B1Ts CA;
blanchet@49990
   950
            val CB2 = mk_bnf_T B2Ts CA;
blanchet@49990
   951
            val bnf_sets_CX =
blanchet@49990
   952
              map2 (normalize_set (map (mk_bnf_T domTs) CA_params)) domTs bnf_sets;
blanchet@49990
   953
            val bnf_sets_CB1 =
blanchet@49990
   954
              map2 (normalize_set (map (mk_bnf_T B1Ts) CA_params)) B1Ts bnf_sets;
blanchet@49990
   955
            val bnf_sets_CB2 =
blanchet@49990
   956
              map2 (normalize_set (map (mk_bnf_T B2Ts) CA_params)) B2Ts bnf_sets;
blanchet@49990
   957
            val bnf_map_app_f1 = Term.list_comb (mk_bnf_map B1Ts ranTs, f1s);
blanchet@49990
   958
            val bnf_map_app_f2 = Term.list_comb (mk_bnf_map B2Ts ranTs, f2s);
blanchet@49990
   959
            val bnf_map_app_e1 = Term.list_comb (mk_bnf_map B1Ts ranTs', e1s);
blanchet@49990
   960
            val bnf_map_app_e2 = Term.list_comb (mk_bnf_map B2Ts ranTs'', e2s);
blanchet@49990
   961
            val bnf_map_app_p1 = Term.list_comb (mk_bnf_map domTs B1Ts, p1s);
blanchet@49990
   962
            val bnf_map_app_p2 = Term.list_comb (mk_bnf_map domTs B2Ts, p2s);
blanchet@49990
   963
blanchet@49990
   964
            val concl = mk_wpull (mk_in Xs bnf_sets_CX CX)
blanchet@49990
   965
              (mk_in B1s bnf_sets_CB1 CB1) (mk_in B2s bnf_sets_CB2 CB2)
blanchet@49990
   966
              bnf_map_app_f1 bnf_map_app_f2 (SOME (bnf_map_app_e1, bnf_map_app_e2))
blanchet@49990
   967
              bnf_map_app_p1 bnf_map_app_p2;
blanchet@49990
   968
blanchet@49990
   969
            val goal =
blanchet@49990
   970
              fold_rev Logic.all (Xs @ B1s @ B2s @ f1s @ f2s @ e1s @ e2s @ p1s @ p2s)
blanchet@49990
   971
                (Logic.list_implies (prems, HOLogic.mk_Trueprop concl))
blanchet@49990
   972
          in
blanchet@49990
   973
            Skip_Proof.prove lthy [] [] goal
blanchet@49990
   974
              (fn _ => mk_map_wppull_tac (#map_id axioms) (#map_cong axioms)
blanchet@49990
   975
                (#map_wpull axioms) (Lazy.force map_comp') (map Lazy.force set_natural'))
traytel@50124
   976
            |> Thm.close_derivation
blanchet@49990
   977
          end;
blanchet@49990
   978
blanchet@50475
   979
        val rel_O_Grs = no_refl [#rel_O_Gr axioms];
blanchet@50468
   980
blanchet@49990
   981
        val map_wppull = mk_lazy mk_map_wppull;
blanchet@49990
   982
blanchet@49990
   983
        fun mk_rel_Gr () =
blanchet@49990
   984
          let
blanchet@49990
   985
            val lhs = Term.list_comb (relAsBs, map2 mk_Gr As fs);
blanchet@49990
   986
            val rhs = mk_Gr (mk_in As bnf_sets_As CA') (Term.list_comb (bnf_map_AsBs, fs));
blanchet@50138
   987
            val goal = fold_rev Logic.all (As @ fs) (mk_Trueprop_eq (lhs, rhs));
blanchet@49990
   988
          in
blanchet@49990
   989
            Skip_Proof.prove lthy [] [] goal
blanchet@50475
   990
              (mk_rel_Gr_tac rel_O_Grs (#map_id axioms) (#map_cong axioms)
blanchet@49990
   991
                (#map_wpull axioms) (Lazy.force in_cong) (Lazy.force map_id')
blanchet@49990
   992
                (Lazy.force map_comp') (map Lazy.force set_natural'))
traytel@50124
   993
            |> Thm.close_derivation
blanchet@49990
   994
          end;
blanchet@49990
   995
blanchet@49990
   996
        val rel_Gr = mk_lazy mk_rel_Gr;
blanchet@49990
   997
blanchet@49990
   998
        fun mk_rel_prems f = map2 (HOLogic.mk_Trueprop oo f) Rs Rs_copy
blanchet@49990
   999
        fun mk_rel_concl f = HOLogic.mk_Trueprop
blanchet@49990
  1000
          (f (Term.list_comb (relAsBs, Rs), Term.list_comb (relAsBs, Rs_copy)));
blanchet@49990
  1001
blanchet@49990
  1002
        fun mk_rel_mono () =
blanchet@49990
  1003
          let
blanchet@49990
  1004
            val mono_prems = mk_rel_prems mk_subset;
blanchet@49990
  1005
            val mono_concl = mk_rel_concl (uncurry mk_subset);
blanchet@49990
  1006
          in
blanchet@49990
  1007
            Skip_Proof.prove lthy [] []
blanchet@49990
  1008
              (fold_rev Logic.all (Rs @ Rs_copy) (Logic.list_implies (mono_prems, mono_concl)))
blanchet@50475
  1009
              (mk_rel_mono_tac rel_O_Grs (Lazy.force in_mono))
traytel@50124
  1010
            |> Thm.close_derivation
blanchet@49990
  1011
          end;
blanchet@49990
  1012
blanchet@49990
  1013
        fun mk_rel_cong () =
blanchet@49990
  1014
          let
blanchet@49990
  1015
            val cong_prems = mk_rel_prems (curry HOLogic.mk_eq);
blanchet@49990
  1016
            val cong_concl = mk_rel_concl HOLogic.mk_eq;
blanchet@49990
  1017
          in
blanchet@49990
  1018
            Skip_Proof.prove lthy [] []
blanchet@49990
  1019
              (fold_rev Logic.all (Rs @ Rs_copy) (Logic.list_implies (cong_prems, cong_concl)))
blanchet@49990
  1020
              (fn _ => (TRY o hyp_subst_tac THEN' rtac refl) 1)
traytel@50124
  1021
            |> Thm.close_derivation
blanchet@49990
  1022
          end;
blanchet@49990
  1023
blanchet@49990
  1024
        val rel_mono = mk_lazy mk_rel_mono;
blanchet@49990
  1025
        val rel_cong = mk_lazy mk_rel_cong;
blanchet@49990
  1026
blanchet@49990
  1027
        fun mk_rel_Id () =
blanchet@49990
  1028
          let val relAsAs = mk_bnf_rel self_setRTs CA' CA' in
blanchet@49990
  1029
            Skip_Proof.prove lthy [] []
blanchet@49990
  1030
              (HOLogic.mk_Trueprop
blanchet@49990
  1031
                (HOLogic.mk_eq (Term.list_comb (relAsAs, map Id_const As'), Id_const CA')))
blanchet@49990
  1032
              (mk_rel_Id_tac live (Lazy.force rel_Gr) (#map_id axioms))
traytel@50124
  1033
            |> Thm.close_derivation
blanchet@49990
  1034
          end;
blanchet@49990
  1035
blanchet@49990
  1036
        val rel_Id = mk_lazy mk_rel_Id;
blanchet@49990
  1037
blanchet@49990
  1038
        fun mk_rel_converse () =
blanchet@49990
  1039
          let
blanchet@49990
  1040
            val relBsAs = mk_bnf_rel setRT's CB' CA';
blanchet@49990
  1041
            val lhs = Term.list_comb (relBsAs, map mk_converse Rs);
blanchet@49990
  1042
            val rhs = mk_converse (Term.list_comb (relAsBs, Rs));
blanchet@49990
  1043
            val le_goal = fold_rev Logic.all Rs (HOLogic.mk_Trueprop (mk_subset lhs rhs));
blanchet@49990
  1044
            val le_thm = Skip_Proof.prove lthy [] [] le_goal
blanchet@50475
  1045
              (mk_rel_converse_le_tac rel_O_Grs (Lazy.force rel_Id) (#map_cong axioms)
blanchet@49990
  1046
                (Lazy.force map_comp') (map Lazy.force set_natural'))
traytel@50124
  1047
              |> Thm.close_derivation
blanchet@50138
  1048
            val goal = fold_rev Logic.all Rs (mk_Trueprop_eq (lhs, rhs));
blanchet@49990
  1049
          in
blanchet@49990
  1050
            Skip_Proof.prove lthy [] [] goal (fn _ => mk_rel_converse_tac le_thm)
traytel@50124
  1051
            |> Thm.close_derivation
blanchet@49990
  1052
          end;
blanchet@49990
  1053
blanchet@49990
  1054
        val rel_converse = mk_lazy mk_rel_converse;
blanchet@49990
  1055
blanchet@49990
  1056
        fun mk_rel_O () =
blanchet@49990
  1057
          let
blanchet@49990
  1058
            val relAsCs = mk_bnf_rel setRTsAsCs CA' CC';
blanchet@49990
  1059
            val relBsCs = mk_bnf_rel setRTsBsCs CB' CC';
blanchet@49990
  1060
            val lhs = Term.list_comb (relAsCs, map2 (curry mk_rel_comp) Rs Ss);
blanchet@49990
  1061
            val rhs = mk_rel_comp (Term.list_comb (relAsBs, Rs), Term.list_comb (relBsCs, Ss));
blanchet@50138
  1062
            val goal = fold_rev Logic.all (Rs @ Ss) (mk_Trueprop_eq (lhs, rhs));
blanchet@49990
  1063
          in
blanchet@49990
  1064
            Skip_Proof.prove lthy [] [] goal
blanchet@50475
  1065
              (mk_rel_O_tac rel_O_Grs (Lazy.force rel_Id) (#map_cong axioms)
blanchet@49990
  1066
                (Lazy.force map_wppull) (Lazy.force map_comp') (map Lazy.force set_natural'))
traytel@50124
  1067
            |> Thm.close_derivation
blanchet@49990
  1068
          end;
blanchet@49990
  1069
blanchet@49990
  1070
        val rel_O = mk_lazy mk_rel_O;
blanchet@49990
  1071
blanchet@49990
  1072
        fun mk_in_rel () =
blanchet@49990
  1073
          let
blanchet@49990
  1074
            val bnf_in = mk_in Rs (map (mk_bnf_t RTs) bnf_sets) CRs';
blanchet@49990
  1075
            val map1 = Term.list_comb (mk_bnf_map RTs As', map fst_const RTs);
blanchet@49990
  1076
            val map2 = Term.list_comb (mk_bnf_map RTs Bs', map snd_const RTs);
blanchet@49990
  1077
            val map_fst_eq = HOLogic.mk_eq (map1 $ z, x);
blanchet@49990
  1078
            val map_snd_eq = HOLogic.mk_eq (map2 $ z, y);
blanchet@49990
  1079
            val lhs = HOLogic.mk_mem (HOLogic.mk_prod (x, y), Term.list_comb (relAsBs, Rs));
blanchet@49990
  1080
            val rhs =
blanchet@49990
  1081
              HOLogic.mk_exists (fst z', snd z', HOLogic.mk_conj (HOLogic.mk_mem (z, bnf_in),
blanchet@49990
  1082
                HOLogic.mk_conj (map_fst_eq, map_snd_eq)));
blanchet@49990
  1083
            val goal =
blanchet@50138
  1084
              fold_rev Logic.all (x :: y :: Rs) (mk_Trueprop_eq (lhs, rhs));
blanchet@49990
  1085
          in
blanchet@50475
  1086
            Skip_Proof.prove lthy [] [] goal (mk_in_rel_tac rel_O_Grs (length bnf_sets))
traytel@50124
  1087
            |> Thm.close_derivation
blanchet@49990
  1088
          end;
blanchet@49990
  1089
blanchet@49990
  1090
        val in_rel = mk_lazy mk_in_rel;
blanchet@49990
  1091
blanchet@50467
  1092
        val defs = mk_defs bnf_map_def bnf_set_defs bnf_rel_def bnf_pred_def;
blanchet@49990
  1093
blanchet@49990
  1094
        val facts = mk_facts bd_Card_order bd_Cinfinite bd_Cnotzero collect_set_natural
blanchet@49990
  1095
          in_cong in_mono in_rel map_comp' map_id' map_wppull
blanchet@50468
  1096
          rel_cong rel_mono rel_Id rel_Gr rel_converse rel_O set_natural';
blanchet@49990
  1097
blanchet@49990
  1098
        val wits = map2 mk_witness bnf_wits wit_thms;
blanchet@49990
  1099
blanchet@49990
  1100
        val bnf_rel = Term.subst_atomic_types
blanchet@49990
  1101
          ((Ds ~~ deads) @ (As' ~~ alphas) @ (Bs' ~~ betas)) relAsBs;
blanchet@49990
  1102
        val bnf_pred = Term.subst_atomic_types
blanchet@49990
  1103
          ((Ds ~~ deads) @ (As' ~~ alphas) @ (Bs' ~~ betas)) pred;
blanchet@49990
  1104
blanchet@49990
  1105
        val bnf = mk_bnf b CA live alphas betas dead deads bnf_map bnf_sets bnf_bd axioms defs facts
blanchet@49990
  1106
          wits bnf_rel bnf_pred;
blanchet@49990
  1107
      in
blanchet@49990
  1108
        (bnf, lthy
blanchet@49990
  1109
          |> (if fact_policy = Note_All_Facts_and_Axioms then
blanchet@49990
  1110
                let
blanchet@49990
  1111
                  val witNs = if length wits = 1 then [witN] else map mk_witN (1 upto length wits);
traytel@50124
  1112
                  val notes =
traytel@50124
  1113
                    [(bd_card_orderN, [#bd_card_order axioms]),
traytel@50124
  1114
                    (bd_cinfiniteN, [#bd_cinfinite axioms]),
traytel@50124
  1115
                    (bd_Card_orderN, [#bd_Card_order facts]),
traytel@50124
  1116
                    (bd_CinfiniteN, [#bd_Cinfinite facts]),
traytel@50124
  1117
                    (bd_CnotzeroN, [#bd_Cnotzero facts]),
traytel@50124
  1118
                    (collect_set_naturalN, [Lazy.force (#collect_set_natural facts)]),
traytel@50124
  1119
                    (in_bdN, [#in_bd axioms]),
traytel@50124
  1120
                    (in_monoN, [Lazy.force (#in_mono facts)]),
traytel@50124
  1121
                    (in_relN, [Lazy.force (#in_rel facts)]),
traytel@50124
  1122
                    (map_compN, [#map_comp axioms]),
traytel@50124
  1123
                    (map_idN, [#map_id axioms]),
traytel@50124
  1124
                    (map_wpullN, [#map_wpull axioms]),
traytel@50124
  1125
                    (set_naturalN, #set_natural axioms),
traytel@50124
  1126
                    (set_bdN, #set_bd axioms)] @
traytel@50124
  1127
                    map2 pair witNs wit_thms
traytel@50124
  1128
                    |> map (fn (thmN, thms) =>
traytel@50124
  1129
                      ((qualify (Binding.qualify true (Binding.name_of b) (Binding.name thmN)), []),
traytel@50124
  1130
                      [(thms, [])]));
blanchet@49990
  1131
                in
traytel@50124
  1132
                  Local_Theory.notes notes #> snd
blanchet@49990
  1133
                end
blanchet@49990
  1134
              else
blanchet@49990
  1135
                I)
blanchet@49990
  1136
          |> (if fact_policy = Note_All_Facts_and_Axioms orelse
blanchet@49990
  1137
                 fact_policy = Derive_All_Facts_Note_Most then
traytel@50124
  1138
                let
traytel@50124
  1139
                  val notes =
traytel@50124
  1140
                    [(map_congN, [#map_cong axioms]),
blanchet@50475
  1141
                    (rel_O_GrN, rel_O_Grs),
traytel@50124
  1142
                    (rel_IdN, [Lazy.force (#rel_Id facts)]),
traytel@50124
  1143
                    (rel_GrN, [Lazy.force (#rel_Gr facts)]),
traytel@50124
  1144
                    (rel_converseN, [Lazy.force (#rel_converse facts)]),
traytel@50401
  1145
                    (rel_monoN, [Lazy.force (#rel_mono facts)]),
traytel@50124
  1146
                    (rel_ON, [Lazy.force (#rel_O facts)]),
traytel@50124
  1147
                    (map_id'N, [Lazy.force (#map_id' facts)]),
traytel@50124
  1148
                    (map_comp'N, [Lazy.force (#map_comp' facts)]),
traytel@50124
  1149
                    (set_natural'N, map Lazy.force (#set_natural' facts))]
blanchet@50475
  1150
                    |> filter_out (null o #2)
traytel@50124
  1151
                    |> map (fn (thmN, thms) =>
traytel@50124
  1152
                      ((qualify (Binding.qualify true (Binding.name_of b) (Binding.name thmN)), []),
traytel@50124
  1153
                      [(thms, [])]));
traytel@50124
  1154
                in
traytel@50124
  1155
                  Local_Theory.notes notes #> snd
traytel@50124
  1156
                end
blanchet@49990
  1157
              else
blanchet@49990
  1158
                I))
blanchet@49990
  1159
      end;
blanchet@50474
  1160
blanchet@50474
  1161
    val one_step_defs =
blanchet@50475
  1162
      no_reflexive (bnf_map_def :: bnf_bd_def :: bnf_set_defs @ bnf_wit_defs @ [bnf_rel_def]);
blanchet@49990
  1163
  in
blanchet@50474
  1164
    (key, goals, wit_goalss, after_qed, lthy, one_step_defs)
blanchet@49990
  1165
  end;
blanchet@49990
  1166
traytel@50449
  1167
fun register_bnf key (bnf, lthy) =
traytel@50449
  1168
  (bnf, Local_Theory.declaration {syntax = false, pervasive = true}
traytel@50449
  1169
    (fn phi => Data.map (Symtab.update_new (key, morph_bnf phi bnf))) lthy);
traytel@50449
  1170
blanchet@50471
  1171
(* TODO: Once the invariant "nwits > 0" holds, remove "mk_conjunction_balanced'" and "rtac TrueI"
blanchet@50471
  1172
   below *)
blanchet@50471
  1173
fun mk_conjunction_balanced' [] = @{prop True}
blanchet@50471
  1174
  | mk_conjunction_balanced' ts = Logic.mk_conjunction_balanced ts;
blanchet@50471
  1175
blanchet@50033
  1176
fun bnf_def const_policy fact_policy qualify tacs wit_tac Ds =
traytel@50449
  1177
  (fn (_, goals, wit_goalss, after_qed, lthy, defs) =>
blanchet@49990
  1178
  let
blanchet@50471
  1179
    val wits_tac =
blanchet@50471
  1180
      K (TRYALL Goal.conjunction_tac) THEN' K (TRYALL (rtac TrueI)) THEN'
blanchet@50471
  1181
      unfold_defs_tac lthy defs wit_tac;
blanchet@50471
  1182
    val wit_goals = map mk_conjunction_balanced' wit_goalss;
blanchet@49990
  1183
    val wit_thms =
blanchet@50471
  1184
      Skip_Proof.prove lthy [] [] (mk_conjunction_balanced' wit_goals) wits_tac
blanchet@49990
  1185
      |> Conjunction.elim_balanced (length wit_goals)
blanchet@49990
  1186
      |> map2 (Conjunction.elim_balanced o length) wit_goalss
blanchet@50471
  1187
      |> map (map (Thm.close_derivation o Thm.forall_elim_vars 0));
blanchet@49990
  1188
  in
blanchet@50126
  1189
    map2 (Thm.close_derivation oo Skip_Proof.prove lthy [] [])
blanchet@50126
  1190
      goals (map (unfold_defs_tac lthy defs) tacs)
blanchet@49990
  1191
    |> (fn thms => after_qed (map single thms @ wit_thms) lthy)
blanchet@50294
  1192
  end) oo prepare_def const_policy fact_policy qualify (K I) Ds;
blanchet@49990
  1193
traytel@50449
  1194
val bnf_def_cmd = (fn (key, goals, wit_goals, after_qed, lthy, defs) =>
blanchet@49990
  1195
  Proof.unfolding ([[(defs, [])]])
traytel@50449
  1196
    (Proof.theorem NONE (snd o register_bnf key oo after_qed)
blanchet@49990
  1197
      (map (single o rpair []) goals @ map (map (rpair [])) wit_goals) lthy)) oo
traytel@50450
  1198
  prepare_def Do_Inline (user_policy Derive_All_Facts_Note_Most) I Syntax.read_term NONE;
blanchet@49990
  1199
blanchet@49990
  1200
fun print_bnfs ctxt =
blanchet@49990
  1201
  let
blanchet@49990
  1202
    fun pretty_set sets i = Pretty.block
blanchet@49990
  1203
      [Pretty.str (mk_setN (i + 1) ^ ":"), Pretty.brk 1,
blanchet@49990
  1204
          Pretty.quote (Syntax.pretty_term ctxt (nth sets i))];
blanchet@49990
  1205
blanchet@49990
  1206
    fun pretty_bnf (key, BNF {T = T, map = map, sets = sets, bd = bd,
blanchet@49990
  1207
      live = live, lives = lives, dead = dead, deads = deads, ...}) =
blanchet@49990
  1208
      Pretty.big_list
blanchet@49990
  1209
        (Pretty.string_of (Pretty.block [Pretty.str key, Pretty.str ":", Pretty.brk 1,
blanchet@49990
  1210
          Pretty.quote (Syntax.pretty_typ ctxt T)]))
blanchet@49990
  1211
        ([Pretty.block [Pretty.str "live:", Pretty.brk 1, Pretty.str (string_of_int live),
blanchet@49990
  1212
            Pretty.brk 3, Pretty.list "[" "]" (List.map (Syntax.pretty_typ ctxt) lives)],
blanchet@49990
  1213
          Pretty.block [Pretty.str "dead:", Pretty.brk 1, Pretty.str (string_of_int dead),
blanchet@49990
  1214
            Pretty.brk 3, Pretty.list "[" "]" (List.map (Syntax.pretty_typ ctxt) deads)],
blanchet@49990
  1215
          Pretty.block [Pretty.str (mapN ^ ":"), Pretty.brk 1,
blanchet@49990
  1216
            Pretty.quote (Syntax.pretty_term ctxt map)]] @
blanchet@49990
  1217
          List.map (pretty_set sets) (0 upto length sets - 1) @
blanchet@49990
  1218
          [Pretty.block [Pretty.str (bdN ^ ":"), Pretty.brk 1,
blanchet@49990
  1219
            Pretty.quote (Syntax.pretty_term ctxt bd)]]);
blanchet@49990
  1220
  in
blanchet@49990
  1221
    Pretty.big_list "BNFs:" (map pretty_bnf (Symtab.dest (Data.get (Context.Proof ctxt))))
blanchet@49990
  1222
    |> Pretty.writeln
blanchet@49990
  1223
  end;
blanchet@49990
  1224
blanchet@49990
  1225
val _ =
blanchet@49990
  1226
  Outer_Syntax.improper_command @{command_spec "print_bnfs"} "print all BNFs"
blanchet@49990
  1227
    (Scan.succeed (Toplevel.keep (print_bnfs o Toplevel.context_of)));
blanchet@49990
  1228
blanchet@49990
  1229
val _ =
blanchet@49990
  1230
  Outer_Syntax.local_theory_to_proof @{command_spec "bnf_def"} "define a BNF for an existing type"
traytel@50449
  1231
    ((parse_opt_binding_colon -- Parse.term --
blanchet@50292
  1232
       (@{keyword "["} |-- Parse.list Parse.term --| @{keyword "]"}) -- Parse.term --
blanchet@50474
  1233
       (@{keyword "["} |-- Parse.list Parse.term --| @{keyword "]"}) -- Scan.option Parse.term)
blanchet@50474
  1234
       >> bnf_def_cmd);
blanchet@49990
  1235
blanchet@49990
  1236
end;