src/HOL/Tools/BNF/bnf_def.ML
author wenzelm
Tue, 23 May 2023 18:46:15 +0200
changeset 78095 bc42c074e58f
parent 76383 fc35dc967344
child 79855 3713ca49e32c
permissions -rw-r--r--
tuned signature: more position information;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
55061
a0adf838e2d1 adjusted comments
blanchet
parents: 55060
diff changeset
     1
(*  Title:      HOL/Tools/BNF/bnf_def.ML
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     2
    Author:     Dmitriy Traytel, TU Muenchen
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     3
    Author:     Jasmin Blanchette, TU Muenchen
57668
blanchet
parents: 57632
diff changeset
     4
    Author:     Martin Desharnais, TU Muenchen
75624
22d1c5f2b9f4 strict bounds for BNFs (by Jan van Brügge)
traytel
parents: 75276
diff changeset
     5
    Author:     Jan van Brügge, TU Muenchen
22d1c5f2b9f4 strict bounds for BNFs (by Jan van Brügge)
traytel
parents: 75276
diff changeset
     6
    Copyright   2012, 2013, 2014, 2022
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     7
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     8
Definition of bounded natural functors.
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     9
*)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    10
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    11
signature BNF_DEF =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    12
sig
51837
087498724486 lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
blanchet
parents: 51836
diff changeset
    13
  type bnf
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    14
  type nonemptiness_witness = {I: int list, wit: term, prop: thm list}
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    15
51837
087498724486 lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
blanchet
parents: 51836
diff changeset
    16
  val morph_bnf: morphism -> bnf -> bnf
56016
8875cdcfc85b unfold intermediate definitions after sealing the bnf
traytel
parents: 55945
diff changeset
    17
  val morph_bnf_defs: morphism -> bnf -> bnf
64440
0d31d1735104 tuned signature
traytel
parents: 63862
diff changeset
    18
  val permute_deads: (typ list -> typ list) -> bnf -> bnf
58175
2412a3369c30 tuned BNF database handling
blanchet
parents: 58116
diff changeset
    19
  val transfer_bnf: theory -> bnf -> bnf
51837
087498724486 lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
blanchet
parents: 51836
diff changeset
    20
  val bnf_of: Proof.context -> string -> bnf option
58116
1a9ac371e5a0 added theory-based getters for convenience
blanchet
parents: 58108
diff changeset
    21
  val bnf_of_global: theory -> string -> bnf option
58659
6c9821c32dd5 Local_Interpretation is superseded by Plugin with formal Plugin_Name management, avoiding undeclared strings;
wenzelm
parents: 58634
diff changeset
    22
  val bnf_interpretation: string -> (bnf -> local_theory -> local_theory) -> theory -> theory
58188
cc71d2be4f0a introduced mechanism to filter interpretations
blanchet
parents: 58187
diff changeset
    23
  val interpret_bnf: (string -> bool) -> bnf -> local_theory -> local_theory
58177
166131276380 introduced local interpretation mechanism for BNFs, to solve issues with datatypes in locales
blanchet
parents: 58175
diff changeset
    24
  val register_bnf_raw: string -> bnf -> local_theory -> local_theory
58188
cc71d2be4f0a introduced mechanism to filter interpretations
blanchet
parents: 58187
diff changeset
    25
  val register_bnf: (string -> bool) -> string -> bnf -> local_theory -> local_theory
49434
433dc7e028c8 separated registration of BNFs from bnf_def (BNFs are now stored only for bnf_def and (co)data commands)
traytel
parents: 49430
diff changeset
    26
51837
087498724486 lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
blanchet
parents: 51836
diff changeset
    27
  val name_of_bnf: bnf -> binding
087498724486 lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
blanchet
parents: 51836
diff changeset
    28
  val T_of_bnf: bnf -> typ
087498724486 lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
blanchet
parents: 51836
diff changeset
    29
  val live_of_bnf: bnf -> int
087498724486 lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
blanchet
parents: 51836
diff changeset
    30
  val lives_of_bnf: bnf -> typ list
087498724486 lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
blanchet
parents: 51836
diff changeset
    31
  val dead_of_bnf: bnf -> int
087498724486 lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
blanchet
parents: 51836
diff changeset
    32
  val deads_of_bnf: bnf -> typ list
56346
42533f8f4729 added BNF interpretation hook
blanchet
parents: 56016
diff changeset
    33
  val bd_of_bnf: bnf -> term
51837
087498724486 lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
blanchet
parents: 51836
diff changeset
    34
  val nwits_of_bnf: bnf -> int
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    35
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    36
  val mapN: string
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
    37
  val predN: string
49507
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49506
diff changeset
    38
  val relN: string
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    39
  val setN: string
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    40
  val mk_setN: int -> string
55025
1ac0a0194428 support declaration of nonemptiness witnesses in bnf_decl
traytel
parents: 54921
diff changeset
    41
  val mk_witN: int -> string
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    42
51837
087498724486 lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
blanchet
parents: 51836
diff changeset
    43
  val map_of_bnf: bnf -> term
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
    44
  val pred_of_bnf: bnf -> term
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
    45
  val rel_of_bnf: bnf -> term
51837
087498724486 lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
blanchet
parents: 51836
diff changeset
    46
  val sets_of_bnf: bnf -> term list
49214
2a3cb4c71b87 construct the right iterator theorem in the recursive case
blanchet
parents: 49123
diff changeset
    47
51837
087498724486 lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
blanchet
parents: 51836
diff changeset
    48
  val mk_T_of_bnf: typ list -> typ list -> bnf -> typ
087498724486 lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
blanchet
parents: 51836
diff changeset
    49
  val mk_bd_of_bnf: typ list -> typ list -> bnf -> term
087498724486 lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
blanchet
parents: 51836
diff changeset
    50
  val mk_map_of_bnf: typ list -> typ list -> typ list -> bnf -> term
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
    51
  val mk_pred_of_bnf: typ list -> typ list -> bnf -> term
51837
087498724486 lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
blanchet
parents: 51836
diff changeset
    52
  val mk_rel_of_bnf: typ list -> typ list -> typ list -> bnf -> term
087498724486 lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
blanchet
parents: 51836
diff changeset
    53
  val mk_sets_of_bnf: typ list list -> typ list list -> bnf -> term list
087498724486 lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
blanchet
parents: 51836
diff changeset
    54
  val mk_wits_of_bnf: typ list list -> typ list list -> bnf -> (int list * term) list
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    55
51837
087498724486 lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
blanchet
parents: 51836
diff changeset
    56
  val bd_Card_order_of_bnf: bnf -> thm
087498724486 lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
blanchet
parents: 51836
diff changeset
    57
  val bd_Cinfinite_of_bnf: bnf -> thm
087498724486 lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
blanchet
parents: 51836
diff changeset
    58
  val bd_Cnotzero_of_bnf: bnf -> thm
087498724486 lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
blanchet
parents: 51836
diff changeset
    59
  val bd_card_order_of_bnf: bnf -> thm
087498724486 lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
blanchet
parents: 51836
diff changeset
    60
  val bd_cinfinite_of_bnf: bnf -> thm
75624
22d1c5f2b9f4 strict bounds for BNFs (by Jan van Brügge)
traytel
parents: 75276
diff changeset
    61
  val bd_regularCard_of_bnf: bnf -> thm
51837
087498724486 lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
blanchet
parents: 51836
diff changeset
    62
  val collect_set_map_of_bnf: bnf -> thm
087498724486 lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
blanchet
parents: 51836
diff changeset
    63
  val in_bd_of_bnf: bnf -> thm
087498724486 lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
blanchet
parents: 51836
diff changeset
    64
  val in_cong_of_bnf: bnf -> thm
087498724486 lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
blanchet
parents: 51836
diff changeset
    65
  val in_mono_of_bnf: bnf -> thm
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
    66
  val in_rel_of_bnf: bnf -> thm
56635
b07c8ad23010 added 'inj_map' as auxiliary BNF theorem
blanchet
parents: 56523
diff changeset
    67
  val inj_map_of_bnf: bnf -> thm
57970
eaa986cd285a generate 'inj_map_strong' for BNFs
desharna
parents: 57969
diff changeset
    68
  val inj_map_strong_of_bnf: bnf -> thm
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
    69
  val le_rel_OO_of_bnf: bnf -> thm
53287
271b34513bfb renamed BNF axiom
blanchet
parents: 53285
diff changeset
    70
  val map_comp0_of_bnf: bnf -> thm
53288
050d0bc9afa5 renamed BNF fact
blanchet
parents: 53287
diff changeset
    71
  val map_comp_of_bnf: bnf -> thm
51837
087498724486 lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
blanchet
parents: 51836
diff changeset
    72
  val map_cong0_of_bnf: bnf -> thm
087498724486 lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
blanchet
parents: 51836
diff changeset
    73
  val map_cong_of_bnf: bnf -> thm
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
    74
  val map_cong_pred_of_bnf: bnf -> thm
57981
81baacfd56e8 generate 'map_cong_simp' for BNFs
desharna
parents: 57970
diff changeset
    75
  val map_cong_simp_of_bnf: bnf -> thm
51837
087498724486 lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
blanchet
parents: 51836
diff changeset
    76
  val map_def_of_bnf: bnf -> thm
53270
c8628119d18e renamed BNF axiom
blanchet
parents: 53265
diff changeset
    77
  val map_id0_of_bnf: bnf -> thm
53285
f09645642984 renamed BNF fact
blanchet
parents: 53270
diff changeset
    78
  val map_id_of_bnf: bnf -> thm
57399
cfc19f0a6261 compile
blanchet
parents: 57303
diff changeset
    79
  val map_ident0_of_bnf: bnf -> thm
56903
69b6369a98fa generate 'map_ident' theorem for BNFs
desharna
parents: 56766
diff changeset
    80
  val map_ident_of_bnf: bnf -> thm
52719
480a3479fa47 transfer rule for map (not yet registered as a transfer rule)
traytel
parents: 52635
diff changeset
    81
  val map_transfer_of_bnf: bnf -> thm
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
    82
  val pred_cong0_of_bnf: bnf -> thm
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
    83
  val pred_cong_of_bnf: bnf -> thm
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
    84
  val pred_cong_simp_of_bnf: bnf -> thm
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
    85
  val pred_def_of_bnf: bnf -> thm
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
    86
  val pred_map_of_bnf: bnf -> thm
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
    87
  val pred_mono_strong0_of_bnf: bnf -> thm
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
    88
  val pred_mono_strong_of_bnf: bnf -> thm
63714
b62f4f765353 derive pred_mono property for BNFs
traytel
parents: 63399
diff changeset
    89
  val pred_mono_of_bnf: bnf -> thm
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
    90
  val pred_set_of_bnf: bnf -> thm
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
    91
  val pred_rel_of_bnf: bnf -> thm
62329
9f7fa08d2307 derive transfer rule for predicator
traytel
parents: 62326
diff changeset
    92
  val pred_transfer_of_bnf: bnf -> thm
9f7fa08d2307 derive transfer rule for predicator
traytel
parents: 62326
diff changeset
    93
  val pred_True_of_bnf: bnf -> thm
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
    94
  val rel_Grp_of_bnf: bnf -> thm
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
    95
  val rel_OO_Grp_of_bnf: bnf -> thm
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
    96
  val rel_OO_of_bnf: bnf -> thm
61242
1f92b0ac9c96 congruence rules for the relator
traytel
parents: 61241
diff changeset
    97
  val rel_cong0_of_bnf: bnf -> thm
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
    98
  val rel_cong_of_bnf: bnf -> thm
61242
1f92b0ac9c96 congruence rules for the relator
traytel
parents: 61241
diff changeset
    99
  val rel_cong_simp_of_bnf: bnf -> thm
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
   100
  val rel_conversep_of_bnf: bnf -> thm
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   101
  val rel_def_of_bnf: bnf -> thm
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   102
  val rel_eq_of_bnf: bnf -> thm
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   103
  val rel_flip_of_bnf: bnf -> thm
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   104
  val rel_map_of_bnf: bnf -> thm list
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
   105
  val rel_mono_of_bnf: bnf -> thm
57967
e6d2e998c30f renamed 'rel_mono_strong' to 'rel_mono_strong0'
desharna
parents: 57932
diff changeset
   106
  val rel_mono_strong0_of_bnf: bnf -> thm
57968
00e9c6d367e7 generate property 'rel_mono_strong' for BNFs
desharna
parents: 57967
diff changeset
   107
  val rel_mono_strong_of_bnf: bnf -> thm
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   108
  val rel_eq_onp_of_bnf: bnf -> thm
59726
64c2bb331035 BNF relators preserve reflexivity
traytel
parents: 59663
diff changeset
   109
  val rel_refl_of_bnf: bnf -> thm
61240
464c5da3f508 more useful properties of the relators
traytel
parents: 61116
diff changeset
   110
  val rel_refl_strong_of_bnf: bnf -> thm
464c5da3f508 more useful properties of the relators
traytel
parents: 61116
diff changeset
   111
  val rel_reflp_of_bnf: bnf -> thm
464c5da3f508 more useful properties of the relators
traytel
parents: 61116
diff changeset
   112
  val rel_symp_of_bnf: bnf -> thm
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   113
  val rel_transfer_of_bnf: bnf -> thm
61240
464c5da3f508 more useful properties of the relators
traytel
parents: 61116
diff changeset
   114
  val rel_transp_of_bnf: bnf -> thm
51837
087498724486 lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
blanchet
parents: 51836
diff changeset
   115
  val set_bd_of_bnf: bnf -> thm list
087498724486 lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
blanchet
parents: 51836
diff changeset
   116
  val set_defs_of_bnf: bnf -> thm list
53289
5e0623448bdb renamed BNF axiom
blanchet
parents: 53288
diff changeset
   117
  val set_map0_of_bnf: bnf -> thm list
53290
b6c3be868217 renamed BNF fact
blanchet
parents: 53289
diff changeset
   118
  val set_map_of_bnf: bnf -> thm list
58106
c8cba801c483 generate 'set_transfer' for BNFs
desharna
parents: 58104
diff changeset
   119
  val set_transfer_of_bnf: bnf -> thm list
51837
087498724486 lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
blanchet
parents: 51836
diff changeset
   120
  val wit_thms_of_bnf: bnf -> thm list
087498724486 lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
blanchet
parents: 51836
diff changeset
   121
  val wit_thmss_of_bnf: bnf -> thm list list
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   122
54236
e00009523727 moved code around
blanchet
parents: 54189
diff changeset
   123
  val mk_map: int -> typ list -> typ list -> term -> term
62326
3cf7a067599c allow predicator instead of map function in 'primrec'
blanchet
parents: 62324
diff changeset
   124
  val mk_pred: typ list -> term -> term
54236
e00009523727 moved code around
blanchet
parents: 54189
diff changeset
   125
  val mk_rel: int -> typ list -> typ list -> term -> term
63851
1a1fd3f3a24c prove 'set' property backward
blanchet
parents: 63714
diff changeset
   126
  val mk_set: typ list -> term -> term
64627
8d7cb22482e3 generalized ML function (towards nonuniform datatypes)
blanchet
parents: 64440
diff changeset
   127
  val build_map: Proof.context -> typ list -> typ list -> (typ * typ -> term) -> typ * typ -> term
8d7cb22482e3 generalized ML function (towards nonuniform datatypes)
blanchet
parents: 64440
diff changeset
   128
  val build_rel: (string * (int * term)) list -> Proof.context -> typ list -> typ list ->
8d7cb22482e3 generalized ML function (towards nonuniform datatypes)
blanchet
parents: 64440
diff changeset
   129
    (typ * typ -> term) -> typ * typ -> term
63851
1a1fd3f3a24c prove 'set' property backward
blanchet
parents: 63714
diff changeset
   130
  val build_set: Proof.context -> typ -> typ -> term
54246
8fdb4dc08ed1 split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
blanchet
parents: 54237
diff changeset
   131
  val flatten_type_args_of_bnf: bnf -> 'a -> 'a list -> 'a list
8fdb4dc08ed1 split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
blanchet
parents: 54237
diff changeset
   132
  val map_flattened_map_args: Proof.context -> string -> (term list -> 'a list) -> term list ->
8fdb4dc08ed1 split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
blanchet
parents: 54237
diff changeset
   133
    'a list
54236
e00009523727 moved code around
blanchet
parents: 54189
diff changeset
   134
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   135
  val mk_witness: int list * term -> thm list -> nonemptiness_witness
60918
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents: 60801
diff changeset
   136
  val mk_wit_goals: term list -> term list -> term list -> int list * term -> term list
49103
3caaa80f53a4 generalized signature
traytel
parents: 49021
diff changeset
   137
  val minimize_wits: (''a list * 'b) list -> (''a list * 'b) list
51837
087498724486 lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
blanchet
parents: 51836
diff changeset
   138
  val wits_of_bnf: bnf -> nonemptiness_witness list
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   139
75624
22d1c5f2b9f4 strict bounds for BNFs (by Jan van Brügge)
traytel
parents: 75276
diff changeset
   140
  val zip_axioms: 'a -> 'a -> 'a -> 'a list -> 'a -> 'a -> 'a -> 'a list -> 'a -> 'a -> 'a -> 'a list
49456
fa8302c8dea1 adapted BNF composition to new relator approach
blanchet
parents: 49455
diff changeset
   141
55854
ee270328a781 make 'typedef' optional, depending on size of original type
blanchet
parents: 55480
diff changeset
   142
  datatype inline_policy = Dont_Inline | Hardly_Inline | Smart_Inline | Do_Inline
49538
c90818b63599 simplified fact policies
blanchet
parents: 49537
diff changeset
   143
  datatype fact_policy = Dont_Note | Note_Some | Note_All
c90818b63599 simplified fact policies
blanchet
parents: 49537
diff changeset
   144
62093
bd73a2279fcd more uniform treatment of package internals;
wenzelm
parents: 61841
diff changeset
   145
  val bnf_internals: bool Config.T
53143
ba80154a1118 configuration option to control timing output for (co)datatypes
traytel
parents: 53126
diff changeset
   146
  val bnf_timing: bool Config.T
49435
483007ddbdc2 bnf_note_all mode for "pre_"-BNFs
traytel
parents: 49434
diff changeset
   147
  val user_policy: fact_policy -> Proof.context -> fact_policy
66272
c6714a9562ae store the unfolded definitions of the lifted bnf constants under "_def" name
traytel
parents: 66198
diff changeset
   148
  val note_bnf_thms: fact_policy -> (binding -> binding) -> binding -> bnf -> local_theory ->
c6714a9562ae store the unfolded definitions of the lifted bnf constants under "_def" name
traytel
parents: 66198
diff changeset
   149
    bnf * local_theory
c6714a9562ae store the unfolded definitions of the lifted bnf constants under "_def" name
traytel
parents: 66198
diff changeset
   150
  val note_bnf_defs: bnf -> local_theory -> bnf * local_theory
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   151
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   152
  val print_bnfs: Proof.context -> unit
56016
8875cdcfc85b unfold intermediate definitions after sealing the bnf
traytel
parents: 55945
diff changeset
   153
  val prepare_def: inline_policy -> (Proof.context -> fact_policy) -> bool ->
8875cdcfc85b unfold intermediate definitions after sealing the bnf
traytel
parents: 55945
diff changeset
   154
    (binding -> binding) -> (Proof.context -> 'a -> typ) -> (Proof.context -> 'b -> term) ->
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   155
    typ list option -> binding -> binding -> binding -> binding list ->
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   156
    ((((((binding * 'a) * 'b) * 'b list) * 'b) * 'b list) * 'b option) * 'b option ->
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   157
    Proof.context ->
58189
9d714be4f028 added 'plugins' option to control which hooks are enabled
blanchet
parents: 58188
diff changeset
   158
    string * term list * ((Proof.context -> thm list -> tactic) option * term list list) *
54601
91a1e4aa7c80 command for axiomatic declaration of bnfs---allow theoreticians to fix an arbitrary type with functorial structure and work abstractly on it
traytel
parents: 54490
diff changeset
   159
    ((thm list -> thm list list) -> thm list list -> Proof.context -> bnf * local_theory) *
91a1e4aa7c80 command for axiomatic declaration of bnfs---allow theoreticians to fix an arbitrary type with functorial structure and work abstractly on it
traytel
parents: 54490
diff changeset
   160
    local_theory * thm list
56016
8875cdcfc85b unfold intermediate definitions after sealing the bnf
traytel
parents: 55945
diff changeset
   161
  val define_bnf_consts: inline_policy -> fact_policy -> bool -> typ list option ->
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   162
    binding -> binding -> binding -> binding list ->
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   163
    ((((((binding * typ) * term) * term list) * term) * term list) * term option) * term option ->
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   164
    local_theory ->
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
   165
      ((typ list * typ list * typ list * typ) *
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   166
       (term * term list * term * (int list * term) list * term * term) *
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   167
       (thm * thm list * thm * thm list * thm * thm) *
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
   168
       ((typ list -> typ list -> typ list -> term) *
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
   169
        (typ list -> typ list -> term -> term) *
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
   170
        (typ list -> typ list -> typ -> typ) *
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
   171
        (typ list -> typ list -> typ list -> term) *
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   172
        (typ list -> typ list -> term) *
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   173
        (typ list -> typ list -> typ list -> term) *
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   174
        (typ list -> typ list -> term))) * local_theory
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
   175
56016
8875cdcfc85b unfold intermediate definitions after sealing the bnf
traytel
parents: 55945
diff changeset
   176
  val bnf_def: inline_policy -> (Proof.context -> fact_policy) -> bool -> (binding -> binding) ->
58189
9d714be4f028 added 'plugins' option to control which hooks are enabled
blanchet
parents: 58188
diff changeset
   177
    (Proof.context -> tactic) list -> (Proof.context -> tactic) -> typ list option -> binding ->
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   178
    binding -> binding -> binding list ->
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   179
    ((((((binding * typ) * term) * term list) * term) * term list) * term option) * term option ->
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   180
    local_theory -> bnf * local_theory
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   181
  val bnf_cmd: (((((((binding * string) * string) * string list) * string) * string list)
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   182
      * string option) * string option) * (Proof.context -> Plugin_Name.filter) ->
61301
484f7878ede4 export '_cmd' functions
blanchet
parents: 61271
diff changeset
   183
    Proof.context -> Proof.state
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   184
end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   185
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   186
structure BNF_Def : BNF_DEF =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   187
struct
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   188
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   189
open BNF_Util
49463
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49462
diff changeset
   190
open BNF_Tactics
49284
5f39b7940b49 spin off "bnf_def_tactics.ML"
blanchet
parents: 49279
diff changeset
   191
open BNF_Def_Tactics
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   192
54624
36301c99ed26 revert making 'map_cong' a 'cong' -- it breaks too many proofs in the AFP
blanchet
parents: 54620
diff changeset
   193
val fundefcong_attrs = @{attributes [fundef_cong]};
58344
ea3d989219b4 set 'mono' attribute on 'rel_mono'
blanchet
parents: 58189
diff changeset
   194
val mono_attrs = @{attributes [mono]};
51765
224b6eb2313a added "fundef_cong" attribute to "map_cong"
blanchet
parents: 51762
diff changeset
   195
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   196
type axioms = {
53270
c8628119d18e renamed BNF axiom
blanchet
parents: 53265
diff changeset
   197
  map_id0: thm,
53287
271b34513bfb renamed BNF axiom
blanchet
parents: 53285
diff changeset
   198
  map_comp0: thm,
51761
4c9f08836d87 renamed "map_cong" axiom to "map_cong0" in preparation for real "map_cong"
blanchet
parents: 51758
diff changeset
   199
  map_cong0: thm,
53289
5e0623448bdb renamed BNF axiom
blanchet
parents: 53288
diff changeset
   200
  set_map0: thm list,
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   201
  bd_card_order: thm,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   202
  bd_cinfinite: thm,
75624
22d1c5f2b9f4 strict bounds for BNFs (by Jan van Brügge)
traytel
parents: 75276
diff changeset
   203
  bd_regularCard: thm,
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   204
  set_bd: thm list,
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
   205
  le_rel_OO: thm,
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   206
  rel_OO_Grp: thm,
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   207
  pred_set: thm
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   208
};
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   209
75624
22d1c5f2b9f4 strict bounds for BNFs (by Jan van Brügge)
traytel
parents: 75276
diff changeset
   210
fun mk_axioms' ((((((((((id, comp), cong), map), c_o), cinf), creg), set_bd), le_rel_OO), rel), pred) =
53289
5e0623448bdb renamed BNF axiom
blanchet
parents: 53288
diff changeset
   211
  {map_id0 = id, map_comp0 = comp, map_cong0 = cong, set_map0 = map, bd_card_order = c_o,
75624
22d1c5f2b9f4 strict bounds for BNFs (by Jan van Brügge)
traytel
parents: 75276
diff changeset
   212
   bd_cinfinite = cinf, bd_regularCard = creg, set_bd = set_bd, le_rel_OO = le_rel_OO, rel_OO_Grp = rel, pred_set = pred};
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   213
51930
52fd62618631 prefer explicitly qualified exceptions, which is particular important for robust handlers;
wenzelm
parents: 51917
diff changeset
   214
fun dest_cons [] = raise List.Empty
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   215
  | dest_cons (x :: xs) = (x, xs);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   216
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   217
fun mk_axioms n thms = thms
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   218
  |> map the_single
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   219
  |> dest_cons
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   220
  ||>> dest_cons
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   221
  ||>> dest_cons
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   222
  ||>> chop n
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   223
  ||>> dest_cons
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   224
  ||>> dest_cons
75624
22d1c5f2b9f4 strict bounds for BNFs (by Jan van Brügge)
traytel
parents: 75276
diff changeset
   225
  ||>> dest_cons
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   226
  ||>> chop n
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   227
  ||>> dest_cons
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   228
  ||>> dest_cons
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   229
  ||> the_single
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   230
  |> mk_axioms';
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   231
75624
22d1c5f2b9f4 strict bounds for BNFs (by Jan van Brügge)
traytel
parents: 75276
diff changeset
   232
fun zip_axioms mid mcomp mcong smap bdco bdinf bdreg sbd le_rel_OO rel pred =
22d1c5f2b9f4 strict bounds for BNFs (by Jan van Brügge)
traytel
parents: 75276
diff changeset
   233
  [mid, mcomp, mcong] @ smap @ [bdco, bdinf, bdreg] @ sbd @ [le_rel_OO, rel, pred];
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   234
75624
22d1c5f2b9f4 strict bounds for BNFs (by Jan van Brügge)
traytel
parents: 75276
diff changeset
   235
fun map_axioms f {map_id0, map_comp0, map_cong0, set_map0, bd_card_order, bd_cinfinite,
22d1c5f2b9f4 strict bounds for BNFs (by Jan van Brügge)
traytel
parents: 75276
diff changeset
   236
  bd_regularCard, set_bd, le_rel_OO, rel_OO_Grp, pred_set} =
53270
c8628119d18e renamed BNF axiom
blanchet
parents: 53265
diff changeset
   237
  {map_id0 = f map_id0,
53287
271b34513bfb renamed BNF axiom
blanchet
parents: 53285
diff changeset
   238
    map_comp0 = f map_comp0,
51761
4c9f08836d87 renamed "map_cong" axiom to "map_cong0" in preparation for real "map_cong"
blanchet
parents: 51758
diff changeset
   239
    map_cong0 = f map_cong0,
53289
5e0623448bdb renamed BNF axiom
blanchet
parents: 53288
diff changeset
   240
    set_map0 = map f set_map0,
49463
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49462
diff changeset
   241
    bd_card_order = f bd_card_order,
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49462
diff changeset
   242
    bd_cinfinite = f bd_cinfinite,
75624
22d1c5f2b9f4 strict bounds for BNFs (by Jan van Brügge)
traytel
parents: 75276
diff changeset
   243
    bd_regularCard = f bd_regularCard,
49463
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49462
diff changeset
   244
    set_bd = map f set_bd,
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
   245
    le_rel_OO = f le_rel_OO,
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   246
    rel_OO_Grp = f rel_OO_Grp,
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   247
    pred_set = f pred_set};
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   248
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   249
val morph_axioms = map_axioms o Morphism.thm;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   250
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   251
type defs = {
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   252
  map_def: thm,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   253
  set_defs: thm list,
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   254
  rel_def: thm,
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   255
  pred_def: thm
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   256
}
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   257
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   258
fun mk_defs map sets rel pred = {map_def = map, set_defs = sets, rel_def = rel, pred_def = pred};
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   259
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   260
fun map_defs f {map_def, set_defs, rel_def, pred_def} =
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   261
  {map_def = f map_def, set_defs = map f set_defs, rel_def = f rel_def, pred_def = f pred_def};
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   262
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   263
val morph_defs = map_defs o Morphism.thm;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   264
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   265
type facts = {
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   266
  bd_Card_order: thm,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   267
  bd_Cinfinite: thm,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   268
  bd_Cnotzero: thm,
51766
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51765
diff changeset
   269
  collect_set_map: thm lazy,
52635
4f84b730c489 got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
traytel
parents: 51930
diff changeset
   270
  in_bd: thm lazy,
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   271
  in_cong: thm lazy,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   272
  in_mono: thm lazy,
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
   273
  in_rel: thm lazy,
56635
b07c8ad23010 added 'inj_map' as auxiliary BNF theorem
blanchet
parents: 56523
diff changeset
   274
  inj_map: thm lazy,
57970
eaa986cd285a generate 'inj_map_strong' for BNFs
desharna
parents: 57969
diff changeset
   275
  inj_map_strong: thm lazy,
53288
050d0bc9afa5 renamed BNF fact
blanchet
parents: 53287
diff changeset
   276
  map_comp: thm lazy,
51762
219a3063ed29 derive "map_cong"
blanchet
parents: 51761
diff changeset
   277
  map_cong: thm lazy,
57981
81baacfd56e8 generate 'map_cong_simp' for BNFs
desharna
parents: 57970
diff changeset
   278
  map_cong_simp: thm lazy,
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   279
  map_cong_pred: thm lazy,
53285
f09645642984 renamed BNF fact
blanchet
parents: 53270
diff changeset
   280
  map_id: thm lazy,
57399
cfc19f0a6261 compile
blanchet
parents: 57303
diff changeset
   281
  map_ident0: thm lazy,
56903
69b6369a98fa generate 'map_ident' theorem for BNFs
desharna
parents: 56766
diff changeset
   282
  map_ident: thm lazy,
75276
686a6d7d0991 generated lemma map_ident_strong for BNFs
desharna
parents: 74664
diff changeset
   283
  map_ident_strong: thm lazy,
52719
480a3479fa47 transfer rule for map (not yet registered as a transfer rule)
traytel
parents: 52635
diff changeset
   284
  map_transfer: thm lazy,
49591
91b228e26348 generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents: 49585
diff changeset
   285
  rel_eq: thm lazy,
49537
fe1deee434b6 generate "rel_as_srel" and "rel_flip" properties
blanchet
parents: 49536
diff changeset
   286
  rel_flip: thm lazy,
53290
b6c3be868217 renamed BNF fact
blanchet
parents: 53289
diff changeset
   287
  set_map: thm lazy list,
61242
1f92b0ac9c96 congruence rules for the relator
traytel
parents: 61241
diff changeset
   288
  rel_cong0: thm lazy,
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
   289
  rel_cong: thm lazy,
61242
1f92b0ac9c96 congruence rules for the relator
traytel
parents: 61241
diff changeset
   290
  rel_cong_simp: thm lazy,
57932
c29659f77f8d generate 'rel_map' theorem for BNFs
desharna
parents: 57668
diff changeset
   291
  rel_map: thm list lazy,
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
   292
  rel_mono: thm lazy,
57967
e6d2e998c30f renamed 'rel_mono_strong' to 'rel_mono_strong0'
desharna
parents: 57932
diff changeset
   293
  rel_mono_strong0: thm lazy,
57968
00e9c6d367e7 generate property 'rel_mono_strong' for BNFs
desharna
parents: 57967
diff changeset
   294
  rel_mono_strong: thm lazy,
58106
c8cba801c483 generate 'set_transfer' for BNFs
desharna
parents: 58104
diff changeset
   295
  set_transfer: thm list lazy,
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
   296
  rel_Grp: thm lazy,
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
   297
  rel_conversep: thm lazy,
58106
c8cba801c483 generate 'set_transfer' for BNFs
desharna
parents: 58104
diff changeset
   298
  rel_OO: thm lazy,
59726
64c2bb331035 BNF relators preserve reflexivity
traytel
parents: 59663
diff changeset
   299
  rel_refl: thm lazy,
61240
464c5da3f508 more useful properties of the relators
traytel
parents: 61116
diff changeset
   300
  rel_refl_strong: thm lazy,
464c5da3f508 more useful properties of the relators
traytel
parents: 61116
diff changeset
   301
  rel_reflp: thm lazy,
464c5da3f508 more useful properties of the relators
traytel
parents: 61116
diff changeset
   302
  rel_symp: thm lazy,
464c5da3f508 more useful properties of the relators
traytel
parents: 61116
diff changeset
   303
  rel_transp: thm lazy,
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   304
  rel_transfer: thm lazy,
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   305
  rel_eq_onp: thm lazy,
62329
9f7fa08d2307 derive transfer rule for predicator
traytel
parents: 62326
diff changeset
   306
  pred_transfer: thm lazy,
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   307
  pred_True: thm lazy,
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   308
  pred_map: thm lazy,
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   309
  pred_rel: thm lazy,
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   310
  pred_mono_strong0: thm lazy,
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   311
  pred_mono_strong: thm lazy,
63714
b62f4f765353 derive pred_mono property for BNFs
traytel
parents: 63399
diff changeset
   312
  pred_mono: thm lazy,
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   313
  pred_cong0: thm lazy,
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   314
  pred_cong: thm lazy,
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   315
  pred_cong_simp: thm lazy
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   316
};
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   317
52635
4f84b730c489 got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
traytel
parents: 51930
diff changeset
   318
fun mk_facts bd_Card_order bd_Cinfinite bd_Cnotzero collect_set_map in_bd in_cong in_mono in_rel
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   319
    inj_map inj_map_strong map_comp map_cong map_cong_simp map_cong_pred map_id map_ident0 map_ident
75276
686a6d7d0991 generated lemma map_ident_strong for BNFs
desharna
parents: 74664
diff changeset
   320
    map_ident_strong map_transfer rel_eq rel_flip set_map rel_cong0 rel_cong rel_cong_simp rel_map
686a6d7d0991 generated lemma map_ident_strong for BNFs
desharna
parents: 74664
diff changeset
   321
    rel_mono rel_mono_strong0 rel_mono_strong set_transfer rel_Grp rel_conversep rel_OO rel_refl
62329
9f7fa08d2307 derive transfer rule for predicator
traytel
parents: 62326
diff changeset
   322
    rel_refl_strong rel_reflp rel_symp rel_transp rel_transfer rel_eq_onp pred_transfer pred_True
63714
b62f4f765353 derive pred_mono property for BNFs
traytel
parents: 63399
diff changeset
   323
    pred_map pred_rel pred_mono_strong0 pred_mono_strong pred_mono pred_cong0 pred_cong
b62f4f765353 derive pred_mono property for BNFs
traytel
parents: 63399
diff changeset
   324
    pred_cong_simp = {
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   325
  bd_Card_order = bd_Card_order,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   326
  bd_Cinfinite = bd_Cinfinite,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   327
  bd_Cnotzero = bd_Cnotzero,
51766
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51765
diff changeset
   328
  collect_set_map = collect_set_map,
52635
4f84b730c489 got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
traytel
parents: 51930
diff changeset
   329
  in_bd = in_bd,
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   330
  in_cong = in_cong,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   331
  in_mono = in_mono,
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
   332
  in_rel = in_rel,
56635
b07c8ad23010 added 'inj_map' as auxiliary BNF theorem
blanchet
parents: 56523
diff changeset
   333
  inj_map = inj_map,
57970
eaa986cd285a generate 'inj_map_strong' for BNFs
desharna
parents: 57969
diff changeset
   334
  inj_map_strong = inj_map_strong,
53288
050d0bc9afa5 renamed BNF fact
blanchet
parents: 53287
diff changeset
   335
  map_comp = map_comp,
51762
219a3063ed29 derive "map_cong"
blanchet
parents: 51761
diff changeset
   336
  map_cong = map_cong,
57981
81baacfd56e8 generate 'map_cong_simp' for BNFs
desharna
parents: 57970
diff changeset
   337
  map_cong_simp = map_cong_simp,
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   338
  map_cong_pred = map_cong_pred,
53285
f09645642984 renamed BNF fact
blanchet
parents: 53270
diff changeset
   339
  map_id = map_id,
57399
cfc19f0a6261 compile
blanchet
parents: 57303
diff changeset
   340
  map_ident0 = map_ident0,
56903
69b6369a98fa generate 'map_ident' theorem for BNFs
desharna
parents: 56766
diff changeset
   341
  map_ident = map_ident,
75276
686a6d7d0991 generated lemma map_ident_strong for BNFs
desharna
parents: 74664
diff changeset
   342
  map_ident_strong = map_ident_strong,
52719
480a3479fa47 transfer rule for map (not yet registered as a transfer rule)
traytel
parents: 52635
diff changeset
   343
  map_transfer = map_transfer,
49591
91b228e26348 generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents: 49585
diff changeset
   344
  rel_eq = rel_eq,
49537
fe1deee434b6 generate "rel_as_srel" and "rel_flip" properties
blanchet
parents: 49536
diff changeset
   345
  rel_flip = rel_flip,
53290
b6c3be868217 renamed BNF fact
blanchet
parents: 53289
diff changeset
   346
  set_map = set_map,
61242
1f92b0ac9c96 congruence rules for the relator
traytel
parents: 61241
diff changeset
   347
  rel_cong0 = rel_cong0,
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
   348
  rel_cong = rel_cong,
61242
1f92b0ac9c96 congruence rules for the relator
traytel
parents: 61241
diff changeset
   349
  rel_cong_simp = rel_cong_simp,
57932
c29659f77f8d generate 'rel_map' theorem for BNFs
desharna
parents: 57668
diff changeset
   350
  rel_map = rel_map,
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
   351
  rel_mono = rel_mono,
57967
e6d2e998c30f renamed 'rel_mono_strong' to 'rel_mono_strong0'
desharna
parents: 57932
diff changeset
   352
  rel_mono_strong0 = rel_mono_strong0,
57968
00e9c6d367e7 generate property 'rel_mono_strong' for BNFs
desharna
parents: 57967
diff changeset
   353
  rel_mono_strong = rel_mono_strong,
58104
c5316f843f72 generate 'rel_transfer' for BNFs
desharna
parents: 58102
diff changeset
   354
  rel_transfer = rel_transfer,
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
   355
  rel_Grp = rel_Grp,
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
   356
  rel_conversep = rel_conversep,
58106
c8cba801c483 generate 'set_transfer' for BNFs
desharna
parents: 58104
diff changeset
   357
  rel_OO = rel_OO,
59726
64c2bb331035 BNF relators preserve reflexivity
traytel
parents: 59663
diff changeset
   358
  rel_refl = rel_refl,
61240
464c5da3f508 more useful properties of the relators
traytel
parents: 61116
diff changeset
   359
  rel_refl_strong = rel_refl_strong,
464c5da3f508 more useful properties of the relators
traytel
parents: 61116
diff changeset
   360
  rel_reflp = rel_reflp,
464c5da3f508 more useful properties of the relators
traytel
parents: 61116
diff changeset
   361
  rel_symp = rel_symp,
464c5da3f508 more useful properties of the relators
traytel
parents: 61116
diff changeset
   362
  rel_transp = rel_transp,
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   363
  set_transfer = set_transfer,
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   364
  rel_eq_onp = rel_eq_onp,
62329
9f7fa08d2307 derive transfer rule for predicator
traytel
parents: 62326
diff changeset
   365
  pred_transfer = pred_transfer,
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   366
  pred_True = pred_True,
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   367
  pred_map = pred_map,
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   368
  pred_rel = pred_rel,
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   369
  pred_mono_strong0 = pred_mono_strong0,
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   370
  pred_mono_strong = pred_mono_strong,
63714
b62f4f765353 derive pred_mono property for BNFs
traytel
parents: 63399
diff changeset
   371
  pred_mono = pred_mono,
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   372
  pred_cong0 = pred_cong0,
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   373
  pred_cong = pred_cong,
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   374
  pred_cong_simp = pred_cong_simp};
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   375
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   376
fun map_facts f {
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   377
  bd_Card_order,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   378
  bd_Cinfinite,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   379
  bd_Cnotzero,
51766
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51765
diff changeset
   380
  collect_set_map,
52635
4f84b730c489 got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
traytel
parents: 51930
diff changeset
   381
  in_bd,
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   382
  in_cong,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   383
  in_mono,
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
   384
  in_rel,
56635
b07c8ad23010 added 'inj_map' as auxiliary BNF theorem
blanchet
parents: 56523
diff changeset
   385
  inj_map,
57970
eaa986cd285a generate 'inj_map_strong' for BNFs
desharna
parents: 57969
diff changeset
   386
  inj_map_strong,
53288
050d0bc9afa5 renamed BNF fact
blanchet
parents: 53287
diff changeset
   387
  map_comp,
51762
219a3063ed29 derive "map_cong"
blanchet
parents: 51761
diff changeset
   388
  map_cong,
57981
81baacfd56e8 generate 'map_cong_simp' for BNFs
desharna
parents: 57970
diff changeset
   389
  map_cong_simp,
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   390
  map_cong_pred,
53285
f09645642984 renamed BNF fact
blanchet
parents: 53270
diff changeset
   391
  map_id,
57399
cfc19f0a6261 compile
blanchet
parents: 57303
diff changeset
   392
  map_ident0,
56903
69b6369a98fa generate 'map_ident' theorem for BNFs
desharna
parents: 56766
diff changeset
   393
  map_ident,
75276
686a6d7d0991 generated lemma map_ident_strong for BNFs
desharna
parents: 74664
diff changeset
   394
  map_ident_strong,
52719
480a3479fa47 transfer rule for map (not yet registered as a transfer rule)
traytel
parents: 52635
diff changeset
   395
  map_transfer,
49591
91b228e26348 generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents: 49585
diff changeset
   396
  rel_eq,
49537
fe1deee434b6 generate "rel_as_srel" and "rel_flip" properties
blanchet
parents: 49536
diff changeset
   397
  rel_flip,
53290
b6c3be868217 renamed BNF fact
blanchet
parents: 53289
diff changeset
   398
  set_map,
61242
1f92b0ac9c96 congruence rules for the relator
traytel
parents: 61241
diff changeset
   399
  rel_cong0,
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
   400
  rel_cong,
61242
1f92b0ac9c96 congruence rules for the relator
traytel
parents: 61241
diff changeset
   401
  rel_cong_simp,
57932
c29659f77f8d generate 'rel_map' theorem for BNFs
desharna
parents: 57668
diff changeset
   402
  rel_map,
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
   403
  rel_mono,
57967
e6d2e998c30f renamed 'rel_mono_strong' to 'rel_mono_strong0'
desharna
parents: 57932
diff changeset
   404
  rel_mono_strong0,
57968
00e9c6d367e7 generate property 'rel_mono_strong' for BNFs
desharna
parents: 57967
diff changeset
   405
  rel_mono_strong,
58104
c5316f843f72 generate 'rel_transfer' for BNFs
desharna
parents: 58102
diff changeset
   406
  rel_transfer,
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
   407
  rel_Grp,
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
   408
  rel_conversep,
58106
c8cba801c483 generate 'set_transfer' for BNFs
desharna
parents: 58104
diff changeset
   409
  rel_OO,
59726
64c2bb331035 BNF relators preserve reflexivity
traytel
parents: 59663
diff changeset
   410
  rel_refl,
61240
464c5da3f508 more useful properties of the relators
traytel
parents: 61116
diff changeset
   411
  rel_refl_strong,
464c5da3f508 more useful properties of the relators
traytel
parents: 61116
diff changeset
   412
  rel_reflp,
464c5da3f508 more useful properties of the relators
traytel
parents: 61116
diff changeset
   413
  rel_symp,
464c5da3f508 more useful properties of the relators
traytel
parents: 61116
diff changeset
   414
  rel_transp,
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   415
  set_transfer,
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   416
  rel_eq_onp,
62329
9f7fa08d2307 derive transfer rule for predicator
traytel
parents: 62326
diff changeset
   417
  pred_transfer,
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   418
  pred_True,
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   419
  pred_map,
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   420
  pred_rel,
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   421
  pred_mono_strong0,
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   422
  pred_mono_strong,
63714
b62f4f765353 derive pred_mono property for BNFs
traytel
parents: 63399
diff changeset
   423
  pred_mono,
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   424
  pred_cong0,
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   425
  pred_cong,
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   426
  pred_cong_simp} =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   427
  {bd_Card_order = f bd_Card_order,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   428
    bd_Cinfinite = f bd_Cinfinite,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   429
    bd_Cnotzero = f bd_Cnotzero,
51766
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51765
diff changeset
   430
    collect_set_map = Lazy.map f collect_set_map,
52635
4f84b730c489 got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
traytel
parents: 51930
diff changeset
   431
    in_bd = Lazy.map f in_bd,
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   432
    in_cong = Lazy.map f in_cong,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   433
    in_mono = Lazy.map f in_mono,
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
   434
    in_rel = Lazy.map f in_rel,
56635
b07c8ad23010 added 'inj_map' as auxiliary BNF theorem
blanchet
parents: 56523
diff changeset
   435
    inj_map = Lazy.map f inj_map,
57970
eaa986cd285a generate 'inj_map_strong' for BNFs
desharna
parents: 57969
diff changeset
   436
    inj_map_strong = Lazy.map f inj_map_strong,
53288
050d0bc9afa5 renamed BNF fact
blanchet
parents: 53287
diff changeset
   437
    map_comp = Lazy.map f map_comp,
51762
219a3063ed29 derive "map_cong"
blanchet
parents: 51761
diff changeset
   438
    map_cong = Lazy.map f map_cong,
57981
81baacfd56e8 generate 'map_cong_simp' for BNFs
desharna
parents: 57970
diff changeset
   439
    map_cong_simp = Lazy.map f map_cong_simp,
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   440
    map_cong_pred = Lazy.map f map_cong_pred,
53285
f09645642984 renamed BNF fact
blanchet
parents: 53270
diff changeset
   441
    map_id = Lazy.map f map_id,
57399
cfc19f0a6261 compile
blanchet
parents: 57303
diff changeset
   442
    map_ident0 = Lazy.map f map_ident0,
56903
69b6369a98fa generate 'map_ident' theorem for BNFs
desharna
parents: 56766
diff changeset
   443
    map_ident = Lazy.map f map_ident,
75276
686a6d7d0991 generated lemma map_ident_strong for BNFs
desharna
parents: 74664
diff changeset
   444
    map_ident_strong = Lazy.map f map_ident_strong,
52719
480a3479fa47 transfer rule for map (not yet registered as a transfer rule)
traytel
parents: 52635
diff changeset
   445
    map_transfer = Lazy.map f map_transfer,
49591
91b228e26348 generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents: 49585
diff changeset
   446
    rel_eq = Lazy.map f rel_eq,
49537
fe1deee434b6 generate "rel_as_srel" and "rel_flip" properties
blanchet
parents: 49536
diff changeset
   447
    rel_flip = Lazy.map f rel_flip,
53290
b6c3be868217 renamed BNF fact
blanchet
parents: 53289
diff changeset
   448
    set_map = map (Lazy.map f) set_map,
61242
1f92b0ac9c96 congruence rules for the relator
traytel
parents: 61241
diff changeset
   449
    rel_cong0 = Lazy.map f rel_cong0,
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
   450
    rel_cong = Lazy.map f rel_cong,
61242
1f92b0ac9c96 congruence rules for the relator
traytel
parents: 61241
diff changeset
   451
    rel_cong_simp = Lazy.map f rel_cong_simp,
57932
c29659f77f8d generate 'rel_map' theorem for BNFs
desharna
parents: 57668
diff changeset
   452
    rel_map = Lazy.map (map f) rel_map,
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
   453
    rel_mono = Lazy.map f rel_mono,
57967
e6d2e998c30f renamed 'rel_mono_strong' to 'rel_mono_strong0'
desharna
parents: 57932
diff changeset
   454
    rel_mono_strong0 = Lazy.map f rel_mono_strong0,
57968
00e9c6d367e7 generate property 'rel_mono_strong' for BNFs
desharna
parents: 57967
diff changeset
   455
    rel_mono_strong = Lazy.map f rel_mono_strong,
58104
c5316f843f72 generate 'rel_transfer' for BNFs
desharna
parents: 58102
diff changeset
   456
    rel_transfer = Lazy.map f rel_transfer,
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
   457
    rel_Grp = Lazy.map f rel_Grp,
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
   458
    rel_conversep = Lazy.map f rel_conversep,
58106
c8cba801c483 generate 'set_transfer' for BNFs
desharna
parents: 58104
diff changeset
   459
    rel_OO = Lazy.map f rel_OO,
59726
64c2bb331035 BNF relators preserve reflexivity
traytel
parents: 59663
diff changeset
   460
    rel_refl = Lazy.map f rel_refl,
61240
464c5da3f508 more useful properties of the relators
traytel
parents: 61116
diff changeset
   461
    rel_refl_strong = Lazy.map f rel_refl_strong,
464c5da3f508 more useful properties of the relators
traytel
parents: 61116
diff changeset
   462
    rel_reflp = Lazy.map f rel_reflp,
464c5da3f508 more useful properties of the relators
traytel
parents: 61116
diff changeset
   463
    rel_symp = Lazy.map f rel_symp,
464c5da3f508 more useful properties of the relators
traytel
parents: 61116
diff changeset
   464
    rel_transp = Lazy.map f rel_transp,
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   465
    set_transfer = Lazy.map (map f) set_transfer,
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   466
    rel_eq_onp = Lazy.map f rel_eq_onp,
62329
9f7fa08d2307 derive transfer rule for predicator
traytel
parents: 62326
diff changeset
   467
    pred_transfer = Lazy.map f pred_transfer,
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   468
    pred_True = Lazy.map f pred_True,
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   469
    pred_map = Lazy.map f pred_map,
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   470
    pred_rel = Lazy.map f pred_rel,
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   471
    pred_mono_strong0 = Lazy.map f pred_mono_strong0,
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   472
    pred_mono_strong = Lazy.map f pred_mono_strong,
63714
b62f4f765353 derive pred_mono property for BNFs
traytel
parents: 63399
diff changeset
   473
    pred_mono = Lazy.map f pred_mono,
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   474
    pred_cong0 = Lazy.map f pred_cong0,
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   475
    pred_cong = Lazy.map f pred_cong,
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   476
    pred_cong_simp = Lazy.map f pred_cong_simp};
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   477
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   478
val morph_facts = map_facts o Morphism.thm;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   479
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   480
type nonemptiness_witness = {
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   481
  I: int list,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   482
  wit: term,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   483
  prop: thm list
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   484
};
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   485
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   486
fun mk_witness (I, wit) prop = {I = I, wit = wit, prop = prop};
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   487
fun map_witness f g {I, wit, prop} = {I = I, wit = f wit, prop = map g prop};
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   488
fun morph_witness phi = map_witness (Morphism.term phi) (Morphism.thm phi);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   489
51837
087498724486 lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
blanchet
parents: 51836
diff changeset
   490
datatype bnf = BNF of {
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   491
  name: binding,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   492
  T: typ,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   493
  live: int,
53261
3c26a7042d8e removed outdated comments
panny
parents: 53143
diff changeset
   494
  lives: typ list, (*source type variables of map*)
3c26a7042d8e removed outdated comments
panny
parents: 53143
diff changeset
   495
  lives': typ list, (*target type variables of map*)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   496
  dead: int,
53261
3c26a7042d8e removed outdated comments
panny
parents: 53143
diff changeset
   497
  deads: typ list,
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   498
  map: term,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   499
  sets: term list,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   500
  bd: term,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   501
  axioms: axioms,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   502
  defs: defs,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   503
  facts: facts,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   504
  nwits: int,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   505
  wits: nonemptiness_witness list,
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   506
  rel: term,
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   507
  pred: term
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   508
};
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   509
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   510
(* getters *)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   511
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   512
fun rep_bnf (BNF bnf) = bnf;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   513
val name_of_bnf = #name o rep_bnf;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   514
val T_of_bnf = #T o rep_bnf;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   515
fun mk_T_of_bnf Ds Ts bnf =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   516
  let val bnf_rep = rep_bnf bnf
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   517
  in Term.typ_subst_atomic ((#deads bnf_rep ~~ Ds) @ (#lives bnf_rep ~~ Ts)) (#T bnf_rep) end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   518
val live_of_bnf = #live o rep_bnf;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   519
val lives_of_bnf = #lives o rep_bnf;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   520
val dead_of_bnf = #dead o rep_bnf;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   521
val deads_of_bnf = #deads o rep_bnf;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   522
val axioms_of_bnf = #axioms o rep_bnf;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   523
val facts_of_bnf = #facts o rep_bnf;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   524
val nwits_of_bnf = #nwits o rep_bnf;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   525
val wits_of_bnf = #wits o rep_bnf;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   526
53031
83cbe188855a moved useful library functions upstream
blanchet
parents: 52957
diff changeset
   527
fun flatten_type_args_of_bnf bnf dead_x xs =
83cbe188855a moved useful library functions upstream
blanchet
parents: 52957
diff changeset
   528
  let
83cbe188855a moved useful library functions upstream
blanchet
parents: 52957
diff changeset
   529
    val Type (_, Ts) = T_of_bnf bnf;
83cbe188855a moved useful library functions upstream
blanchet
parents: 52957
diff changeset
   530
    val lives = lives_of_bnf bnf;
83cbe188855a moved useful library functions upstream
blanchet
parents: 52957
diff changeset
   531
    val deads = deads_of_bnf bnf;
83cbe188855a moved useful library functions upstream
blanchet
parents: 52957
diff changeset
   532
  in
55480
59cc4a8bc28a allow different functions to recurse on the same type, like in the old package
blanchet
parents: 55444
diff changeset
   533
    permute_like_unique (op =) (deads @ lives) Ts (replicate (length deads) dead_x @ xs)
53031
83cbe188855a moved useful library functions upstream
blanchet
parents: 52957
diff changeset
   534
  end;
83cbe188855a moved useful library functions upstream
blanchet
parents: 52957
diff changeset
   535
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   536
(*terms*)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   537
val map_of_bnf = #map o rep_bnf;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   538
val sets_of_bnf = #sets o rep_bnf;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   539
fun mk_map_of_bnf Ds Ts Us bnf =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   540
  let val bnf_rep = rep_bnf bnf;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   541
  in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   542
    Term.subst_atomic_types
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   543
      ((#deads bnf_rep ~~ Ds) @ (#lives bnf_rep ~~ Ts) @ (#lives' bnf_rep ~~ Us)) (#map bnf_rep)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   544
  end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   545
fun mk_sets_of_bnf Dss Tss bnf =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   546
  let val bnf_rep = rep_bnf bnf;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   547
  in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   548
    map2 (fn (Ds, Ts) => Term.subst_atomic_types
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   549
      ((#deads bnf_rep ~~ Ds) @ (#lives bnf_rep ~~ Ts))) (Dss ~~ Tss) (#sets bnf_rep)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   550
  end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   551
val bd_of_bnf = #bd o rep_bnf;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   552
fun mk_bd_of_bnf Ds Ts bnf =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   553
  let val bnf_rep = rep_bnf bnf;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   554
  in Term.subst_atomic_types ((#deads bnf_rep ~~ Ds) @ (#lives bnf_rep ~~ Ts)) (#bd bnf_rep) end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   555
fun mk_wits_of_bnf Dss Tss bnf =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   556
  let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   557
    val bnf_rep = rep_bnf bnf;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   558
    val wits = map (fn x => (#I x, #wit x)) (#wits bnf_rep);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   559
  in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   560
    map2 (fn (Ds, Ts) => apsnd (Term.subst_atomic_types
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   561
      ((#deads bnf_rep ~~ Ds) @ (#lives bnf_rep ~~ Ts)))) (Dss ~~ Tss) wits
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   562
  end;
49507
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49506
diff changeset
   563
val rel_of_bnf = #rel o rep_bnf;
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49506
diff changeset
   564
fun mk_rel_of_bnf Ds Ts Us bnf =
49462
blanchet
parents: 49461
diff changeset
   565
  let val bnf_rep = rep_bnf bnf;
blanchet
parents: 49461
diff changeset
   566
  in
blanchet
parents: 49461
diff changeset
   567
    Term.subst_atomic_types
49507
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49506
diff changeset
   568
      ((#deads bnf_rep ~~ Ds) @ (#lives bnf_rep ~~ Ts) @ (#lives' bnf_rep ~~ Us)) (#rel bnf_rep)
49462
blanchet
parents: 49461
diff changeset
   569
  end;
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   570
val pred_of_bnf = #pred o rep_bnf;
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   571
fun mk_pred_of_bnf Ds Ts bnf =
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   572
  let val bnf_rep = rep_bnf bnf;
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   573
  in
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   574
    Term.subst_atomic_types
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   575
      ((#deads bnf_rep ~~ Ds) @ (#lives bnf_rep ~~ Ts)) (#pred bnf_rep)
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   576
  end;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   577
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   578
(*thms*)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   579
val bd_Card_order_of_bnf = #bd_Card_order o #facts o rep_bnf;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   580
val bd_Cinfinite_of_bnf = #bd_Cinfinite o #facts o rep_bnf;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   581
val bd_Cnotzero_of_bnf = #bd_Cnotzero o #facts o rep_bnf;
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   582
val bd_card_order_of_bnf = #bd_card_order o #axioms o rep_bnf;
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   583
val bd_cinfinite_of_bnf = #bd_cinfinite o #axioms o rep_bnf;
75624
22d1c5f2b9f4 strict bounds for BNFs (by Jan van Brügge)
traytel
parents: 75276
diff changeset
   584
val bd_regularCard_of_bnf = #bd_regularCard o #axioms o rep_bnf;
51766
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51765
diff changeset
   585
val collect_set_map_of_bnf = Lazy.force o #collect_set_map o #facts o rep_bnf;
52635
4f84b730c489 got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
traytel
parents: 51930
diff changeset
   586
val in_bd_of_bnf = Lazy.force o #in_bd o #facts o rep_bnf;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   587
val in_cong_of_bnf = Lazy.force o #in_cong o #facts o rep_bnf;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   588
val in_mono_of_bnf = Lazy.force o #in_mono o #facts o rep_bnf;
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
   589
val in_rel_of_bnf = Lazy.force o #in_rel o #facts o rep_bnf;
56635
b07c8ad23010 added 'inj_map' as auxiliary BNF theorem
blanchet
parents: 56523
diff changeset
   590
val inj_map_of_bnf = Lazy.force o #inj_map o #facts o rep_bnf;
57970
eaa986cd285a generate 'inj_map_strong' for BNFs
desharna
parents: 57969
diff changeset
   591
val inj_map_strong_of_bnf = Lazy.force o #inj_map_strong o #facts o rep_bnf;
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   592
val le_rel_OO_of_bnf = #le_rel_OO o #axioms o rep_bnf;
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   593
val map_comp0_of_bnf = #map_comp0 o #axioms o rep_bnf;
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   594
val map_comp_of_bnf = Lazy.force o #map_comp o #facts o rep_bnf;
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   595
val map_cong0_of_bnf = #map_cong0 o #axioms o rep_bnf;
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   596
val map_cong_of_bnf = Lazy.force o #map_cong o #facts o rep_bnf;
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   597
val map_cong_pred_of_bnf = Lazy.force o #map_cong_pred o #facts o rep_bnf;
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   598
val map_cong_simp_of_bnf = Lazy.force o #map_cong_simp o #facts o rep_bnf;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   599
val map_def_of_bnf = #map_def o #defs o rep_bnf;
53270
c8628119d18e renamed BNF axiom
blanchet
parents: 53265
diff changeset
   600
val map_id0_of_bnf = #map_id0 o #axioms o rep_bnf;
53285
f09645642984 renamed BNF fact
blanchet
parents: 53270
diff changeset
   601
val map_id_of_bnf = Lazy.force o #map_id o #facts o rep_bnf;
57399
cfc19f0a6261 compile
blanchet
parents: 57303
diff changeset
   602
val map_ident0_of_bnf = Lazy.force o #map_ident0 o #facts o rep_bnf;
56903
69b6369a98fa generate 'map_ident' theorem for BNFs
desharna
parents: 56766
diff changeset
   603
val map_ident_of_bnf = Lazy.force o #map_ident o #facts o rep_bnf;
75276
686a6d7d0991 generated lemma map_ident_strong for BNFs
desharna
parents: 74664
diff changeset
   604
val map_ident_strong_of_bnf = Lazy.force o #map_ident_strong o #facts o rep_bnf;
52731
dacd47a0633f transfer rule for {c,d}tor_{,un}fold
traytel
parents: 52725
diff changeset
   605
val map_transfer_of_bnf = Lazy.force o #map_transfer o #facts o rep_bnf;
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   606
val rel_eq_onp_of_bnf = Lazy.force o #rel_eq_onp o #facts o rep_bnf;
62329
9f7fa08d2307 derive transfer rule for predicator
traytel
parents: 62326
diff changeset
   607
val pred_def_of_bnf = #pred_def o #defs o rep_bnf;
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   608
val pred_map_of_bnf = Lazy.force o #pred_map o #facts o rep_bnf;
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   609
val pred_mono_strong0_of_bnf = Lazy.force o #pred_mono_strong0 o #facts o rep_bnf;
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   610
val pred_mono_strong_of_bnf = Lazy.force o #pred_mono_strong o #facts o rep_bnf;
63714
b62f4f765353 derive pred_mono property for BNFs
traytel
parents: 63399
diff changeset
   611
val pred_mono_of_bnf = Lazy.force o #pred_mono o #facts o rep_bnf;
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   612
val pred_cong0_of_bnf = Lazy.force o #pred_cong0 o #facts o rep_bnf;
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   613
val pred_cong_of_bnf = Lazy.force o #pred_cong o #facts o rep_bnf;
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   614
val pred_cong_simp_of_bnf = Lazy.force o #pred_cong_simp o #facts o rep_bnf;
62329
9f7fa08d2307 derive transfer rule for predicator
traytel
parents: 62326
diff changeset
   615
val pred_rel_of_bnf = Lazy.force o #pred_rel o #facts o rep_bnf;
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   616
val pred_set_of_bnf = #pred_set o #axioms o rep_bnf;
62329
9f7fa08d2307 derive transfer rule for predicator
traytel
parents: 62326
diff changeset
   617
val pred_transfer_of_bnf = Lazy.force o #pred_transfer o #facts o rep_bnf;
9f7fa08d2307 derive transfer rule for predicator
traytel
parents: 62326
diff changeset
   618
val pred_True_of_bnf = Lazy.force o #pred_True o #facts o rep_bnf;
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   619
val rel_Grp_of_bnf = Lazy.force o #rel_Grp o #facts o rep_bnf;
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   620
val rel_OO_Grp_of_bnf = #rel_OO_Grp o #axioms o rep_bnf;
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   621
val rel_OO_of_bnf = Lazy.force o #rel_OO o #facts o rep_bnf;
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   622
val rel_cong0_of_bnf = Lazy.force o #rel_cong0 o #facts o rep_bnf;
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   623
val rel_cong_of_bnf = Lazy.force o #rel_cong o #facts o rep_bnf;
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   624
val rel_cong_simp_of_bnf = Lazy.force o #rel_cong_simp o #facts o rep_bnf;
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   625
val rel_conversep_of_bnf = Lazy.force o #rel_conversep o #facts o rep_bnf;
49507
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49506
diff changeset
   626
val rel_def_of_bnf = #rel_def o #defs o rep_bnf;
49591
91b228e26348 generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents: 49585
diff changeset
   627
val rel_eq_of_bnf = Lazy.force o #rel_eq o #facts o rep_bnf;
49537
fe1deee434b6 generate "rel_as_srel" and "rel_flip" properties
blanchet
parents: 49536
diff changeset
   628
val rel_flip_of_bnf = Lazy.force o #rel_flip o #facts o rep_bnf;
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   629
val rel_map_of_bnf = Lazy.force o #rel_map o #facts o rep_bnf;
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
   630
val rel_mono_of_bnf = Lazy.force o #rel_mono o #facts o rep_bnf;
57967
e6d2e998c30f renamed 'rel_mono_strong' to 'rel_mono_strong0'
desharna
parents: 57932
diff changeset
   631
val rel_mono_strong0_of_bnf = Lazy.force o #rel_mono_strong0 o #facts o rep_bnf;
57968
00e9c6d367e7 generate property 'rel_mono_strong' for BNFs
desharna
parents: 57967
diff changeset
   632
val rel_mono_strong_of_bnf = Lazy.force o #rel_mono_strong o #facts o rep_bnf;
59726
64c2bb331035 BNF relators preserve reflexivity
traytel
parents: 59663
diff changeset
   633
val rel_refl_of_bnf = Lazy.force o #rel_refl o #facts o rep_bnf;
61240
464c5da3f508 more useful properties of the relators
traytel
parents: 61116
diff changeset
   634
val rel_refl_strong_of_bnf = Lazy.force o #rel_refl_strong o #facts o rep_bnf;
464c5da3f508 more useful properties of the relators
traytel
parents: 61116
diff changeset
   635
val rel_reflp_of_bnf = Lazy.force o #rel_reflp o #facts o rep_bnf;
464c5da3f508 more useful properties of the relators
traytel
parents: 61116
diff changeset
   636
val rel_symp_of_bnf = Lazy.force o #rel_symp o #facts o rep_bnf;
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   637
val rel_transfer_of_bnf = Lazy.force o #rel_transfer o #facts o rep_bnf;
61240
464c5da3f508 more useful properties of the relators
traytel
parents: 61116
diff changeset
   638
val rel_transp_of_bnf = Lazy.force o #rel_transp o #facts o rep_bnf;
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   639
val set_bd_of_bnf = #set_bd o #axioms o rep_bnf;
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   640
val set_defs_of_bnf = #set_defs o #defs o rep_bnf;
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   641
val set_map0_of_bnf = #set_map0 o #axioms o rep_bnf;
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   642
val set_map_of_bnf = map Lazy.force o #set_map o #facts o rep_bnf;
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   643
val set_transfer_of_bnf = Lazy.force o #set_transfer o #facts o rep_bnf;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   644
val wit_thms_of_bnf = maps #prop o wits_of_bnf;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   645
val wit_thmss_of_bnf = map #prop o wits_of_bnf;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   646
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   647
fun mk_bnf name T live lives lives' dead deads map sets bd axioms defs facts wits rel pred =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   648
  BNF {name = name, T = T,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   649
       live = live, lives = lives, lives' = lives', dead = dead, deads = deads,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   650
       map = map, sets = sets, bd = bd,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   651
       axioms = axioms, defs = defs, facts = facts,
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   652
       nwits = length wits, wits = wits, rel = rel, pred = pred};
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   653
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   654
fun map_bnf f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17
56016
8875cdcfc85b unfold intermediate definitions after sealing the bnf
traytel
parents: 55945
diff changeset
   655
  (BNF {name = name, T = T, live = live, lives = lives, lives' = lives',
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   656
  dead = dead, deads = deads, map = map, sets = sets, bd = bd,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   657
  axioms = axioms, defs = defs, facts = facts,
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   658
  nwits = nwits, wits = wits, rel = rel, pred = pred}) =
56016
8875cdcfc85b unfold intermediate definitions after sealing the bnf
traytel
parents: 55945
diff changeset
   659
  BNF {name = f1 name, T = f2 T,
8875cdcfc85b unfold intermediate definitions after sealing the bnf
traytel
parents: 55945
diff changeset
   660
       live = f3 live, lives = f4 lives, lives' = f5 lives', dead = f6 dead, deads = f7 deads,
8875cdcfc85b unfold intermediate definitions after sealing the bnf
traytel
parents: 55945
diff changeset
   661
       map = f8 map, sets = f9 sets, bd = f10 bd,
8875cdcfc85b unfold intermediate definitions after sealing the bnf
traytel
parents: 55945
diff changeset
   662
       axioms = f11 axioms, defs = f12 defs, facts = f13 facts,
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   663
       nwits = f14 nwits, wits = f15 wits, rel = f16 rel, pred = f17 pred};
56016
8875cdcfc85b unfold intermediate definitions after sealing the bnf
traytel
parents: 55945
diff changeset
   664
8875cdcfc85b unfold intermediate definitions after sealing the bnf
traytel
parents: 55945
diff changeset
   665
fun morph_bnf phi =
8875cdcfc85b unfold intermediate definitions after sealing the bnf
traytel
parents: 55945
diff changeset
   666
  let
8875cdcfc85b unfold intermediate definitions after sealing the bnf
traytel
parents: 55945
diff changeset
   667
    val Tphi = Morphism.typ phi;
8875cdcfc85b unfold intermediate definitions after sealing the bnf
traytel
parents: 55945
diff changeset
   668
    val tphi = Morphism.term phi;
8875cdcfc85b unfold intermediate definitions after sealing the bnf
traytel
parents: 55945
diff changeset
   669
  in
8875cdcfc85b unfold intermediate definitions after sealing the bnf
traytel
parents: 55945
diff changeset
   670
    map_bnf (Morphism.binding phi) Tphi I (map Tphi) (map Tphi) I (map Tphi) tphi (map tphi) tphi
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   671
      (morph_axioms phi) (morph_defs phi) (morph_facts phi) I (map (morph_witness phi)) tphi tphi
56016
8875cdcfc85b unfold intermediate definitions after sealing the bnf
traytel
parents: 55945
diff changeset
   672
  end;
8875cdcfc85b unfold intermediate definitions after sealing the bnf
traytel
parents: 55945
diff changeset
   673
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   674
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 I;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   675
64440
0d31d1735104 tuned signature
traytel
parents: 63862
diff changeset
   676
fun permute_deads perm = map_bnf I I I I I I perm I I I I I I I I I I;
0d31d1735104 tuned signature
traytel
parents: 63862
diff changeset
   677
58175
2412a3369c30 tuned BNF database handling
blanchet
parents: 58116
diff changeset
   678
val transfer_bnf = morph_bnf o Morphism.transfer_morphism;
2412a3369c30 tuned BNF database handling
blanchet
parents: 58116
diff changeset
   679
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   680
structure Data = Generic_Data
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   681
(
51837
087498724486 lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
blanchet
parents: 51836
diff changeset
   682
  type T = bnf Symtab.table;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   683
  val empty = Symtab.empty;
55394
d5ebe055b599 more liberal merging of BNFs and constructor sugar
blanchet
parents: 55356
diff changeset
   684
  fun merge data : T = Symtab.merge (K true) data;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   685
);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   686
58116
1a9ac371e5a0 added theory-based getters for convenience
blanchet
parents: 58108
diff changeset
   687
fun bnf_of_generic context =
58175
2412a3369c30 tuned BNF database handling
blanchet
parents: 58116
diff changeset
   688
  Option.map (transfer_bnf (Context.theory_of context)) o Symtab.lookup (Data.get context);
58116
1a9ac371e5a0 added theory-based getters for convenience
blanchet
parents: 58108
diff changeset
   689
1a9ac371e5a0 added theory-based getters for convenience
blanchet
parents: 58108
diff changeset
   690
val bnf_of = bnf_of_generic o Context.Proof;
1a9ac371e5a0 added theory-based getters for convenience
blanchet
parents: 58108
diff changeset
   691
val bnf_of_global = bnf_of_generic o Context.Theory;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   692
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   693
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   694
(* Utilities *)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   695
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   696
fun normalize_set insts instA set =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   697
  let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   698
    val (T, T') = dest_funT (fastype_of set);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   699
    val A = fst (Term.dest_TVar (HOLogic.dest_setT T'));
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   700
    val params = Term.add_tvar_namesT T [];
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   701
  in Term.subst_TVars ((A :: params) ~~ (instA :: insts)) set end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   702
49507
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49506
diff changeset
   703
fun normalize_rel ctxt instTs instA instB rel =
49462
blanchet
parents: 49461
diff changeset
   704
  let
blanchet
parents: 49461
diff changeset
   705
    val thy = Proof_Context.theory_of ctxt;
blanchet
parents: 49461
diff changeset
   706
    val tyenv =
49507
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49506
diff changeset
   707
      Sign.typ_match thy (fastype_of rel, Library.foldr (op -->) (instTs, mk_pred2T instA instB))
49463
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49462
diff changeset
   708
        Vartab.empty;
49507
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49506
diff changeset
   709
  in Envir.subst_term (tyenv, Vartab.empty) rel end
49453
ff0e540d8758 add rel as first-class citizen of BNF
blanchet
parents: 49452
diff changeset
   710
  handle Type.TYPE_MATCH => error "Bad relator";
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   711
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   712
fun normalize_pred ctxt instTs instA pred =
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   713
  let
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   714
    val thy = Proof_Context.theory_of ctxt;
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   715
    val tyenv =
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   716
      Sign.typ_match thy (fastype_of pred, Library.foldr (op -->) (instTs, mk_pred1T instA))
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   717
        Vartab.empty;
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   718
  in Envir.subst_term (tyenv, Vartab.empty) pred end
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   719
  handle Type.TYPE_MATCH => error "Bad predicator";
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   720
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   721
fun normalize_wit insts CA As wit =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   722
  let
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67222
diff changeset
   723
    fun strip_param (Ts, T as Type (\<^type_name>\<open>fun\<close>, [T1, T2])) =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   724
        if Type.raw_instance (CA, T) then (Ts, T) else strip_param (T1 :: Ts, T2)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   725
      | strip_param x = x;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   726
    val (Ts, T) = strip_param ([], fastype_of wit);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   727
    val subst = Term.add_tvar_namesT T [] ~~ insts;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   728
    fun find y = find_index (fn x => x = y) As;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   729
  in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   730
    (map (find o Term.typ_subst_TVars subst) (rev Ts), Term.subst_TVars subst wit)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   731
  end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   732
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   733
fun minimize_wits wits =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   734
 let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   735
   fun minimize done [] = done
49103
3caaa80f53a4 generalized signature
traytel
parents: 49021
diff changeset
   736
     | minimize done ((I, wit) :: todo) =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   737
       if exists (fn (J, _) => subset (op =) (J, I)) (done @ todo)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   738
       then minimize done todo
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   739
       else minimize ((I, wit) :: done) todo;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   740
 in minimize [] wits end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   741
54236
e00009523727 moved code around
blanchet
parents: 54189
diff changeset
   742
fun mk_map live Ts Us t =
e00009523727 moved code around
blanchet
parents: 54189
diff changeset
   743
  let val (Type (_, Ts0), Type (_, Us0)) = strip_typeN (live + 1) (fastype_of t) |>> List.last in
e00009523727 moved code around
blanchet
parents: 54189
diff changeset
   744
    Term.subst_atomic_types (Ts0 @ Us0 ~~ Ts @ Us) t
e00009523727 moved code around
blanchet
parents: 54189
diff changeset
   745
  end;
e00009523727 moved code around
blanchet
parents: 54189
diff changeset
   746
62326
3cf7a067599c allow predicator instead of map function in 'primrec'
blanchet
parents: 62324
diff changeset
   747
fun mk_pred Ts t =
3cf7a067599c allow predicator instead of map function in 'primrec'
blanchet
parents: 62324
diff changeset
   748
  let val Type (_, Ts0) = domain_type (body_fun_type (fastype_of t)) in
3cf7a067599c allow predicator instead of map function in 'primrec'
blanchet
parents: 62324
diff changeset
   749
    Term.subst_atomic_types (Ts0 ~~ Ts) t
3cf7a067599c allow predicator instead of map function in 'primrec'
blanchet
parents: 62324
diff changeset
   750
  end;
63851
1a1fd3f3a24c prove 'set' property backward
blanchet
parents: 63714
diff changeset
   751
val mk_set = mk_pred;
62326
3cf7a067599c allow predicator instead of map function in 'primrec'
blanchet
parents: 62324
diff changeset
   752
54236
e00009523727 moved code around
blanchet
parents: 54189
diff changeset
   753
fun mk_rel live Ts Us t =
e00009523727 moved code around
blanchet
parents: 54189
diff changeset
   754
  let val [Type (_, Ts0), Type (_, Us0)] = binder_types (snd (strip_typeN live (fastype_of t))) in
e00009523727 moved code around
blanchet
parents: 54189
diff changeset
   755
    Term.subst_atomic_types (Ts0 @ Us0 ~~ Ts @ Us) t
e00009523727 moved code around
blanchet
parents: 54189
diff changeset
   756
  end;
e00009523727 moved code around
blanchet
parents: 54189
diff changeset
   757
64627
8d7cb22482e3 generalized ML function (towards nonuniform datatypes)
blanchet
parents: 64440
diff changeset
   758
fun build_map_or_rel mk const of_bnf dest pre_cst_table ctxt simple_Ts simple_Us build_simple =
54236
e00009523727 moved code around
blanchet
parents: 54189
diff changeset
   759
  let
e00009523727 moved code around
blanchet
parents: 54189
diff changeset
   760
    fun build (TU as (T, U)) =
64627
8d7cb22482e3 generalized ML function (towards nonuniform datatypes)
blanchet
parents: 64440
diff changeset
   761
      if exists (curry (op =) T) simple_Ts orelse exists (curry (op =) U) simple_Us then
57301
7b997028aaac generate 'rel_coinduct0' theorem for codatatypes
desharna
parents: 56954
diff changeset
   762
        build_simple TU
64627
8d7cb22482e3 generalized ML function (towards nonuniform datatypes)
blanchet
parents: 64440
diff changeset
   763
      else if T = U andalso not (exists_subtype_in simple_Ts T) andalso
8d7cb22482e3 generalized ML function (towards nonuniform datatypes)
blanchet
parents: 64440
diff changeset
   764
          not (exists_subtype_in simple_Us U) then
54236
e00009523727 moved code around
blanchet
parents: 54189
diff changeset
   765
        const T
e00009523727 moved code around
blanchet
parents: 54189
diff changeset
   766
      else
e00009523727 moved code around
blanchet
parents: 54189
diff changeset
   767
        (case TU of
e00009523727 moved code around
blanchet
parents: 54189
diff changeset
   768
          (Type (s, Ts), Type (s', Us)) =>
e00009523727 moved code around
blanchet
parents: 54189
diff changeset
   769
          if s = s' then
e00009523727 moved code around
blanchet
parents: 54189
diff changeset
   770
            let
62427
6dce7bf7960b generalized ML function
blanchet
parents: 62329
diff changeset
   771
              fun recurse (live, cst0) =
6dce7bf7960b generalized ML function
blanchet
parents: 62329
diff changeset
   772
                let
6dce7bf7960b generalized ML function
blanchet
parents: 62329
diff changeset
   773
                  val cst = mk live Ts Us cst0;
6dce7bf7960b generalized ML function
blanchet
parents: 62329
diff changeset
   774
                  val TUs' = map dest (fst (strip_typeN live (fastype_of cst)));
6dce7bf7960b generalized ML function
blanchet
parents: 62329
diff changeset
   775
                in Term.list_comb (cst, map build TUs') end;
6dce7bf7960b generalized ML function
blanchet
parents: 62329
diff changeset
   776
            in
6dce7bf7960b generalized ML function
blanchet
parents: 62329
diff changeset
   777
              (case AList.lookup (op =) pre_cst_table s of
6dce7bf7960b generalized ML function
blanchet
parents: 62329
diff changeset
   778
                NONE =>
6dce7bf7960b generalized ML function
blanchet
parents: 62329
diff changeset
   779
                (case bnf_of ctxt s of
6dce7bf7960b generalized ML function
blanchet
parents: 62329
diff changeset
   780
                  SOME bnf => recurse (live_of_bnf bnf, of_bnf bnf)
6dce7bf7960b generalized ML function
blanchet
parents: 62329
diff changeset
   781
                | NONE => build_simple TU)
6dce7bf7960b generalized ML function
blanchet
parents: 62329
diff changeset
   782
              | SOME entry => recurse entry)
6dce7bf7960b generalized ML function
blanchet
parents: 62329
diff changeset
   783
            end
58352
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents: 58344
diff changeset
   784
          else
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents: 58344
diff changeset
   785
            build_simple TU
54236
e00009523727 moved code around
blanchet
parents: 54189
diff changeset
   786
        | _ => build_simple TU);
e00009523727 moved code around
blanchet
parents: 54189
diff changeset
   787
  in build end;
e00009523727 moved code around
blanchet
parents: 54189
diff changeset
   788
62482
577199585ba0 generalized ML function
blanchet
parents: 62427
diff changeset
   789
val build_map = build_map_or_rel mk_map HOLogic.id_const map_of_bnf dest_funT
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67222
diff changeset
   790
  [(\<^type_name>\<open>set\<close>, (1, \<^term>\<open>image\<close>))];
58104
c5316f843f72 generate 'rel_transfer' for BNFs
desharna
parents: 58102
diff changeset
   791
val build_rel = build_map_or_rel mk_rel HOLogic.eq_const rel_of_bnf dest_pred2T o append
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67222
diff changeset
   792
  [(\<^type_name>\<open>set\<close>, (1, \<^term>\<open>rel_set\<close>)), (\<^type_name>\<open>fun\<close>, (2, \<^term>\<open>rel_fun\<close>))];
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   793
63851
1a1fd3f3a24c prove 'set' property backward
blanchet
parents: 63714
diff changeset
   794
fun build_set ctxt A =
1a1fd3f3a24c prove 'set' property backward
blanchet
parents: 63714
diff changeset
   795
  let
1a1fd3f3a24c prove 'set' property backward
blanchet
parents: 63714
diff changeset
   796
    fun build T =
1a1fd3f3a24c prove 'set' property backward
blanchet
parents: 63714
diff changeset
   797
      Abs (Name.uu, T,
1a1fd3f3a24c prove 'set' property backward
blanchet
parents: 63714
diff changeset
   798
        if T = A then
1a1fd3f3a24c prove 'set' property backward
blanchet
parents: 63714
diff changeset
   799
          HOLogic.mk_set A [Bound 0]
1a1fd3f3a24c prove 'set' property backward
blanchet
parents: 63714
diff changeset
   800
        else
1a1fd3f3a24c prove 'set' property backward
blanchet
parents: 63714
diff changeset
   801
          (case T of
1a1fd3f3a24c prove 'set' property backward
blanchet
parents: 63714
diff changeset
   802
            Type (s, Ts) =>
1a1fd3f3a24c prove 'set' property backward
blanchet
parents: 63714
diff changeset
   803
            let
1a1fd3f3a24c prove 'set' property backward
blanchet
parents: 63714
diff changeset
   804
              val sets = map (mk_set Ts) (sets_of_bnf (the (bnf_of ctxt s)))
1a1fd3f3a24c prove 'set' property backward
blanchet
parents: 63714
diff changeset
   805
                |> filter (exists_subtype_in [A] o range_type o fastype_of);
1a1fd3f3a24c prove 'set' property backward
blanchet
parents: 63714
diff changeset
   806
              val set_apps = map (fn set => Term.betapply (set, Bound 0)) sets;
1a1fd3f3a24c prove 'set' property backward
blanchet
parents: 63714
diff changeset
   807
1a1fd3f3a24c prove 'set' property backward
blanchet
parents: 63714
diff changeset
   808
              fun recurse set_app =
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67222
diff changeset
   809
                let val Type (\<^type_name>\<open>set\<close>, [elemT]) = fastype_of set_app in
63851
1a1fd3f3a24c prove 'set' property backward
blanchet
parents: 63714
diff changeset
   810
                  if elemT = A then set_app else mk_UNION set_app (build elemT)
1a1fd3f3a24c prove 'set' property backward
blanchet
parents: 63714
diff changeset
   811
                end;
1a1fd3f3a24c prove 'set' property backward
blanchet
parents: 63714
diff changeset
   812
            in
1a1fd3f3a24c prove 'set' property backward
blanchet
parents: 63714
diff changeset
   813
              if null set_apps then HOLogic.mk_set A []
63862
ce05cc93e07b union associates to the left
blanchet
parents: 63851
diff changeset
   814
              else Library.foldl1 mk_union (map recurse set_apps)
63851
1a1fd3f3a24c prove 'set' property backward
blanchet
parents: 63714
diff changeset
   815
            end
1a1fd3f3a24c prove 'set' property backward
blanchet
parents: 63714
diff changeset
   816
          | _ => HOLogic.mk_set A []));
1a1fd3f3a24c prove 'set' property backward
blanchet
parents: 63714
diff changeset
   817
  in build end;
1a1fd3f3a24c prove 'set' property backward
blanchet
parents: 63714
diff changeset
   818
54246
8fdb4dc08ed1 split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
blanchet
parents: 54237
diff changeset
   819
fun map_flattened_map_args ctxt s map_args fs =
8fdb4dc08ed1 split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
blanchet
parents: 54237
diff changeset
   820
  let
8fdb4dc08ed1 split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
blanchet
parents: 54237
diff changeset
   821
    val flat_fs = flatten_type_args_of_bnf (the (bnf_of ctxt s)) Term.dummy fs;
8fdb4dc08ed1 split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
blanchet
parents: 54237
diff changeset
   822
    val flat_fs' = map_args flat_fs;
8fdb4dc08ed1 split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
blanchet
parents: 54237
diff changeset
   823
  in
55480
59cc4a8bc28a allow different functions to recurse on the same type, like in the old package
blanchet
parents: 55444
diff changeset
   824
    permute_like_unique (op aconv) flat_fs fs flat_fs'
54246
8fdb4dc08ed1 split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
blanchet
parents: 54237
diff changeset
   825
  end;
8fdb4dc08ed1 split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
blanchet
parents: 54237
diff changeset
   826
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   827
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   828
(* Names *)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   829
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   830
val mapN = "map";
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   831
val setN = "set";
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   832
fun mk_setN i = setN ^ nonzero_string_of_int i;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   833
val bdN = "bd";
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   834
val witN = "wit";
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   835
fun mk_witN i = witN ^ nonzero_string_of_int i;
49507
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49506
diff changeset
   836
val relN = "rel";
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   837
val predN = "pred";
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   838
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   839
val bd_Card_orderN = "bd_Card_order";
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   840
val bd_CinfiniteN = "bd_Cinfinite";
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   841
val bd_CnotzeroN = "bd_Cnotzero";
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   842
val bd_card_orderN = "bd_card_order";
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   843
val bd_cinfiniteN = "bd_cinfinite";
75624
22d1c5f2b9f4 strict bounds for BNFs (by Jan van Brügge)
traytel
parents: 75276
diff changeset
   844
val bd_regularCardN = "bd_regularCard";
51766
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51765
diff changeset
   845
val collect_set_mapN = "collect_set_map";
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   846
val in_bdN = "in_bd";
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   847
val in_monoN = "in_mono";
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
   848
val in_relN = "in_rel";
56635
b07c8ad23010 added 'inj_map' as auxiliary BNF theorem
blanchet
parents: 56523
diff changeset
   849
val inj_mapN = "inj_map";
57970
eaa986cd285a generate 'inj_map_strong' for BNFs
desharna
parents: 57969
diff changeset
   850
val inj_map_strongN = "inj_map_strong";
53287
271b34513bfb renamed BNF axiom
blanchet
parents: 53285
diff changeset
   851
val map_comp0N = "map_comp0";
53288
050d0bc9afa5 renamed BNF fact
blanchet
parents: 53287
diff changeset
   852
val map_compN = "map_comp";
51761
4c9f08836d87 renamed "map_cong" axiom to "map_cong0" in preparation for real "map_cong"
blanchet
parents: 51758
diff changeset
   853
val map_cong0N = "map_cong0";
51762
219a3063ed29 derive "map_cong"
blanchet
parents: 51761
diff changeset
   854
val map_congN = "map_cong";
57981
81baacfd56e8 generate 'map_cong_simp' for BNFs
desharna
parents: 57970
diff changeset
   855
val map_cong_simpN = "map_cong_simp";
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   856
val map_cong_predN = "map_cong_pred";
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   857
val map_id0N = "map_id0";
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   858
val map_idN = "map_id";
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   859
val map_identN = "map_ident";
75276
686a6d7d0991 generated lemma map_ident_strong for BNFs
desharna
parents: 74664
diff changeset
   860
val map_ident_strongN = "map_ident_strong";
52719
480a3479fa47 transfer rule for map (not yet registered as a transfer rule)
traytel
parents: 52635
diff changeset
   861
val map_transferN = "map_transfer";
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   862
val pred_mono_strong0N = "pred_mono_strong0";
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   863
val pred_mono_strongN = "pred_mono_strong";
63714
b62f4f765353 derive pred_mono property for BNFs
traytel
parents: 63399
diff changeset
   864
val pred_monoN = "pred_mono";
62329
9f7fa08d2307 derive transfer rule for predicator
traytel
parents: 62326
diff changeset
   865
val pred_transferN = "pred_transfer";
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   866
val pred_TrueN = "pred_True";
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   867
val pred_mapN = "pred_map";
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   868
val pred_relN = "pred_rel";
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   869
val pred_setN = "pred_set";
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   870
val pred_congN = "pred_cong";
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   871
val pred_cong_simpN = "pred_cong_simp";
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   872
val rel_GrpN = "rel_Grp";
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   873
val rel_comppN = "rel_compp";
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   874
val rel_compp_GrpN = "rel_compp_Grp";
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   875
val rel_congN = "rel_cong";
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   876
val rel_cong_simpN = "rel_cong_simp";
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   877
val rel_conversepN = "rel_conversep";
49591
91b228e26348 generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents: 49585
diff changeset
   878
val rel_eqN = "rel_eq";
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   879
val rel_eq_onpN = "rel_eq_onp";
49537
fe1deee434b6 generate "rel_as_srel" and "rel_flip" properties
blanchet
parents: 49536
diff changeset
   880
val rel_flipN = "rel_flip";
59726
64c2bb331035 BNF relators preserve reflexivity
traytel
parents: 59663
diff changeset
   881
val rel_mapN = "rel_map";
64c2bb331035 BNF relators preserve reflexivity
traytel
parents: 59663
diff changeset
   882
val rel_monoN = "rel_mono";
64c2bb331035 BNF relators preserve reflexivity
traytel
parents: 59663
diff changeset
   883
val rel_mono_strong0N = "rel_mono_strong0";
64c2bb331035 BNF relators preserve reflexivity
traytel
parents: 59663
diff changeset
   884
val rel_mono_strongN = "rel_mono_strong";
64c2bb331035 BNF relators preserve reflexivity
traytel
parents: 59663
diff changeset
   885
val rel_reflN = "rel_refl";
61240
464c5da3f508 more useful properties of the relators
traytel
parents: 61116
diff changeset
   886
val rel_refl_strongN = "rel_refl_strong";
464c5da3f508 more useful properties of the relators
traytel
parents: 61116
diff changeset
   887
val rel_reflpN = "rel_reflp";
464c5da3f508 more useful properties of the relators
traytel
parents: 61116
diff changeset
   888
val rel_sympN = "rel_symp";
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   889
val rel_transferN = "rel_transfer";
61240
464c5da3f508 more useful properties of the relators
traytel
parents: 61116
diff changeset
   890
val rel_transpN = "rel_transp";
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   891
val set_bdN = "set_bd";
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   892
val set_map0N = "set_map0";
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   893
val set_mapN = "set_map";
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   894
val set_transferN = "set_transfer";
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   895
55854
ee270328a781 make 'typedef' optional, depending on size of original type
blanchet
parents: 55480
diff changeset
   896
datatype inline_policy = Dont_Inline | Hardly_Inline | Smart_Inline | Do_Inline;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   897
49538
c90818b63599 simplified fact policies
blanchet
parents: 49537
diff changeset
   898
datatype fact_policy = Dont_Note | Note_Some | Note_All;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   899
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67222
diff changeset
   900
val bnf_internals = Attrib.setup_config_bool \<^binding>\<open>bnf_internals\<close> (K false);
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67222
diff changeset
   901
val bnf_timing = Attrib.setup_config_bool \<^binding>\<open>bnf_timing\<close> (K false);
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   902
62093
bd73a2279fcd more uniform treatment of package internals;
wenzelm
parents: 61841
diff changeset
   903
fun user_policy policy ctxt = if Config.get ctxt bnf_internals then Note_All else policy;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   904
55854
ee270328a781 make 'typedef' optional, depending on size of original type
blanchet
parents: 55480
diff changeset
   905
val smart_max_inline_term_size = 25; (*FUDGE*)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   906
57631
959caab43a3d use the noted theorem whenever possible, also in 'BNF_Def'
blanchet
parents: 57567
diff changeset
   907
fun note_bnf_thms fact_policy qualify0 bnf_b bnf lthy =
52720
fdc04f9bf906 separate ML interface to note facts of a bnf
traytel
parents: 52719
diff changeset
   908
  let
fdc04f9bf906 separate ML interface to note facts of a bnf
traytel
parents: 52719
diff changeset
   909
    val axioms = axioms_of_bnf bnf;
fdc04f9bf906 separate ML interface to note facts of a bnf
traytel
parents: 52719
diff changeset
   910
    val facts = facts_of_bnf bnf;
fdc04f9bf906 separate ML interface to note facts of a bnf
traytel
parents: 52719
diff changeset
   911
    val wits = wits_of_bnf bnf;
54045
369a4a14583a keep the qualification of bindings when noting bnf theorems
traytel
parents: 53561
diff changeset
   912
    val qualify =
59858
890b68e1e8b6 support for strictly private name space entries;
wenzelm
parents: 59822
diff changeset
   913
      let val qs = Binding.path_of bnf_b;
56766
ba2fa4e99729 more reliable 'name_of_bnf'
blanchet
parents: 56657
diff changeset
   914
      in fold_rev (fn (s, mand) => Binding.qualify mand s) qs #> qualify0 end;
57631
959caab43a3d use the noted theorem whenever possible, also in 'BNF_Def'
blanchet
parents: 57567
diff changeset
   915
959caab43a3d use the noted theorem whenever possible, also in 'BNF_Def'
blanchet
parents: 57567
diff changeset
   916
    fun note_if_note_all (noted0, lthy0) =
52720
fdc04f9bf906 separate ML interface to note facts of a bnf
traytel
parents: 52719
diff changeset
   917
      let
fdc04f9bf906 separate ML interface to note facts of a bnf
traytel
parents: 52719
diff changeset
   918
        val witNs = if length wits = 1 then [witN] else map mk_witN (1 upto length wits);
fdc04f9bf906 separate ML interface to note facts of a bnf
traytel
parents: 52719
diff changeset
   919
        val notes =
fdc04f9bf906 separate ML interface to note facts of a bnf
traytel
parents: 52719
diff changeset
   920
          [(bd_card_orderN, [#bd_card_order axioms]),
57631
959caab43a3d use the noted theorem whenever possible, also in 'BNF_Def'
blanchet
parents: 57567
diff changeset
   921
           (bd_cinfiniteN, [#bd_cinfinite axioms]),
75624
22d1c5f2b9f4 strict bounds for BNFs (by Jan van Brügge)
traytel
parents: 75276
diff changeset
   922
           (bd_regularCardN, [#bd_regularCard axioms]),
57631
959caab43a3d use the noted theorem whenever possible, also in 'BNF_Def'
blanchet
parents: 57567
diff changeset
   923
           (bd_Card_orderN, [#bd_Card_order facts]),
959caab43a3d use the noted theorem whenever possible, also in 'BNF_Def'
blanchet
parents: 57567
diff changeset
   924
           (bd_CinfiniteN, [#bd_Cinfinite facts]),
959caab43a3d use the noted theorem whenever possible, also in 'BNF_Def'
blanchet
parents: 57567
diff changeset
   925
           (bd_CnotzeroN, [#bd_Cnotzero facts]),
959caab43a3d use the noted theorem whenever possible, also in 'BNF_Def'
blanchet
parents: 57567
diff changeset
   926
           (collect_set_mapN, [Lazy.force (#collect_set_map facts)]),
959caab43a3d use the noted theorem whenever possible, also in 'BNF_Def'
blanchet
parents: 57567
diff changeset
   927
           (in_bdN, [Lazy.force (#in_bd facts)]),
959caab43a3d use the noted theorem whenever possible, also in 'BNF_Def'
blanchet
parents: 57567
diff changeset
   928
           (in_monoN, [Lazy.force (#in_mono facts)]),
959caab43a3d use the noted theorem whenever possible, also in 'BNF_Def'
blanchet
parents: 57567
diff changeset
   929
           (map_comp0N, [#map_comp0 axioms]),
57967
e6d2e998c30f renamed 'rel_mono_strong' to 'rel_mono_strong0'
desharna
parents: 57932
diff changeset
   930
           (rel_mono_strong0N, [Lazy.force (#rel_mono_strong0 facts)]),
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   931
           (pred_mono_strong0N, [Lazy.force (#pred_mono_strong0 facts)]),
57631
959caab43a3d use the noted theorem whenever possible, also in 'BNF_Def'
blanchet
parents: 57567
diff changeset
   932
           (set_map0N, #set_map0 axioms),
959caab43a3d use the noted theorem whenever possible, also in 'BNF_Def'
blanchet
parents: 57567
diff changeset
   933
           (set_bdN, #set_bd axioms)] @
959caab43a3d use the noted theorem whenever possible, also in 'BNF_Def'
blanchet
parents: 57567
diff changeset
   934
          (witNs ~~ wit_thmss_of_bnf bnf)
959caab43a3d use the noted theorem whenever possible, also in 'BNF_Def'
blanchet
parents: 57567
diff changeset
   935
          |> map (fn (thmN, thms) =>
959caab43a3d use the noted theorem whenever possible, also in 'BNF_Def'
blanchet
parents: 57567
diff changeset
   936
            ((qualify (Binding.qualify true (Binding.name_of bnf_b) (Binding.name thmN)), []),
959caab43a3d use the noted theorem whenever possible, also in 'BNF_Def'
blanchet
parents: 57567
diff changeset
   937
             [(thms, [])]));
959caab43a3d use the noted theorem whenever possible, also in 'BNF_Def'
blanchet
parents: 57567
diff changeset
   938
      in
959caab43a3d use the noted theorem whenever possible, also in 'BNF_Def'
blanchet
parents: 57567
diff changeset
   939
        Local_Theory.notes notes lthy0 |>> append noted0
58104
c5316f843f72 generate 'rel_transfer' for BNFs
desharna
parents: 58102
diff changeset
   940
      end;
57631
959caab43a3d use the noted theorem whenever possible, also in 'BNF_Def'
blanchet
parents: 57567
diff changeset
   941
959caab43a3d use the noted theorem whenever possible, also in 'BNF_Def'
blanchet
parents: 57567
diff changeset
   942
    fun note_unless_dont_note (noted0, lthy0) =
959caab43a3d use the noted theorem whenever possible, also in 'BNF_Def'
blanchet
parents: 57567
diff changeset
   943
      let
959caab43a3d use the noted theorem whenever possible, also in 'BNF_Def'
blanchet
parents: 57567
diff changeset
   944
        val notes =
59133
347fece4a85e note more facts (always)
traytel
parents: 59058
diff changeset
   945
          [(in_relN, [Lazy.force (#in_rel facts)], []),
347fece4a85e note more facts (always)
traytel
parents: 59058
diff changeset
   946
           (inj_mapN, [Lazy.force (#inj_map facts)], []),
57970
eaa986cd285a generate 'inj_map_strong' for BNFs
desharna
parents: 57969
diff changeset
   947
           (inj_map_strongN, [Lazy.force (#inj_map_strong facts)], []),
57969
1e7b9579b14f note 'inj_map' more often
desharna
parents: 57968
diff changeset
   948
           (map_compN, [Lazy.force (#map_comp facts)], []),
57631
959caab43a3d use the noted theorem whenever possible, also in 'BNF_Def'
blanchet
parents: 57567
diff changeset
   949
           (map_cong0N, [#map_cong0 axioms], []),
959caab43a3d use the noted theorem whenever possible, also in 'BNF_Def'
blanchet
parents: 57567
diff changeset
   950
           (map_congN, [Lazy.force (#map_cong facts)], fundefcong_attrs),
57981
81baacfd56e8 generate 'map_cong_simp' for BNFs
desharna
parents: 57970
diff changeset
   951
           (map_cong_simpN, [Lazy.force (#map_cong_simp facts)], []),
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   952
           (map_cong_predN, [Lazy.force (#map_cong_pred facts)], []),
57631
959caab43a3d use the noted theorem whenever possible, also in 'BNF_Def'
blanchet
parents: 57567
diff changeset
   953
           (map_idN, [Lazy.force (#map_id facts)], []),
959caab43a3d use the noted theorem whenever possible, also in 'BNF_Def'
blanchet
parents: 57567
diff changeset
   954
           (map_id0N, [#map_id0 axioms], []),
58102
73f46283c247 note 'map_transfer' more often
desharna
parents: 58093
diff changeset
   955
           (map_transferN, [Lazy.force (#map_transfer facts)], []),
57631
959caab43a3d use the noted theorem whenever possible, also in 'BNF_Def'
blanchet
parents: 57567
diff changeset
   956
           (map_identN, [Lazy.force (#map_ident facts)], []),
75276
686a6d7d0991 generated lemma map_ident_strong for BNFs
desharna
parents: 74664
diff changeset
   957
           (map_ident_strongN, [Lazy.force (#map_ident_strong facts)], []),
74664
d4ef127b74df added "mono" attribute to BNF generated pred_mono theorems
desharna
parents: 74561
diff changeset
   958
           (pred_monoN, [Lazy.force (#pred_mono facts)], mono_attrs),
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   959
           (pred_mono_strongN, [Lazy.force (#pred_mono_strong facts)], []),
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   960
           (pred_congN, [Lazy.force (#pred_cong facts)], fundefcong_attrs),
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   961
           (pred_cong_simpN, [Lazy.force (#pred_cong_simp facts)], []),
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   962
           (pred_mapN, [Lazy.force (#pred_map facts)], []),
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   963
           (pred_relN, [Lazy.force (#pred_rel facts)], []),
62329
9f7fa08d2307 derive transfer rule for predicator
traytel
parents: 62326
diff changeset
   964
           (pred_transferN, [Lazy.force (#pred_transfer facts)], []),
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   965
           (pred_TrueN, [Lazy.force (#pred_True facts)], []),
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   966
           (pred_setN, [#pred_set axioms], []),
57631
959caab43a3d use the noted theorem whenever possible, also in 'BNF_Def'
blanchet
parents: 57567
diff changeset
   967
           (rel_comppN, [Lazy.force (#rel_OO facts)], []),
959caab43a3d use the noted theorem whenever possible, also in 'BNF_Def'
blanchet
parents: 57567
diff changeset
   968
           (rel_compp_GrpN, no_refl [#rel_OO_Grp axioms], []),
959caab43a3d use the noted theorem whenever possible, also in 'BNF_Def'
blanchet
parents: 57567
diff changeset
   969
           (rel_conversepN, [Lazy.force (#rel_conversep facts)], []),
959caab43a3d use the noted theorem whenever possible, also in 'BNF_Def'
blanchet
parents: 57567
diff changeset
   970
           (rel_eqN, [Lazy.force (#rel_eq facts)], []),
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
   971
           (rel_eq_onpN, [Lazy.force (#rel_eq_onp facts)], []),
57631
959caab43a3d use the noted theorem whenever possible, also in 'BNF_Def'
blanchet
parents: 57567
diff changeset
   972
           (rel_flipN, [Lazy.force (#rel_flip facts)], []),
959caab43a3d use the noted theorem whenever possible, also in 'BNF_Def'
blanchet
parents: 57567
diff changeset
   973
           (rel_GrpN, [Lazy.force (#rel_Grp facts)], []),
57932
c29659f77f8d generate 'rel_map' theorem for BNFs
desharna
parents: 57668
diff changeset
   974
           (rel_mapN, Lazy.force (#rel_map facts), []),
58344
ea3d989219b4 set 'mono' attribute on 'rel_mono'
blanchet
parents: 58189
diff changeset
   975
           (rel_monoN, [Lazy.force (#rel_mono facts)], mono_attrs),
59133
347fece4a85e note more facts (always)
traytel
parents: 59058
diff changeset
   976
           (rel_mono_strongN, [Lazy.force (#rel_mono_strong facts)], []),
61242
1f92b0ac9c96 congruence rules for the relator
traytel
parents: 61241
diff changeset
   977
           (rel_congN, [Lazy.force (#rel_cong facts)], fundefcong_attrs),
1f92b0ac9c96 congruence rules for the relator
traytel
parents: 61241
diff changeset
   978
           (rel_cong_simpN, [Lazy.force (#rel_cong_simp facts)], []),
59726
64c2bb331035 BNF relators preserve reflexivity
traytel
parents: 59663
diff changeset
   979
           (rel_reflN, [Lazy.force (#rel_refl facts)], []),
61240
464c5da3f508 more useful properties of the relators
traytel
parents: 61116
diff changeset
   980
           (rel_refl_strongN, [Lazy.force (#rel_refl_strong facts)], []),
464c5da3f508 more useful properties of the relators
traytel
parents: 61116
diff changeset
   981
           (rel_reflpN, [Lazy.force (#rel_reflp facts)], []),
464c5da3f508 more useful properties of the relators
traytel
parents: 61116
diff changeset
   982
           (rel_sympN, [Lazy.force (#rel_symp facts)], []),
464c5da3f508 more useful properties of the relators
traytel
parents: 61116
diff changeset
   983
           (rel_transpN, [Lazy.force (#rel_transp facts)], []),
58104
c5316f843f72 generate 'rel_transfer' for BNFs
desharna
parents: 58102
diff changeset
   984
           (rel_transferN, [Lazy.force (#rel_transfer facts)], []),
58106
c8cba801c483 generate 'set_transfer' for BNFs
desharna
parents: 58104
diff changeset
   985
           (set_mapN, map Lazy.force (#set_map facts), []),
c8cba801c483 generate 'set_transfer' for BNFs
desharna
parents: 58104
diff changeset
   986
           (set_transferN, Lazy.force (#set_transfer facts), [])]
57631
959caab43a3d use the noted theorem whenever possible, also in 'BNF_Def'
blanchet
parents: 57567
diff changeset
   987
          |> filter_out (null o #2)
959caab43a3d use the noted theorem whenever possible, also in 'BNF_Def'
blanchet
parents: 57567
diff changeset
   988
          |> map (fn (thmN, thms, attrs) =>
959caab43a3d use the noted theorem whenever possible, also in 'BNF_Def'
blanchet
parents: 57567
diff changeset
   989
            ((qualify (Binding.qualify true (Binding.name_of bnf_b) (Binding.name thmN)), attrs),
959caab43a3d use the noted theorem whenever possible, also in 'BNF_Def'
blanchet
parents: 57567
diff changeset
   990
             [(thms, [])]));
959caab43a3d use the noted theorem whenever possible, also in 'BNF_Def'
blanchet
parents: 57567
diff changeset
   991
      in
959caab43a3d use the noted theorem whenever possible, also in 'BNF_Def'
blanchet
parents: 57567
diff changeset
   992
        Local_Theory.notes notes lthy0 |>> append noted0
58104
c5316f843f72 generate 'rel_transfer' for BNFs
desharna
parents: 58102
diff changeset
   993
      end;
57631
959caab43a3d use the noted theorem whenever possible, also in 'BNF_Def'
blanchet
parents: 57567
diff changeset
   994
  in
959caab43a3d use the noted theorem whenever possible, also in 'BNF_Def'
blanchet
parents: 57567
diff changeset
   995
    ([], lthy)
959caab43a3d use the noted theorem whenever possible, also in 'BNF_Def'
blanchet
parents: 57567
diff changeset
   996
    |> fact_policy = Note_All ? note_if_note_all
959caab43a3d use the noted theorem whenever possible, also in 'BNF_Def'
blanchet
parents: 57567
diff changeset
   997
    |> fact_policy <> Dont_Note ? note_unless_dont_note
959caab43a3d use the noted theorem whenever possible, also in 'BNF_Def'
blanchet
parents: 57567
diff changeset
   998
    |>> (fn [] => bnf | noted => morph_bnf (substitute_noted_thm noted) bnf)
52720
fdc04f9bf906 separate ML interface to note facts of a bnf
traytel
parents: 52719
diff changeset
   999
  end;
fdc04f9bf906 separate ML interface to note facts of a bnf
traytel
parents: 52719
diff changeset
  1000
66272
c6714a9562ae store the unfolded definitions of the lifted bnf constants under "_def" name
traytel
parents: 66198
diff changeset
  1001
fun note_bnf_defs bnf lthy =
c6714a9562ae store the unfolded definitions of the lifted bnf constants under "_def" name
traytel
parents: 66198
diff changeset
  1002
  let
c6714a9562ae store the unfolded definitions of the lifted bnf constants under "_def" name
traytel
parents: 66198
diff changeset
  1003
    fun mk_def_binding cst_of =
c6714a9562ae store the unfolded definitions of the lifted bnf constants under "_def" name
traytel
parents: 66198
diff changeset
  1004
      Thm.def_binding (Binding.qualified_name (dest_Const (cst_of bnf) |> fst));
c6714a9562ae store the unfolded definitions of the lifted bnf constants under "_def" name
traytel
parents: 66198
diff changeset
  1005
    val notes =
c6714a9562ae store the unfolded definitions of the lifted bnf constants under "_def" name
traytel
parents: 66198
diff changeset
  1006
      [(mk_def_binding map_of_bnf, map_def_of_bnf bnf),
c6714a9562ae store the unfolded definitions of the lifted bnf constants under "_def" name
traytel
parents: 66198
diff changeset
  1007
       (mk_def_binding rel_of_bnf, rel_def_of_bnf bnf),
c6714a9562ae store the unfolded definitions of the lifted bnf constants under "_def" name
traytel
parents: 66198
diff changeset
  1008
       (mk_def_binding pred_of_bnf, pred_def_of_bnf bnf)] @
c6714a9562ae store the unfolded definitions of the lifted bnf constants under "_def" name
traytel
parents: 66198
diff changeset
  1009
      @{map 2} (pair o mk_def_binding o K) (sets_of_bnf bnf) (set_defs_of_bnf bnf)
c6714a9562ae store the unfolded definitions of the lifted bnf constants under "_def" name
traytel
parents: 66198
diff changeset
  1010
      |> map (fn (b, thm) => ((b, []), [([thm], [])]));
c6714a9562ae store the unfolded definitions of the lifted bnf constants under "_def" name
traytel
parents: 66198
diff changeset
  1011
  in
c6714a9562ae store the unfolded definitions of the lifted bnf constants under "_def" name
traytel
parents: 66198
diff changeset
  1012
    lthy
c6714a9562ae store the unfolded definitions of the lifted bnf constants under "_def" name
traytel
parents: 66198
diff changeset
  1013
    |> Local_Theory.notes notes
c6714a9562ae store the unfolded definitions of the lifted bnf constants under "_def" name
traytel
parents: 66198
diff changeset
  1014
    |>> (fn noted => morph_bnf (substitute_noted_thm noted) bnf)
c6714a9562ae store the unfolded definitions of the lifted bnf constants under "_def" name
traytel
parents: 66198
diff changeset
  1015
  end;
c6714a9562ae store the unfolded definitions of the lifted bnf constants under "_def" name
traytel
parents: 66198
diff changeset
  1016
60918
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents: 60801
diff changeset
  1017
fun mk_wit_goals zs bs sets (I, wit) =
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents: 60801
diff changeset
  1018
  let
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents: 60801
diff changeset
  1019
    val xs = map (nth bs) I;
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents: 60801
diff changeset
  1020
    fun wit_goal i =
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents: 60801
diff changeset
  1021
      let
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents: 60801
diff changeset
  1022
        val z = nth zs i;
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents: 60801
diff changeset
  1023
        val set_wit = nth sets i $ Term.list_comb (wit, xs);
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents: 60801
diff changeset
  1024
        val concl = HOLogic.mk_Trueprop
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67222
diff changeset
  1025
          (if member (op =) I i then HOLogic.mk_eq (z, nth bs i) else \<^term>\<open>False\<close>);
60918
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents: 60801
diff changeset
  1026
      in
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents: 60801
diff changeset
  1027
        fold_rev Logic.all (z :: xs) (Logic.mk_implies (mk_Trueprop_mem (z, set_wit), concl))
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents: 60801
diff changeset
  1028
      end;
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents: 60801
diff changeset
  1029
  in
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents: 60801
diff changeset
  1030
    map wit_goal (0 upto length sets - 1)
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents: 60801
diff changeset
  1031
  end;
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents: 60801
diff changeset
  1032
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1033
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1034
(* Define new BNFs *)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1035
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1036
fun define_bnf_consts const_policy fact_policy internal Ds_opt map_b rel_b pred_b set_bs
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1037
    (((((((bnf_b, T_rhs), map_rhs), set_rhss), bd_rhs), wit_rhss), rel_rhs_opt), pred_rhs_opt)
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1038
    no_defs_lthy =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1039
  let
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1040
    val live = length set_rhss;
53265
cc9a2976f836 qualify BNF constants properly
blanchet
parents: 53261
diff changeset
  1041
61241
69a97fc33f7a conceal only the definitional theorems of map, set, rel (and not the actual constants)
traytel
parents: 61240
diff changeset
  1042
    val def_qualify = Binding.qualify false (Binding.name_of bnf_b);
53265
cc9a2976f836 qualify BNF constants properly
blanchet
parents: 53261
diff changeset
  1043
54490
930409d43211 use suffix '_list' etc. instead of prefix 'list_' for constants not present in the old package
blanchet
parents: 54426
diff changeset
  1044
    fun mk_prefix_binding pre = Binding.prefix_name (pre ^ "_") bnf_b;
49434
433dc7e028c8 separated registration of BNFs from bnf_def (BNFs are now stored only for bnf_def and (co)data commands)
traytel
parents: 49430
diff changeset
  1045
49463
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49462
diff changeset
  1046
    fun maybe_define user_specified (b, rhs) lthy =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1047
      let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1048
        val inline =
49538
c90818b63599 simplified fact policies
blanchet
parents: 49537
diff changeset
  1049
          (user_specified orelse fact_policy = Dont_Note) andalso
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1050
          (case const_policy of
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1051
            Dont_Inline => false
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1052
          | Hardly_Inline => Term.is_Free rhs orelse Term.is_Const rhs
55854
ee270328a781 make 'typedef' optional, depending on size of original type
blanchet
parents: 55480
diff changeset
  1053
          | Smart_Inline => Term.size_of_term rhs <= smart_max_inline_term_size
60154
7478de1f5b59 allow sorts on dead variables in BNFs
blanchet
parents: 59936
diff changeset
  1054
          | Do_Inline => true);
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1055
      in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1056
        if inline then
49460
4dd451a075ce fixed infinite loop with trivial rel_O_Gr + tuning
blanchet
parents: 49459
diff changeset
  1057
          ((rhs, Drule.reflexive_thm), lthy)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1058
        else
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1059
          let val b = b () in
56016
8875cdcfc85b unfold intermediate definitions after sealing the bnf
traytel
parents: 55945
diff changeset
  1060
            apfst (apsnd snd)
8875cdcfc85b unfold intermediate definitions after sealing the bnf
traytel
parents: 55945
diff changeset
  1061
              ((if internal then Local_Theory.define_internal else Local_Theory.define)
61241
69a97fc33f7a conceal only the definitional theorems of map, set, rel (and not the actual constants)
traytel
parents: 61240
diff changeset
  1062
                ((b, NoSyn), ((Binding.concealed (Thm.def_binding b), []), rhs)) lthy)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1063
          end
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1064
      end;
49459
3f8e2b5249ec adapted FP code to new relator approach
blanchet
parents: 49458
diff changeset
  1065
51758
55963309557b honor user-specified name for map function
blanchet
parents: 51757
diff changeset
  1066
    val map_bind_def =
54490
930409d43211 use suffix '_list' etc. instead of prefix 'list_' for constants not present in the old package
blanchet
parents: 54426
diff changeset
  1067
      (fn () => def_qualify (if Binding.is_empty map_b then mk_prefix_binding mapN else map_b),
53265
cc9a2976f836 qualify BNF constants properly
blanchet
parents: 53261
diff changeset
  1068
         map_rhs);
49459
3f8e2b5249ec adapted FP code to new relator approach
blanchet
parents: 49458
diff changeset
  1069
    val set_binds_defs =
3f8e2b5249ec adapted FP code to new relator approach
blanchet
parents: 49458
diff changeset
  1070
      let
51757
7babcb61aa5c honor user-specified set function names
blanchet
parents: 51551
diff changeset
  1071
        fun set_name i get_b =
7babcb61aa5c honor user-specified set function names
blanchet
parents: 51551
diff changeset
  1072
          (case try (nth set_bs) (i - 1) of
7babcb61aa5c honor user-specified set function names
blanchet
parents: 51551
diff changeset
  1073
            SOME b => if Binding.is_empty b then get_b else K b
53265
cc9a2976f836 qualify BNF constants properly
blanchet
parents: 53261
diff changeset
  1074
          | NONE => get_b) #> def_qualify;
54490
930409d43211 use suffix '_list' etc. instead of prefix 'list_' for constants not present in the old package
blanchet
parents: 54426
diff changeset
  1075
        val bs = if live = 1 then [set_name 1 (fn () => mk_prefix_binding setN)]
930409d43211 use suffix '_list' etc. instead of prefix 'list_' for constants not present in the old package
blanchet
parents: 54426
diff changeset
  1076
          else map (fn i => set_name i (fn () => mk_prefix_binding (mk_setN i))) (1 upto live);
51757
7babcb61aa5c honor user-specified set function names
blanchet
parents: 51551
diff changeset
  1077
      in bs ~~ set_rhss end;
54490
930409d43211 use suffix '_list' etc. instead of prefix 'list_' for constants not present in the old package
blanchet
parents: 54426
diff changeset
  1078
    val bd_bind_def = (fn () => def_qualify (mk_prefix_binding bdN), bd_rhs);
49459
3f8e2b5249ec adapted FP code to new relator approach
blanchet
parents: 49458
diff changeset
  1079
76383
fc35dc967344 separate out definition of bound to avoid spurious sort hypotheses (by Jan van Brügge)
traytel
parents: 75624
diff changeset
  1080
    val (((bnf_map_term, raw_map_def),
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1081
      (bnf_set_terms, raw_set_defs)),
76383
fc35dc967344 separate out definition of bound to avoid spurious sort hypotheses (by Jan van Brügge)
traytel
parents: 75624
diff changeset
  1082
      (lthy, lthy_old)) =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1083
        no_defs_lthy
72450
24bd1316eaae consolidated names and operations
haftmann
parents: 72154
diff changeset
  1084
        |> (snd o Local_Theory.begin_nested)
49463
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49462
diff changeset
  1085
        |> maybe_define true map_bind_def
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49462
diff changeset
  1086
        ||>> apfst split_list o fold_map (maybe_define true) set_binds_defs
72450
24bd1316eaae consolidated names and operations
haftmann
parents: 72154
diff changeset
  1087
        ||> `Local_Theory.end_nested;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1088
49459
3f8e2b5249ec adapted FP code to new relator approach
blanchet
parents: 49458
diff changeset
  1089
    val phi = Proof_Context.export_morphism lthy_old lthy;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1090
76383
fc35dc967344 separate out definition of bound to avoid spurious sort hypotheses (by Jan van Brügge)
traytel
parents: 75624
diff changeset
  1091
    val ((bnf_bd_term, raw_bd_def), (lthy, lthy_old)) =
fc35dc967344 separate out definition of bound to avoid spurious sort hypotheses (by Jan van Brügge)
traytel
parents: 75624
diff changeset
  1092
      lthy
fc35dc967344 separate out definition of bound to avoid spurious sort hypotheses (by Jan van Brügge)
traytel
parents: 75624
diff changeset
  1093
      |> (snd o Local_Theory.begin_nested)
fc35dc967344 separate out definition of bound to avoid spurious sort hypotheses (by Jan van Brügge)
traytel
parents: 75624
diff changeset
  1094
      |> maybe_define true bd_bind_def
fc35dc967344 separate out definition of bound to avoid spurious sort hypotheses (by Jan van Brügge)
traytel
parents: 75624
diff changeset
  1095
      ||> `Local_Theory.end_nested;
fc35dc967344 separate out definition of bound to avoid spurious sort hypotheses (by Jan van Brügge)
traytel
parents: 75624
diff changeset
  1096
fc35dc967344 separate out definition of bound to avoid spurious sort hypotheses (by Jan van Brügge)
traytel
parents: 75624
diff changeset
  1097
    val phi' = Proof_Context.export_morphism lthy_old lthy;
fc35dc967344 separate out definition of bound to avoid spurious sort hypotheses (by Jan van Brügge)
traytel
parents: 75624
diff changeset
  1098
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1099
    val bnf_map_def = Morphism.thm phi raw_map_def;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1100
    val bnf_set_defs = map (Morphism.thm phi) raw_set_defs;
76383
fc35dc967344 separate out definition of bound to avoid spurious sort hypotheses (by Jan van Brügge)
traytel
parents: 75624
diff changeset
  1101
    val bnf_bd_def = Morphism.thm phi' raw_bd_def;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1102
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1103
    val bnf_map = Morphism.term phi bnf_map_term;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1104
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1105
    (*TODO: handle errors*)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1106
    (*simple shape analysis of a map function*)
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1107
    val ((alphas, betas), (Calpha, _)) =
49395
323414474c1f use strip_typeN in bnf_def (instead of repairing strip_type)
traytel
parents: 49386
diff changeset
  1108
      fastype_of bnf_map
323414474c1f use strip_typeN in bnf_def (instead of repairing strip_type)
traytel
parents: 49386
diff changeset
  1109
      |> strip_typeN live
323414474c1f use strip_typeN in bnf_def (instead of repairing strip_type)
traytel
parents: 49386
diff changeset
  1110
      |>> map_split dest_funT
323414474c1f use strip_typeN in bnf_def (instead of repairing strip_type)
traytel
parents: 49386
diff changeset
  1111
      ||> dest_funT
323414474c1f use strip_typeN in bnf_def (instead of repairing strip_type)
traytel
parents: 49386
diff changeset
  1112
      handle TYPE _ => error "Bad map function";
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1113
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1114
    val Calpha_params = map TVar (Term.add_tvarsT Calpha []);
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1115
54426
edb87aac9cca prohibit locally fixed type variables in bnf definitions
traytel
parents: 54425
diff changeset
  1116
    val bnf_T = Morphism.typ phi T_rhs;
edb87aac9cca prohibit locally fixed type variables in bnf definitions
traytel
parents: 54425
diff changeset
  1117
    val bad_args = Term.add_tfreesT bnf_T [];
59281
1b4dc8a9f7d9 added plugins syntax to prim(co)rec
blanchet
parents: 59133
diff changeset
  1118
    val _ = null bad_args orelse error ("Locally fixed type arguments " ^
54426
edb87aac9cca prohibit locally fixed type variables in bnf definitions
traytel
parents: 54425
diff changeset
  1119
      commas_quote (map (Syntax.string_of_typ no_defs_lthy o TFree) bad_args));
edb87aac9cca prohibit locally fixed type variables in bnf definitions
traytel
parents: 54425
diff changeset
  1120
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1121
    val bnf_sets =
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1122
      map2 (normalize_set Calpha_params) alphas (map (Morphism.term phi) bnf_set_terms);
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1123
    val bnf_bd =
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1124
      Term.subst_TVars (Term.add_tvar_namesT bnf_T [] ~~ Calpha_params)
76383
fc35dc967344 separate out definition of bound to avoid spurious sort hypotheses (by Jan van Brügge)
traytel
parents: 75624
diff changeset
  1125
        (Morphism.term phi' bnf_bd_term);
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1126
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1127
    (*TODO: assert Ds = (TVars of bnf_map) \ (alphas @ betas) as sets*)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1128
    val deads = (case Ds_opt of
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1129
      NONE => subtract (op =) (alphas @ betas) (map TVar (Term.add_tvars bnf_map []))
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1130
    | SOME Ds => map (Morphism.typ phi) Ds);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1131
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1132
    (*TODO: further checks of type of bnf_map*)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1133
    (*TODO: check types of bnf_sets*)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1134
    (*TODO: check type of bnf_bd*)
49507
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49506
diff changeset
  1135
    (*TODO: check type of bnf_rel*)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1136
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1137
    fun mk_bnf_map Ds As' Bs' =
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1138
      Term.subst_atomic_types ((deads ~~ Ds) @ (alphas ~~ As') @ (betas ~~ Bs')) bnf_map;
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1139
    fun mk_bnf_t Ds As' = Term.subst_atomic_types ((deads ~~ Ds) @ (alphas ~~ As'));
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1140
    fun mk_bnf_T Ds As' = Term.typ_subst_atomic ((deads ~~ Ds) @ (alphas ~~ As'));
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1141
60154
7478de1f5b59 allow sorts on dead variables in BNFs
blanchet
parents: 59936
diff changeset
  1142
    val (((As, Bs), unsorted_Ds), names_lthy) = lthy
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1143
      |> mk_TFrees live
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1144
      ||>> mk_TFrees live
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1145
      ||>> mk_TFrees (length deads);
60154
7478de1f5b59 allow sorts on dead variables in BNFs
blanchet
parents: 59936
diff changeset
  1146
7478de1f5b59 allow sorts on dead variables in BNFs
blanchet
parents: 59936
diff changeset
  1147
    val Ds = map2 (resort_tfree_or_tvar o Type.sort_of_atyp) deads unsorted_Ds;
7478de1f5b59 allow sorts on dead variables in BNFs
blanchet
parents: 59936
diff changeset
  1148
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1149
    val RTs = map2 (curry HOLogic.mk_prodT) As Bs;
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1150
    val pred2RTs = map2 mk_pred2T As Bs;
60154
7478de1f5b59 allow sorts on dead variables in BNFs
blanchet
parents: 59936
diff changeset
  1151
    val (Rs, Rs') = names_lthy |> mk_Frees' "R" pred2RTs |> fst;
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1152
    val CA = mk_bnf_T Ds As Calpha;
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1153
    val CR = mk_bnf_T Ds RTs Calpha;
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1154
    val setRs =
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58432
diff changeset
  1155
      @{map 3} (fn R => fn T => fn U =>
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61334
diff changeset
  1156
          HOLogic.Collect_const (HOLogic.mk_prodT (T, U)) $ HOLogic.mk_case_prod R) Rs As Bs;
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1157
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1158
    (*Grp (in (Collect (split R1) .. Collect (split Rn))) (map fst .. fst)^--1 OO
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1159
      Grp (in (Collect (split R1) .. Collect (split Rn))) (map snd .. snd)*)
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1160
    val rel_spec =
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1161
      let
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1162
        val map1 = Term.list_comb (mk_bnf_map Ds RTs As, map fst_const RTs);
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1163
        val map2 = Term.list_comb (mk_bnf_map Ds RTs Bs, map snd_const RTs);
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1164
        val bnf_in = mk_in setRs (map (mk_bnf_t Ds RTs) bnf_sets) CR;
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1165
      in
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1166
        mk_rel_compp (mk_conversep (mk_Grp bnf_in map1), mk_Grp bnf_in map2)
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1167
        |> fold_rev Term.absfree Rs'
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1168
      end;
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1169
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1170
    val rel_rhs = the_default rel_spec rel_rhs_opt;
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1171
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1172
    val rel_bind_def =
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1173
      (fn () => def_qualify (if Binding.is_empty rel_b then mk_prefix_binding relN else rel_b),
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1174
         rel_rhs);
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1175
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1176
    val pred_spec =
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67222
diff changeset
  1177
      if live = 0 then Term.absdummy (mk_bnf_T Ds As Calpha) \<^term>\<open>True\<close> else
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1178
      let
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1179
        val sets = map (mk_bnf_t Ds As) bnf_sets;
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1180
        val argTs = map mk_pred1T As;
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1181
        val T = mk_bnf_T Ds As Calpha;
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1182
        val ((Ps, Ps'), x) = lthy
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1183
          |> mk_Frees' "P" argTs
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1184
          ||>> yield_singleton (mk_Frees "x") T
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1185
          |> fst;
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1186
        val conjs = map2 (fn set => fn P => mk_Ball (set $ x) P) sets Ps;
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1187
      in
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1188
        fold_rev Term.absfree Ps'
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1189
          (Term.absfree (dest_Free x) (Library.foldr1 HOLogic.mk_conj conjs))
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1190
      end;
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1191
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1192
    val pred_rhs = the_default pred_spec pred_rhs_opt;
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1193
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1194
    val pred_bind_def =
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1195
      (fn () => def_qualify (if Binding.is_empty pred_b then mk_prefix_binding predN else pred_b),
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1196
         pred_rhs);
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1197
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1198
    val wit_rhss =
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1199
      if null wit_rhss then
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1200
        [fold_rev Term.absdummy As (Term.list_comb (mk_bnf_map Ds As As,
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1201
          map2 (fn T => fn i => Term.absdummy T (Bound i)) As (live downto 1)) $
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67222
diff changeset
  1202
          Const (\<^const_name>\<open>undefined\<close>, CA))]
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1203
      else wit_rhss;
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1204
    val nwits = length wit_rhss;
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1205
    val wit_binds_defs =
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1206
      let
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1207
        val bs = if nwits = 1 then [fn () => def_qualify (mk_prefix_binding witN)]
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1208
          else map (fn i => fn () => def_qualify (mk_prefix_binding (mk_witN i))) (1 upto nwits);
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1209
      in bs ~~ wit_rhss end;
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1210
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1211
    val ((((bnf_rel_term, raw_rel_def), (bnf_pred_term, raw_pred_def)),
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1212
        (bnf_wit_terms, raw_wit_defs)), (lthy, lthy_old)) =
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1213
      lthy
72450
24bd1316eaae consolidated names and operations
haftmann
parents: 72154
diff changeset
  1214
      |> (snd o Local_Theory.begin_nested)
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1215
      |> maybe_define (is_some rel_rhs_opt) rel_bind_def
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1216
      ||>> maybe_define (is_some pred_rhs_opt) pred_bind_def
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1217
      ||>> apfst split_list o fold_map (maybe_define (not (null wit_rhss))) wit_binds_defs
72450
24bd1316eaae consolidated names and operations
haftmann
parents: 72154
diff changeset
  1218
      ||> `Local_Theory.end_nested;
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1219
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1220
    val phi = Proof_Context.export_morphism lthy_old lthy;
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1221
    val bnf_rel_def = Morphism.thm phi raw_rel_def;
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1222
    val bnf_rel = Morphism.term phi bnf_rel_term;
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1223
    fun mk_bnf_rel Ds As' Bs' =
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1224
      normalize_rel lthy (map2 mk_pred2T As' Bs') (mk_bnf_T Ds As' Calpha) (mk_bnf_T Ds Bs' Calpha)
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1225
        bnf_rel;
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1226
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1227
    val bnf_pred_def = Morphism.thm phi raw_pred_def;
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1228
    val bnf_pred = Morphism.term phi bnf_pred_term;
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1229
    fun mk_bnf_pred Ds As' =
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1230
      normalize_pred lthy (map mk_pred1T As') (mk_bnf_T Ds As' Calpha) bnf_pred;
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1231
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1232
    val bnf_wit_defs = map (Morphism.thm phi) raw_wit_defs;
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1233
    val bnf_wits =
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1234
      map (normalize_wit Calpha_params Calpha alphas o Morphism.term phi) bnf_wit_terms;
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1235
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1236
    fun mk_rel_spec Ds' As' Bs' =
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1237
      Term.subst_atomic_types ((Ds ~~ Ds') @ (As ~~ As') @ (Bs ~~ Bs')) rel_spec;
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1238
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1239
    fun mk_pred_spec Ds' As' =
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1240
      Term.subst_atomic_types ((Ds ~~ Ds') @ (As ~~ As')) pred_spec;
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1241
  in
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1242
    (((alphas, betas, deads, Calpha),
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1243
     (bnf_map, bnf_sets, bnf_bd, bnf_wits, bnf_rel, bnf_pred),
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1244
     (bnf_map_def, bnf_set_defs, bnf_bd_def, bnf_wit_defs, bnf_rel_def, bnf_pred_def),
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1245
     (mk_bnf_map, mk_bnf_t, mk_bnf_T, mk_bnf_rel, mk_bnf_pred, mk_rel_spec, mk_pred_spec)), lthy)
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1246
  end;
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1247
56016
8875cdcfc85b unfold intermediate definitions after sealing the bnf
traytel
parents: 55945
diff changeset
  1248
fun prepare_def const_policy mk_fact_policy internal qualify prep_typ prep_term Ds_opt map_b rel_b
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1249
  pred_b set_bs (((((((raw_bnf_b, raw_bnf_T), raw_map), raw_sets), raw_bd), raw_wits), raw_rel_opt),
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1250
    raw_pred_opt) no_defs_lthy =
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1251
  let
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1252
    val fact_policy = mk_fact_policy no_defs_lthy;
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1253
    val bnf_b = qualify raw_bnf_b;
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1254
    val live = length raw_sets;
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1255
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1256
    val T_rhs = prep_typ no_defs_lthy raw_bnf_T;
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1257
    val map_rhs = prep_term no_defs_lthy raw_map;
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1258
    val set_rhss = map (prep_term no_defs_lthy) raw_sets;
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1259
    val bd_rhs = prep_term no_defs_lthy raw_bd;
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1260
    val wit_rhss = map (prep_term no_defs_lthy) raw_wits;
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1261
    val rel_rhs_opt = Option.map (prep_term no_defs_lthy) raw_rel_opt;
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1262
    val pred_rhs_opt = Option.map (prep_term no_defs_lthy) raw_pred_opt;
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1263
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1264
    fun err T =
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1265
      error ("Trying to register the type " ^ quote (Syntax.string_of_typ no_defs_lthy T) ^
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1266
        " as unnamed BNF");
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1267
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1268
    val (bnf_b, key) =
63009
wenzelm
parents: 62482
diff changeset
  1269
      if Binding.is_empty bnf_b then
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1270
        (case T_rhs of
66198
4a5589dd8e1a more error checking
blanchet
parents: 64627
diff changeset
  1271
          Type (C, Ts) =>
4a5589dd8e1a more error checking
blanchet
parents: 64627
diff changeset
  1272
          if forall (can dest_TFree) Ts andalso not (has_duplicates (op =) Ts) then
4a5589dd8e1a more error checking
blanchet
parents: 64627
diff changeset
  1273
            (Binding.qualified_name C, C)
4a5589dd8e1a more error checking
blanchet
parents: 64627
diff changeset
  1274
          else
4a5589dd8e1a more error checking
blanchet
parents: 64627
diff changeset
  1275
            err T_rhs
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1276
        | T => err T)
66198
4a5589dd8e1a more error checking
blanchet
parents: 64627
diff changeset
  1277
      else
4a5589dd8e1a more error checking
blanchet
parents: 64627
diff changeset
  1278
        (bnf_b, Local_Theory.full_name no_defs_lthy bnf_b);
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1279
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1280
    val (((alphas, betas, deads, Calpha),
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1281
     (bnf_map, bnf_sets, bnf_bd, bnf_wits, bnf_rel, bnf_pred),
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1282
     (bnf_map_def, bnf_set_defs, bnf_bd_def, bnf_wit_defs, bnf_rel_def, bnf_pred_def),
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1283
     (mk_bnf_map_Ds, mk_bnf_t_Ds, mk_bnf_T_Ds, _, _, mk_rel_spec, mk_pred_spec)), lthy) =
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1284
       define_bnf_consts const_policy fact_policy internal Ds_opt map_b rel_b pred_b set_bs
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1285
         (((((((bnf_b, T_rhs), map_rhs), set_rhss), bd_rhs), wit_rhss), rel_rhs_opt), pred_rhs_opt)
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1286
         no_defs_lthy;
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1287
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1288
    val dead = length deads;
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1289
60154
7478de1f5b59 allow sorts on dead variables in BNFs
blanchet
parents: 59936
diff changeset
  1290
    val (((((((As', Bs'), Cs), unsorted_Ds), Es), B1Ts), B2Ts), (Ts, T)) = lthy
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1291
      |> mk_TFrees live
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1292
      ||>> mk_TFrees live
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1293
      ||>> mk_TFrees live
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1294
      ||>> mk_TFrees dead
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1295
      ||>> mk_TFrees live
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1296
      ||>> mk_TFrees live
58104
c5316f843f72 generate 'rel_transfer' for BNFs
desharna
parents: 58102
diff changeset
  1297
      ||>> mk_TFrees live
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1298
      ||> fst o mk_TFrees 1
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1299
      ||> the_single
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1300
      ||> `(replicate live);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1301
60154
7478de1f5b59 allow sorts on dead variables in BNFs
blanchet
parents: 59936
diff changeset
  1302
    val Ds = map2 (resort_tfree_or_tvar o Type.sort_of_atyp) deads unsorted_Ds;
7478de1f5b59 allow sorts on dead variables in BNFs
blanchet
parents: 59936
diff changeset
  1303
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1304
    val mk_bnf_map = mk_bnf_map_Ds Ds;
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1305
    val mk_bnf_t = mk_bnf_t_Ds Ds;
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1306
    val mk_bnf_T = mk_bnf_T_Ds Ds;
49453
ff0e540d8758 add rel as first-class citizen of BNF
blanchet
parents: 49452
diff changeset
  1307
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1308
    val pred1PTs = map mk_pred1T As';
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1309
    val pred1QTs = map mk_pred1T Bs';
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
  1310
    val pred2RTs = map2 mk_pred2T As' Bs';
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
  1311
    val pred2RTsAsCs = map2 mk_pred2T As' Cs;
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
  1312
    val pred2RTsBsCs = map2 mk_pred2T Bs' Cs;
58104
c5316f843f72 generate 'rel_transfer' for BNFs
desharna
parents: 58102
diff changeset
  1313
    val pred2RTsBsEs = map2 mk_pred2T Bs' Es;
57932
c29659f77f8d generate 'rel_map' theorem for BNFs
desharna
parents: 57668
diff changeset
  1314
    val pred2RTsCsBs = map2 mk_pred2T Cs Bs';
58104
c5316f843f72 generate 'rel_transfer' for BNFs
desharna
parents: 58102
diff changeset
  1315
    val pred2RTsCsEs = map2 mk_pred2T Cs Es;
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
  1316
    val pred2RT's = map2 mk_pred2T Bs' As';
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
  1317
    val self_pred2RTs = map2 mk_pred2T As' As';
52719
480a3479fa47 transfer rule for map (not yet registered as a transfer rule)
traytel
parents: 52635
diff changeset
  1318
    val transfer_domRTs = map2 mk_pred2T As' B1Ts;
480a3479fa47 transfer rule for map (not yet registered as a transfer rule)
traytel
parents: 52635
diff changeset
  1319
    val transfer_ranRTs = map2 mk_pred2T Bs' B2Ts;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1320
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1321
    val CA' = mk_bnf_T As' Calpha;
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1322
    val CB' = mk_bnf_T Bs' Calpha;
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1323
    val CC' = mk_bnf_T Cs Calpha;
58104
c5316f843f72 generate 'rel_transfer' for BNFs
desharna
parents: 58102
diff changeset
  1324
    val CE' = mk_bnf_T Es Calpha;
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1325
    val CB1 = mk_bnf_T B1Ts Calpha;
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1326
    val CB2 = mk_bnf_T B2Ts Calpha;
49453
ff0e540d8758 add rel as first-class citizen of BNF
blanchet
parents: 49452
diff changeset
  1327
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1328
    val bnf_map_AsAs = mk_bnf_map As' As';
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1329
    val bnf_map_AsBs = mk_bnf_map As' Bs';
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1330
    val bnf_map_AsCs = mk_bnf_map As' Cs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1331
    val bnf_map_BsCs = mk_bnf_map Bs' Cs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1332
    val bnf_sets_As = map (mk_bnf_t As') bnf_sets;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1333
    val bnf_sets_Bs = map (mk_bnf_t Bs') bnf_sets;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1334
    val bnf_bd_As = mk_bnf_t As' bnf_bd;
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1335
    fun mk_bnf_rel RTs CA CB = normalize_rel lthy RTs CA CB bnf_rel;
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1336
    fun mk_bnf_pred PTs CA = normalize_pred lthy PTs CA bnf_pred;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1337
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1338
    val ((((((((((((((((((((((((((fs, fs'), gs), hs), is), x), x'), y), y'), zs), zs'), ys), As),
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1339
      As_copy), bs), (Ps, Ps')), Ps_copy), Qs), Rs), Rs_copy), Ss), S_AsCs), S_CsBs), S_BsEs),
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1340
      transfer_domRs), transfer_ranRs), _) = lthy
52923
traytel
parents: 52844
diff changeset
  1341
      |> mk_Frees "f" (map2 (curry op -->) As' Bs')
57970
eaa986cd285a generate 'inj_map_strong' for BNFs
desharna
parents: 57969
diff changeset
  1342
      ||>> mk_Frees "f" (map2 (curry op -->) As' Bs')
52923
traytel
parents: 52844
diff changeset
  1343
      ||>> mk_Frees "g" (map2 (curry op -->) Bs' Cs)
traytel
parents: 52844
diff changeset
  1344
      ||>> mk_Frees "h" (map2 (curry op -->) As' Ts)
57932
c29659f77f8d generate 'rel_map' theorem for BNFs
desharna
parents: 57668
diff changeset
  1345
      ||>> mk_Frees "i" (map2 (curry op -->) As' Cs)
51894
traytel
parents: 51893
diff changeset
  1346
      ||>> yield_singleton (mk_Frees "x") CA'
57970
eaa986cd285a generate 'inj_map_strong' for BNFs
desharna
parents: 57969
diff changeset
  1347
      ||>> yield_singleton (mk_Frees "x") CA'
51894
traytel
parents: 51893
diff changeset
  1348
      ||>> yield_singleton (mk_Frees "y") CB'
61242
1f92b0ac9c96 congruence rules for the relator
traytel
parents: 61241
diff changeset
  1349
      ||>> yield_singleton (mk_Frees "y") CB'
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1350
      ||>> mk_Frees "z" As'
57970
eaa986cd285a generate 'inj_map_strong' for BNFs
desharna
parents: 57969
diff changeset
  1351
      ||>> mk_Frees "z" As'
51916
eac9e9a45bf5 stronger monotonicity property for relators
traytel
parents: 51915
diff changeset
  1352
      ||>> mk_Frees "y" Bs'
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1353
      ||>> mk_Frees "A" (map HOLogic.mk_setT As')
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1354
      ||>> mk_Frees "A" (map HOLogic.mk_setT As')
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1355
      ||>> mk_Frees "b" As'
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1356
      ||>> mk_Frees' "P" pred1PTs
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1357
      ||>> mk_Frees "P" pred1PTs
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1358
      ||>> mk_Frees "Q" pred1QTs
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1359
      ||>> mk_Frees "R" pred2RTs
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
  1360
      ||>> mk_Frees "R" pred2RTs
52719
480a3479fa47 transfer rule for map (not yet registered as a transfer rule)
traytel
parents: 52635
diff changeset
  1361
      ||>> mk_Frees "S" pred2RTsBsCs
57932
c29659f77f8d generate 'rel_map' theorem for BNFs
desharna
parents: 57668
diff changeset
  1362
      ||>> mk_Frees "S" pred2RTsAsCs
c29659f77f8d generate 'rel_map' theorem for BNFs
desharna
parents: 57668
diff changeset
  1363
      ||>> mk_Frees "S" pred2RTsCsBs
58104
c5316f843f72 generate 'rel_transfer' for BNFs
desharna
parents: 58102
diff changeset
  1364
      ||>> mk_Frees "S" pred2RTsBsEs
52719
480a3479fa47 transfer rule for map (not yet registered as a transfer rule)
traytel
parents: 52635
diff changeset
  1365
      ||>> mk_Frees "R" transfer_domRTs
480a3479fa47 transfer rule for map (not yet registered as a transfer rule)
traytel
parents: 52635
diff changeset
  1366
      ||>> mk_Frees "S" transfer_ranRTs;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1367
56651
fc105315822a localize new size function generation code
blanchet
parents: 56635
diff changeset
  1368
    val fs_copy = map2 (retype_const_or_free o fastype_of) fs gs;
61242
1f92b0ac9c96 congruence rules for the relator
traytel
parents: 61241
diff changeset
  1369
    val x_copy = retype_const_or_free CA' y';
1f92b0ac9c96 congruence rules for the relator
traytel
parents: 61241
diff changeset
  1370
    val y_copy = retype_const_or_free CB' x';
51762
219a3063ed29 derive "map_cong"
blanchet
parents: 51761
diff changeset
  1371
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
  1372
    val rel = mk_bnf_rel pred2RTs CA' CB';
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1373
    val pred = mk_bnf_pred pred1PTs CA';
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1374
    val pred' = mk_bnf_pred pred1QTs CB';
58104
c5316f843f72 generate 'rel_transfer' for BNFs
desharna
parents: 58102
diff changeset
  1375
    val relCsEs = mk_bnf_rel pred2RTsCsEs CC' CE';
52719
480a3479fa47 transfer rule for map (not yet registered as a transfer rule)
traytel
parents: 52635
diff changeset
  1376
    val relAsAs = mk_bnf_rel self_pred2RTs CA' CA';
54189
c0186a0d8cb3 define a trivial nonemptiness witness if none is provided
traytel
parents: 54158
diff changeset
  1377
    val bnf_wit_As = map (apsnd (mk_bnf_t As')) bnf_wits;
c0186a0d8cb3 define a trivial nonemptiness witness if none is provided
traytel
parents: 54158
diff changeset
  1378
53270
c8628119d18e renamed BNF axiom
blanchet
parents: 53265
diff changeset
  1379
    val map_id0_goal =
51762
219a3063ed29 derive "map_cong"
blanchet
parents: 51761
diff changeset
  1380
      let val bnf_map_app_id = Term.list_comb (bnf_map_AsAs, map HOLogic.id_const As') in
219a3063ed29 derive "map_cong"
blanchet
parents: 51761
diff changeset
  1381
        mk_Trueprop_eq (bnf_map_app_id, HOLogic.id_const CA')
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1382
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1383
53287
271b34513bfb renamed BNF axiom
blanchet
parents: 53285
diff changeset
  1384
    val map_comp0_goal =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1385
      let
49018
b2884253b421 renamed ML function for consistency
blanchet
parents: 49016
diff changeset
  1386
        val bnf_map_app_comp = Term.list_comb (bnf_map_AsCs, map2 (curry HOLogic.mk_comp) gs fs);
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1387
        val comp_bnf_map_app = HOLogic.mk_comp
51762
219a3063ed29 derive "map_cong"
blanchet
parents: 51761
diff changeset
  1388
          (Term.list_comb (bnf_map_BsCs, gs), Term.list_comb (bnf_map_AsBs, fs));
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1389
      in
49123
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents: 49111
diff changeset
  1390
        fold_rev Logic.all (fs @ gs) (mk_Trueprop_eq (bnf_map_app_comp, comp_bnf_map_app))
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1391
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1392
57981
81baacfd56e8 generate 'map_cong_simp' for BNFs
desharna
parents: 57970
diff changeset
  1393
    fun mk_map_cong_prem mk_implies x z set f f_copy =
58108
1c8541513acb tuned whitespace
blanchet
parents: 58106
diff changeset
  1394
      Logic.all z (mk_implies (mk_Trueprop_mem (z, set $ x), mk_Trueprop_eq (f $ z, f_copy $ z)));
51762
219a3063ed29 derive "map_cong"
blanchet
parents: 51761
diff changeset
  1395
51761
4c9f08836d87 renamed "map_cong" axiom to "map_cong0" in preparation for real "map_cong"
blanchet
parents: 51758
diff changeset
  1396
    val map_cong0_goal =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1397
      let
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58432
diff changeset
  1398
        val prems = @{map 4} (mk_map_cong_prem Logic.mk_implies x) zs bnf_sets_As fs fs_copy;
51762
219a3063ed29 derive "map_cong"
blanchet
parents: 51761
diff changeset
  1399
        val eq = mk_Trueprop_eq (Term.list_comb (bnf_map_AsBs, fs) $ x,
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1400
          Term.list_comb (bnf_map_AsBs, fs_copy) $ x);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1401
      in
51762
219a3063ed29 derive "map_cong"
blanchet
parents: 51761
diff changeset
  1402
        fold_rev Logic.all (x :: fs @ fs_copy) (Logic.list_implies (prems, eq))
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1403
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1404
53289
5e0623448bdb renamed BNF axiom
blanchet
parents: 53288
diff changeset
  1405
    val set_map0s_goal =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1406
      let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1407
        fun mk_goal setA setB f =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1408
          let
58108
1c8541513acb tuned whitespace
blanchet
parents: 58106
diff changeset
  1409
            val set_comp_map = HOLogic.mk_comp (setB, Term.list_comb (bnf_map_AsBs, fs));
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1410
            val image_comp_set = HOLogic.mk_comp (mk_image f, setA);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1411
          in
49123
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents: 49111
diff changeset
  1412
            fold_rev Logic.all fs (mk_Trueprop_eq (set_comp_map, image_comp_set))
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1413
          end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1414
      in
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58432
diff changeset
  1415
        @{map 3} mk_goal bnf_sets_As bnf_sets_Bs fs
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1416
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1417
49458
blanchet
parents: 49456
diff changeset
  1418
    val card_order_bd_goal = HOLogic.mk_Trueprop (mk_card_order bnf_bd_As);
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1419
49458
blanchet
parents: 49456
diff changeset
  1420
    val cinfinite_bd_goal = HOLogic.mk_Trueprop (mk_cinfinite bnf_bd_As);
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1421
75624
22d1c5f2b9f4 strict bounds for BNFs (by Jan van Brügge)
traytel
parents: 75276
diff changeset
  1422
    val regularCard_bd_goal = HOLogic.mk_Trueprop (mk_regularCard bnf_bd_As);
22d1c5f2b9f4 strict bounds for BNFs (by Jan van Brügge)
traytel
parents: 75276
diff changeset
  1423
49458
blanchet
parents: 49456
diff changeset
  1424
    val set_bds_goal =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1425
      let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1426
        fun mk_goal set =
75624
22d1c5f2b9f4 strict bounds for BNFs (by Jan van Brügge)
traytel
parents: 75276
diff changeset
  1427
          Logic.all x (HOLogic.mk_Trueprop (mk_ordLess (mk_card_of (set $ x)) bnf_bd_As));
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1428
      in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1429
        map mk_goal bnf_sets_As
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1430
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1431
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1432
    val relAsCs = mk_bnf_rel pred2RTsAsCs CA' CC';
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1433
    val relBsCs = mk_bnf_rel pred2RTsBsCs CB' CC';
57932
c29659f77f8d generate 'rel_map' theorem for BNFs
desharna
parents: 57668
diff changeset
  1434
    val relCsBs = mk_bnf_rel pred2RTsCsBs CC' CB';
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1435
    val rel_OO_lhs = Term.list_comb (relAsCs, map2 (curry mk_rel_compp) Rs Ss);
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1436
    val rel_OO_rhs = mk_rel_compp (Term.list_comb (rel, Rs), Term.list_comb (relBsCs, Ss));
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1437
    val le_rel_OO_goal =
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1438
      fold_rev Logic.all (Rs @ Ss) (HOLogic.mk_Trueprop (mk_leq rel_OO_rhs rel_OO_lhs));
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1439
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1440
    val rel_OO_Grp_goal = fold_rev Logic.all Rs (mk_Trueprop_eq (Term.list_comb (rel, Rs),
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1441
      Term.list_comb (mk_rel_spec Ds As' Bs', Rs)));
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1442
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1443
    val pred_set_goal = fold_rev Logic.all Ps (mk_Trueprop_eq (Term.list_comb (pred, Ps),
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1444
      Term.list_comb (mk_pred_spec Ds As', Ps)));
49453
ff0e540d8758 add rel as first-class citizen of BNF
blanchet
parents: 49452
diff changeset
  1445
53289
5e0623448bdb renamed BNF axiom
blanchet
parents: 53288
diff changeset
  1446
    val goals = zip_axioms map_id0_goal map_comp0_goal map_cong0_goal set_map0s_goal
75624
22d1c5f2b9f4 strict bounds for BNFs (by Jan van Brügge)
traytel
parents: 75276
diff changeset
  1447
      card_order_bd_goal cinfinite_bd_goal regularCard_bd_goal set_bds_goal le_rel_OO_goal rel_OO_Grp_goal pred_set_goal;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1448
60918
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents: 60801
diff changeset
  1449
    val mk_wit_goals = mk_wit_goals bs zs bnf_sets_As;
55197
5a54ed681ba2 less hermetic tactics
traytel
parents: 55061
diff changeset
  1450
    fun triv_wit_tac ctxt = mk_trivial_wit_tac ctxt bnf_wit_defs;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1451
54189
c0186a0d8cb3 define a trivial nonemptiness witness if none is provided
traytel
parents: 54158
diff changeset
  1452
    val wit_goalss =
54921
862c36b6e57c avoid schematic variable in goal, which sometimes gets instantiated by tactic
blanchet
parents: 54841
diff changeset
  1453
      (if null raw_wits then SOME triv_wit_tac else NONE, map mk_wit_goals bnf_wit_As);
54189
c0186a0d8cb3 define a trivial nonemptiness witness if none is provided
traytel
parents: 54158
diff changeset
  1454
c0186a0d8cb3 define a trivial nonemptiness witness if none is provided
traytel
parents: 54158
diff changeset
  1455
    fun after_qed mk_wit_thms thms lthy =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1456
      let
54189
c0186a0d8cb3 define a trivial nonemptiness witness if none is provided
traytel
parents: 54158
diff changeset
  1457
        val (axioms, nontriv_wit_thms) = apfst (mk_axioms live) (chop (length goals) thms);
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1458
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49103
diff changeset
  1459
        val bd_Card_order = #bd_card_order axioms RS @{thm conjunct2[OF card_order_on_Card_order]};
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1460
        val bd_Cinfinite = @{thm conjI} OF [#bd_cinfinite axioms, bd_Card_order];
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1461
        val bd_Cnotzero = bd_Cinfinite RS @{thm Cinfinite_Cnotzero};
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1462
51766
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51765
diff changeset
  1463
        fun mk_collect_set_map () =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1464
          let
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1465
            val defT = mk_bnf_T Ts Calpha --> HOLogic.mk_setT T;
58108
1c8541513acb tuned whitespace
blanchet
parents: 58106
diff changeset
  1466
            val collect_map = HOLogic.mk_comp (mk_collect (map (mk_bnf_t Ts) bnf_sets) defT,
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1467
              Term.list_comb (mk_bnf_map As' Ts, hs));
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1468
            val image_collect = mk_collect
58108
1c8541513acb tuned whitespace
blanchet
parents: 58106
diff changeset
  1469
              (map2 (fn h => fn set => HOLogic.mk_comp (mk_image h, set)) hs bnf_sets_As) defT;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1470
            (*collect {set1 ... setm} o map f1 ... fm = collect {f1` o set1 ... fm` o setm}*)
49123
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents: 49111
diff changeset
  1471
            val goal = fold_rev Logic.all hs (mk_Trueprop_eq (collect_map, image_collect));
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1472
          in
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1473
            Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} =>
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 60154
diff changeset
  1474
              mk_collect_set_map_tac ctxt (#set_map0 axioms))
70494
41108e3e9ca5 formal position for PThm nodes;
wenzelm
parents: 69593
diff changeset
  1475
            |> Thm.close_derivation \<^here>
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1476
          end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1477
51766
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51765
diff changeset
  1478
        val collect_set_map = Lazy.lazy mk_collect_set_map;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1479
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1480
        fun mk_in_mono () =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1481
          let
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
  1482
            val prems_mono = map2 (HOLogic.mk_Trueprop oo mk_leq) As As_copy;
49458
blanchet
parents: 49456
diff changeset
  1483
            val in_mono_goal =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1484
              fold_rev Logic.all (As @ As_copy)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1485
                (Logic.list_implies (prems_mono, HOLogic.mk_Trueprop
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
  1486
                  (mk_leq (mk_in As bnf_sets_As CA') (mk_in As_copy bnf_sets_As CA'))));
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1487
          in
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 60154
diff changeset
  1488
            Goal.prove_sorry lthy [] [] in_mono_goal (fn {context = ctxt, prems = _} =>
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 60154
diff changeset
  1489
              mk_in_mono_tac ctxt live)
70494
41108e3e9ca5 formal position for PThm nodes;
wenzelm
parents: 69593
diff changeset
  1490
            |> Thm.close_derivation \<^here>
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1491
          end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1492
49538
c90818b63599 simplified fact policies
blanchet
parents: 49537
diff changeset
  1493
        val in_mono = Lazy.lazy mk_in_mono;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1494
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1495
        fun mk_in_cong () =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1496
          let
51762
219a3063ed29 derive "map_cong"
blanchet
parents: 51761
diff changeset
  1497
            val prems_cong = map2 (curry mk_Trueprop_eq) As As_copy;
49458
blanchet
parents: 49456
diff changeset
  1498
            val in_cong_goal =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1499
              fold_rev Logic.all (As @ As_copy)
51762
219a3063ed29 derive "map_cong"
blanchet
parents: 51761
diff changeset
  1500
                (Logic.list_implies (prems_cong,
219a3063ed29 derive "map_cong"
blanchet
parents: 51761
diff changeset
  1501
                  mk_Trueprop_eq (mk_in As bnf_sets_As CA', mk_in As_copy bnf_sets_As CA')));
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1502
          in
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 51790
diff changeset
  1503
            Goal.prove_sorry lthy [] [] in_cong_goal
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 60154
diff changeset
  1504
              (fn {context = ctxt, prems = _} => (TRY o hyp_subst_tac ctxt THEN' rtac ctxt refl) 1)
70494
41108e3e9ca5 formal position for PThm nodes;
wenzelm
parents: 69593
diff changeset
  1505
            |> Thm.close_derivation \<^here>
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1506
          end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1507
49538
c90818b63599 simplified fact policies
blanchet
parents: 49537
diff changeset
  1508
        val in_cong = Lazy.lazy mk_in_cong;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1509
53285
f09645642984 renamed BNF fact
blanchet
parents: 53270
diff changeset
  1510
        val map_id = Lazy.lazy (fn () => mk_map_id (#map_id0 axioms));
57399
cfc19f0a6261 compile
blanchet
parents: 57303
diff changeset
  1511
        val map_ident0 = Lazy.lazy (fn () => mk_map_ident lthy (#map_id0 axioms));
56903
69b6369a98fa generate 'map_ident' theorem for BNFs
desharna
parents: 56766
diff changeset
  1512
        val map_ident = Lazy.lazy (fn () => mk_map_ident lthy (Lazy.force map_id));
75276
686a6d7d0991 generated lemma map_ident_strong for BNFs
desharna
parents: 74664
diff changeset
  1513
        val map_ident_strong = Lazy.lazy (fn () =>
686a6d7d0991 generated lemma map_ident_strong for BNFs
desharna
parents: 74664
diff changeset
  1514
          mk_map_ident_strong lthy (#map_cong0 axioms) (Lazy.force map_id));
53288
050d0bc9afa5 renamed BNF fact
blanchet
parents: 53287
diff changeset
  1515
        val map_comp = Lazy.lazy (fn () => mk_map_comp (#map_comp0 axioms));
51762
219a3063ed29 derive "map_cong"
blanchet
parents: 51761
diff changeset
  1516
57981
81baacfd56e8 generate 'map_cong_simp' for BNFs
desharna
parents: 57970
diff changeset
  1517
        fun mk_map_cong mk_implies () =
51762
219a3063ed29 derive "map_cong"
blanchet
parents: 51761
diff changeset
  1518
          let
219a3063ed29 derive "map_cong"
blanchet
parents: 51761
diff changeset
  1519
            val prem0 = mk_Trueprop_eq (x, x_copy);
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58432
diff changeset
  1520
            val prems = @{map 4} (mk_map_cong_prem mk_implies x_copy) zs bnf_sets_As fs fs_copy;
51762
219a3063ed29 derive "map_cong"
blanchet
parents: 51761
diff changeset
  1521
            val eq = mk_Trueprop_eq (Term.list_comb (bnf_map_AsBs, fs) $ x,
219a3063ed29 derive "map_cong"
blanchet
parents: 51761
diff changeset
  1522
              Term.list_comb (bnf_map_AsBs, fs_copy) $ x_copy);
219a3063ed29 derive "map_cong"
blanchet
parents: 51761
diff changeset
  1523
            val goal = fold_rev Logic.all (x :: x_copy :: fs @ fs_copy)
219a3063ed29 derive "map_cong"
blanchet
parents: 51761
diff changeset
  1524
              (Logic.list_implies (prem0 :: prems, eq));
219a3063ed29 derive "map_cong"
blanchet
parents: 51761
diff changeset
  1525
          in
61271
0478ba10152a more canonical context threading
traytel
parents: 61242
diff changeset
  1526
            Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} =>
0478ba10152a more canonical context threading
traytel
parents: 61242
diff changeset
  1527
              unfold_thms_tac ctxt @{thms simp_implies_def} THEN
0478ba10152a more canonical context threading
traytel
parents: 61242
diff changeset
  1528
              mk_map_cong_tac ctxt (#map_cong0 axioms))
70494
41108e3e9ca5 formal position for PThm nodes;
wenzelm
parents: 69593
diff changeset
  1529
            |> Thm.close_derivation \<^here>
51762
219a3063ed29 derive "map_cong"
blanchet
parents: 51761
diff changeset
  1530
          end;
219a3063ed29 derive "map_cong"
blanchet
parents: 51761
diff changeset
  1531
57981
81baacfd56e8 generate 'map_cong_simp' for BNFs
desharna
parents: 57970
diff changeset
  1532
        val map_cong = Lazy.lazy (mk_map_cong Logic.mk_implies);
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67222
diff changeset
  1533
        val map_cong_simp = Lazy.lazy (mk_map_cong (fn (a, b) => \<^term>\<open>simp_implies\<close> $ a $ b));
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1534
56635
b07c8ad23010 added 'inj_map' as auxiliary BNF theorem
blanchet
parents: 56523
diff changeset
  1535
        fun mk_inj_map () =
b07c8ad23010 added 'inj_map' as auxiliary BNF theorem
blanchet
parents: 56523
diff changeset
  1536
          let
b07c8ad23010 added 'inj_map' as auxiliary BNF theorem
blanchet
parents: 56523
diff changeset
  1537
            val prems = map (HOLogic.mk_Trueprop o mk_inj) fs;
b07c8ad23010 added 'inj_map' as auxiliary BNF theorem
blanchet
parents: 56523
diff changeset
  1538
            val concl = HOLogic.mk_Trueprop (mk_inj (Term.list_comb (bnf_map_AsBs, fs)));
b07c8ad23010 added 'inj_map' as auxiliary BNF theorem
blanchet
parents: 56523
diff changeset
  1539
            val goal = fold_rev Logic.all fs (Logic.list_implies (prems, concl));
b07c8ad23010 added 'inj_map' as auxiliary BNF theorem
blanchet
parents: 56523
diff changeset
  1540
          in
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 60154
diff changeset
  1541
            Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} =>
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 60154
diff changeset
  1542
              mk_inj_map_tac ctxt live (Lazy.force map_id) (Lazy.force map_comp) (#map_cong0 axioms)
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 60154
diff changeset
  1543
                (Lazy.force map_cong))
70494
41108e3e9ca5 formal position for PThm nodes;
wenzelm
parents: 69593
diff changeset
  1544
            |> Thm.close_derivation \<^here>
56635
b07c8ad23010 added 'inj_map' as auxiliary BNF theorem
blanchet
parents: 56523
diff changeset
  1545
          end;
b07c8ad23010 added 'inj_map' as auxiliary BNF theorem
blanchet
parents: 56523
diff changeset
  1546
b07c8ad23010 added 'inj_map' as auxiliary BNF theorem
blanchet
parents: 56523
diff changeset
  1547
        val inj_map = Lazy.lazy mk_inj_map;
b07c8ad23010 added 'inj_map' as auxiliary BNF theorem
blanchet
parents: 56523
diff changeset
  1548
53290
b6c3be868217 renamed BNF fact
blanchet
parents: 53289
diff changeset
  1549
        val set_map = map (fn thm => Lazy.lazy (fn () => mk_set_map thm)) (#set_map0 axioms);
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1550
54189
c0186a0d8cb3 define a trivial nonemptiness witness if none is provided
traytel
parents: 54158
diff changeset
  1551
        val wit_thms =
c0186a0d8cb3 define a trivial nonemptiness witness if none is provided
traytel
parents: 54158
diff changeset
  1552
          if null nontriv_wit_thms then mk_wit_thms (map Lazy.force set_map) else nontriv_wit_thms;
c0186a0d8cb3 define a trivial nonemptiness witness if none is provided
traytel
parents: 54158
diff changeset
  1553
52635
4f84b730c489 got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
traytel
parents: 51930
diff changeset
  1554
        fun mk_in_bd () =
4f84b730c489 got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
traytel
parents: 51930
diff changeset
  1555
          let
52813
963297a24206 more robust tactic
traytel
parents: 52731
diff changeset
  1556
            val bdT = fst (dest_relT (fastype_of bnf_bd_As));
963297a24206 more robust tactic
traytel
parents: 52731
diff changeset
  1557
            val bdTs = replicate live bdT;
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1558
            val bd_bnfT = mk_bnf_T bdTs Calpha;
52813
963297a24206 more robust tactic
traytel
parents: 52731
diff changeset
  1559
            val surj_imp_ordLeq_inst = (if live = 0 then TrueI else
963297a24206 more robust tactic
traytel
parents: 52731
diff changeset
  1560
              let
963297a24206 more robust tactic
traytel
parents: 52731
diff changeset
  1561
                val ranTs = map (fn AT => mk_sumT (AT, HOLogic.unitT)) As';
963297a24206 more robust tactic
traytel
parents: 52731
diff changeset
  1562
                val funTs = map (fn T => bdT --> T) ranTs;
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1563
                val ran_bnfT = mk_bnf_T ranTs Calpha;
52813
963297a24206 more robust tactic
traytel
parents: 52731
diff changeset
  1564
                val (revTs, Ts) = `rev (bd_bnfT :: funTs);
61760
1647bb489522 tuned whitespace
blanchet
parents: 61424
diff changeset
  1565
                val cTs = map (SOME o Thm.ctyp_of lthy) [ran_bnfT,
1647bb489522 tuned whitespace
blanchet
parents: 61424
diff changeset
  1566
                  Library.foldr1 HOLogic.mk_prodT Ts];
1647bb489522 tuned whitespace
blanchet
parents: 61424
diff changeset
  1567
                val tinst = fold (fn T => fn t =>
1647bb489522 tuned whitespace
blanchet
parents: 61424
diff changeset
  1568
                  HOLogic.mk_case_prod (Term.absdummy T t)) (tl revTs)
1647bb489522 tuned whitespace
blanchet
parents: 61424
diff changeset
  1569
                    (Term.absdummy (hd revTs) (Term.list_comb (mk_bnf_map bdTs ranTs,
1647bb489522 tuned whitespace
blanchet
parents: 61424
diff changeset
  1570
                      map Bound (live - 1 downto 0)) $ Bound live));
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59580
diff changeset
  1571
                val cts = [NONE, SOME (Thm.cterm_of lthy tinst)];
52813
963297a24206 more robust tactic
traytel
parents: 52731
diff changeset
  1572
              in
60801
7664e0916eec tuned signature;
wenzelm
parents: 60784
diff changeset
  1573
                Thm.instantiate' cTs cts @{thm surj_imp_ordLeq}
52813
963297a24206 more robust tactic
traytel
parents: 52731
diff changeset
  1574
              end);
52635
4f84b730c489 got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
traytel
parents: 51930
diff changeset
  1575
            val bd = mk_cexp
4f84b730c489 got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
traytel
parents: 51930
diff changeset
  1576
              (if live = 0 then ctwo
4f84b730c489 got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
traytel
parents: 51930
diff changeset
  1577
                else mk_csum (Library.foldr1 (uncurry mk_csum) (map mk_card_of As)) ctwo)
52813
963297a24206 more robust tactic
traytel
parents: 52731
diff changeset
  1578
              (mk_csum bnf_bd_As (mk_card_of (HOLogic.mk_UNIV bd_bnfT)));
52635
4f84b730c489 got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
traytel
parents: 51930
diff changeset
  1579
            val in_bd_goal =
4f84b730c489 got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
traytel
parents: 51930
diff changeset
  1580
              fold_rev Logic.all As
4f84b730c489 got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
traytel
parents: 51930
diff changeset
  1581
                (HOLogic.mk_Trueprop (mk_ordLeq (mk_card_of (mk_in As bnf_sets_As CA')) bd));
75624
22d1c5f2b9f4 strict bounds for BNFs (by Jan van Brügge)
traytel
parents: 75276
diff changeset
  1582
            val weak_set_bds = map (fn thm => @{thm ordLess_imp_ordLeq} OF [thm]) (#set_bd axioms);
52635
4f84b730c489 got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
traytel
parents: 51930
diff changeset
  1583
          in
4f84b730c489 got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
traytel
parents: 51930
diff changeset
  1584
            Goal.prove_sorry lthy [] [] in_bd_goal
55197
5a54ed681ba2 less hermetic tactics
traytel
parents: 55061
diff changeset
  1585
              (fn {context = ctxt, prems = _} => mk_in_bd_tac ctxt live surj_imp_ordLeq_inst
53288
050d0bc9afa5 renamed BNF fact
blanchet
parents: 53287
diff changeset
  1586
                (Lazy.force map_comp) (Lazy.force map_id) (#map_cong0 axioms)
75624
22d1c5f2b9f4 strict bounds for BNFs (by Jan van Brügge)
traytel
parents: 75276
diff changeset
  1587
                (map Lazy.force set_map) weak_set_bds (#bd_card_order axioms)
52635
4f84b730c489 got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
traytel
parents: 51930
diff changeset
  1588
                bd_Card_order bd_Cinfinite bd_Cnotzero)
70494
41108e3e9ca5 formal position for PThm nodes;
wenzelm
parents: 69593
diff changeset
  1589
            |> Thm.close_derivation \<^here>
52635
4f84b730c489 got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
traytel
parents: 51930
diff changeset
  1590
          end;
4f84b730c489 got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
traytel
parents: 51930
diff changeset
  1591
4f84b730c489 got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
traytel
parents: 51930
diff changeset
  1592
        val in_bd = Lazy.lazy mk_in_bd;
4f84b730c489 got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
traytel
parents: 51930
diff changeset
  1593
53561
92bcac4f9ac9 simplified derivation of in_rel
traytel
parents: 53290
diff changeset
  1594
        val rel_OO_Grp = #rel_OO_Grp axioms;
92bcac4f9ac9 simplified derivation of in_rel
traytel
parents: 53290
diff changeset
  1595
        val rel_OO_Grps = no_refl [rel_OO_Grp];
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1596
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
  1597
        fun mk_rel_Grp () =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1598
          let
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
  1599
            val lhs = Term.list_comb (rel, map2 mk_Grp As fs);
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
  1600
            val rhs = mk_Grp (mk_in As bnf_sets_As CA') (Term.list_comb (bnf_map_AsBs, fs));
49123
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents: 49111
diff changeset
  1601
            val goal = fold_rev Logic.all (As @ fs) (mk_Trueprop_eq (lhs, rhs));
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1602
          in
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 49631
diff changeset
  1603
            Goal.prove_sorry lthy [] [] goal
55197
5a54ed681ba2 less hermetic tactics
traytel
parents: 55061
diff changeset
  1604
              (fn {context = ctxt, prems = _} => mk_rel_Grp_tac ctxt rel_OO_Grps (#map_id0 axioms)
5a54ed681ba2 less hermetic tactics
traytel
parents: 55061
diff changeset
  1605
                (#map_cong0 axioms) (Lazy.force map_id) (Lazy.force map_comp)
5a54ed681ba2 less hermetic tactics
traytel
parents: 55061
diff changeset
  1606
                (map Lazy.force set_map))
70494
41108e3e9ca5 formal position for PThm nodes;
wenzelm
parents: 69593
diff changeset
  1607
            |> Thm.close_derivation \<^here>
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1608
          end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1609
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
  1610
        val rel_Grp = Lazy.lazy mk_rel_Grp;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1611
61242
1f92b0ac9c96 congruence rules for the relator
traytel
parents: 61241
diff changeset
  1612
        fun mk_rel_prems f = map2 (HOLogic.mk_Trueprop oo f) Rs Rs_copy;
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
  1613
        fun mk_rel_concl f = HOLogic.mk_Trueprop
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
  1614
          (f (Term.list_comb (rel, Rs), Term.list_comb (rel, Rs_copy)));
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1615
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
  1616
        fun mk_rel_mono () =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1617
          let
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
  1618
            val mono_prems = mk_rel_prems mk_leq;
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
  1619
            val mono_concl = mk_rel_concl (uncurry mk_leq);
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1620
          in
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 49631
diff changeset
  1621
            Goal.prove_sorry lthy [] []
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1622
              (fold_rev Logic.all (Rs @ Rs_copy) (Logic.list_implies (mono_prems, mono_concl)))
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 60154
diff changeset
  1623
              (fn {context = ctxt, prems = _} =>
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 60154
diff changeset
  1624
                mk_rel_mono_tac ctxt rel_OO_Grps (Lazy.force in_mono))
70494
41108e3e9ca5 formal position for PThm nodes;
wenzelm
parents: 69593
diff changeset
  1625
            |> Thm.close_derivation \<^here>
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1626
          end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1627
61242
1f92b0ac9c96 congruence rules for the relator
traytel
parents: 61241
diff changeset
  1628
        fun mk_rel_cong0 () =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1629
          let
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
  1630
            val cong_prems = mk_rel_prems (curry HOLogic.mk_eq);
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
  1631
            val cong_concl = mk_rel_concl HOLogic.mk_eq;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1632
          in
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 49631
diff changeset
  1633
            Goal.prove_sorry lthy [] []
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1634
              (fold_rev Logic.all (Rs @ Rs_copy) (Logic.list_implies (cong_prems, cong_concl)))
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 60154
diff changeset
  1635
              (fn {context = ctxt, prems = _} => (TRY o hyp_subst_tac ctxt THEN' rtac ctxt refl) 1)
70494
41108e3e9ca5 formal position for PThm nodes;
wenzelm
parents: 69593
diff changeset
  1636
            |> Thm.close_derivation \<^here>
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1637
          end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1638
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
  1639
        val rel_mono = Lazy.lazy mk_rel_mono;
61242
1f92b0ac9c96 congruence rules for the relator
traytel
parents: 61241
diff changeset
  1640
        val rel_cong0 = Lazy.lazy mk_rel_cong0;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1641
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
  1642
        fun mk_rel_eq () =
52719
480a3479fa47 transfer rule for map (not yet registered as a transfer rule)
traytel
parents: 52635
diff changeset
  1643
          Goal.prove_sorry lthy [] []
480a3479fa47 transfer rule for map (not yet registered as a transfer rule)
traytel
parents: 52635
diff changeset
  1644
            (mk_Trueprop_eq (Term.list_comb (relAsAs, map HOLogic.eq_const As'),
480a3479fa47 transfer rule for map (not yet registered as a transfer rule)
traytel
parents: 52635
diff changeset
  1645
              HOLogic.eq_const CA'))
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 60154
diff changeset
  1646
            (fn {context = ctxt, prems = _} =>
61242
1f92b0ac9c96 congruence rules for the relator
traytel
parents: 61241
diff changeset
  1647
              mk_rel_eq_tac ctxt live (Lazy.force rel_Grp) (Lazy.force rel_cong0) (#map_id0 axioms))
70494
41108e3e9ca5 formal position for PThm nodes;
wenzelm
parents: 69593
diff changeset
  1648
          |> Thm.close_derivation \<^here>;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1649
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
  1650
        val rel_eq = Lazy.lazy mk_rel_eq;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1651
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
  1652
        fun mk_rel_conversep () =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1653
          let
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
  1654
            val relBsAs = mk_bnf_rel pred2RT's CB' CA';
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
  1655
            val lhs = Term.list_comb (relBsAs, map mk_conversep Rs);
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
  1656
            val rhs = mk_conversep (Term.list_comb (rel, Rs));
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
  1657
            val le_goal = fold_rev Logic.all Rs (HOLogic.mk_Trueprop (mk_leq lhs rhs));
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 49631
diff changeset
  1658
            val le_thm = Goal.prove_sorry lthy [] [] le_goal
55197
5a54ed681ba2 less hermetic tactics
traytel
parents: 55061
diff changeset
  1659
              (fn {context = ctxt, prems = _} => mk_rel_conversep_le_tac ctxt rel_OO_Grps
5a54ed681ba2 less hermetic tactics
traytel
parents: 55061
diff changeset
  1660
                (Lazy.force rel_eq) (#map_cong0 axioms) (Lazy.force map_comp)
5a54ed681ba2 less hermetic tactics
traytel
parents: 55061
diff changeset
  1661
                (map Lazy.force set_map))
70494
41108e3e9ca5 formal position for PThm nodes;
wenzelm
parents: 69593
diff changeset
  1662
              |> Thm.close_derivation \<^here>
49123
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents: 49111
diff changeset
  1663
            val goal = fold_rev Logic.all Rs (mk_Trueprop_eq (lhs, rhs));
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1664
          in
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
  1665
            Goal.prove_sorry lthy [] [] goal
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 60154
diff changeset
  1666
              (fn {context = ctxt, prems = _} =>
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 60154
diff changeset
  1667
                mk_rel_conversep_tac ctxt le_thm (Lazy.force rel_mono))
70494
41108e3e9ca5 formal position for PThm nodes;
wenzelm
parents: 69593
diff changeset
  1668
            |> Thm.close_derivation \<^here>
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1669
          end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1670
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
  1671
        val rel_conversep = Lazy.lazy mk_rel_conversep;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1672
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
  1673
        fun mk_rel_OO () =
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1674
          Goal.prove_sorry lthy [] []
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1675
            (fold_rev Logic.all (Rs @ Ss) (HOLogic.mk_Trueprop (mk_leq rel_OO_lhs rel_OO_rhs)))
55197
5a54ed681ba2 less hermetic tactics
traytel
parents: 55061
diff changeset
  1676
            (fn {context = ctxt, prems = _} => mk_rel_OO_le_tac ctxt rel_OO_Grps (Lazy.force rel_eq)
5a54ed681ba2 less hermetic tactics
traytel
parents: 55061
diff changeset
  1677
              (#map_cong0 axioms) (Lazy.force map_comp) (map Lazy.force set_map))
70494
41108e3e9ca5 formal position for PThm nodes;
wenzelm
parents: 69593
diff changeset
  1678
          |> Thm.close_derivation \<^here>
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  1679
          |> (fn thm => @{thm antisym} OF [thm, #le_rel_OO axioms]);
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1680
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
  1681
        val rel_OO = Lazy.lazy mk_rel_OO;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1682
53561
92bcac4f9ac9 simplified derivation of in_rel
traytel
parents: 53290
diff changeset
  1683
        fun mk_in_rel () = trans OF [rel_OO_Grp, @{thm OO_Grp_alt}] RS @{thm predicate2_eqD};
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1684
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
  1685
        val in_rel = Lazy.lazy mk_in_rel;
49537
fe1deee434b6 generate "rel_as_srel" and "rel_flip" properties
blanchet
parents: 49536
diff changeset
  1686
fe1deee434b6 generate "rel_as_srel" and "rel_flip" properties
blanchet
parents: 49536
diff changeset
  1687
        fun mk_rel_flip () =
61760
1647bb489522 tuned whitespace
blanchet
parents: 61424
diff changeset
  1688
          unfold_thms lthy @{thms conversep_iff}
1647bb489522 tuned whitespace
blanchet
parents: 61424
diff changeset
  1689
            (Lazy.force rel_conversep RS @{thm predicate2_eqD});
49537
fe1deee434b6 generate "rel_as_srel" and "rel_flip" properties
blanchet
parents: 49536
diff changeset
  1690
49538
c90818b63599 simplified fact policies
blanchet
parents: 49537
diff changeset
  1691
        val rel_flip = Lazy.lazy mk_rel_flip;
49537
fe1deee434b6 generate "rel_as_srel" and "rel_flip" properties
blanchet
parents: 49536
diff changeset
  1692
57967
e6d2e998c30f renamed 'rel_mono_strong' to 'rel_mono_strong0'
desharna
parents: 57932
diff changeset
  1693
        fun mk_rel_mono_strong0 () =
51916
eac9e9a45bf5 stronger monotonicity property for relators
traytel
parents: 51915
diff changeset
  1694
          let
eac9e9a45bf5 stronger monotonicity property for relators
traytel
parents: 51915
diff changeset
  1695
            fun mk_prem setA setB R S a b =
eac9e9a45bf5 stronger monotonicity property for relators
traytel
parents: 51915
diff changeset
  1696
              HOLogic.mk_Trueprop
eac9e9a45bf5 stronger monotonicity property for relators
traytel
parents: 51915
diff changeset
  1697
                (mk_Ball (setA $ x) (Term.absfree (dest_Free a)
eac9e9a45bf5 stronger monotonicity property for relators
traytel
parents: 51915
diff changeset
  1698
                  (mk_Ball (setB $ y) (Term.absfree (dest_Free b)
eac9e9a45bf5 stronger monotonicity property for relators
traytel
parents: 51915
diff changeset
  1699
                    (HOLogic.mk_imp (R $ a $ b, S $ a $ b))))));
57301
7b997028aaac generate 'rel_coinduct0' theorem for codatatypes
desharna
parents: 56954
diff changeset
  1700
            val prems = HOLogic.mk_Trueprop (Term.list_comb (rel, Rs) $ x $ y) ::
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58432
diff changeset
  1701
              @{map 6} mk_prem bnf_sets_As bnf_sets_Bs Rs Rs_copy zs ys;
51916
eac9e9a45bf5 stronger monotonicity property for relators
traytel
parents: 51915
diff changeset
  1702
            val concl = HOLogic.mk_Trueprop (Term.list_comb (rel, Rs_copy) $ x $ y);
eac9e9a45bf5 stronger monotonicity property for relators
traytel
parents: 51915
diff changeset
  1703
          in
eac9e9a45bf5 stronger monotonicity property for relators
traytel
parents: 51915
diff changeset
  1704
            Goal.prove_sorry lthy [] []
eac9e9a45bf5 stronger monotonicity property for relators
traytel
parents: 51915
diff changeset
  1705
              (fold_rev Logic.all (x :: y :: Rs @ Rs_copy) (Logic.list_implies (prems, concl)))
57967
e6d2e998c30f renamed 'rel_mono_strong' to 'rel_mono_strong0'
desharna
parents: 57932
diff changeset
  1706
              (fn {context = ctxt, prems = _} => mk_rel_mono_strong0_tac ctxt (Lazy.force in_rel)
55197
5a54ed681ba2 less hermetic tactics
traytel
parents: 55061
diff changeset
  1707
                (map Lazy.force set_map))
70494
41108e3e9ca5 formal position for PThm nodes;
wenzelm
parents: 69593
diff changeset
  1708
            |> Thm.close_derivation \<^here>
51916
eac9e9a45bf5 stronger monotonicity property for relators
traytel
parents: 51915
diff changeset
  1709
          end;
eac9e9a45bf5 stronger monotonicity property for relators
traytel
parents: 51915
diff changeset
  1710
57967
e6d2e998c30f renamed 'rel_mono_strong' to 'rel_mono_strong0'
desharna
parents: 57932
diff changeset
  1711
        val rel_mono_strong0 = Lazy.lazy mk_rel_mono_strong0;
51916
eac9e9a45bf5 stronger monotonicity property for relators
traytel
parents: 51915
diff changeset
  1712
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1713
        val rel_mono_strong = Lazy.map (Object_Logic.rulify lthy) rel_mono_strong0;
57968
00e9c6d367e7 generate property 'rel_mono_strong' for BNFs
desharna
parents: 57967
diff changeset
  1714
61242
1f92b0ac9c96 congruence rules for the relator
traytel
parents: 61241
diff changeset
  1715
        fun mk_rel_cong_prem mk_implies x x' z z' set set' R R_copy =
1f92b0ac9c96 congruence rules for the relator
traytel
parents: 61241
diff changeset
  1716
          Logic.all z (Logic.all z'
1f92b0ac9c96 congruence rules for the relator
traytel
parents: 61241
diff changeset
  1717
            (mk_implies (mk_Trueprop_mem (z, set $ x), mk_implies (mk_Trueprop_mem (z', set' $ x'),
1f92b0ac9c96 congruence rules for the relator
traytel
parents: 61241
diff changeset
  1718
              mk_Trueprop_eq (R $ z $ z', R_copy $ z $ z')))));
1f92b0ac9c96 congruence rules for the relator
traytel
parents: 61241
diff changeset
  1719
1f92b0ac9c96 congruence rules for the relator
traytel
parents: 61241
diff changeset
  1720
        fun mk_rel_cong mk_implies () =
1f92b0ac9c96 congruence rules for the relator
traytel
parents: 61241
diff changeset
  1721
          let
1f92b0ac9c96 congruence rules for the relator
traytel
parents: 61241
diff changeset
  1722
            val prem0 = mk_Trueprop_eq (x, x_copy);
1f92b0ac9c96 congruence rules for the relator
traytel
parents: 61241
diff changeset
  1723
            val prem1 = mk_Trueprop_eq (y, y_copy);
1f92b0ac9c96 congruence rules for the relator
traytel
parents: 61241
diff changeset
  1724
            val prems = @{map 6} (mk_rel_cong_prem mk_implies x_copy y_copy)
1f92b0ac9c96 congruence rules for the relator
traytel
parents: 61241
diff changeset
  1725
              zs ys bnf_sets_As bnf_sets_Bs Rs Rs_copy;
1f92b0ac9c96 congruence rules for the relator
traytel
parents: 61241
diff changeset
  1726
            val eq = mk_Trueprop_eq (Term.list_comb (rel, Rs) $ x $ y,
1f92b0ac9c96 congruence rules for the relator
traytel
parents: 61241
diff changeset
  1727
              Term.list_comb (rel, Rs_copy) $ x_copy $ y_copy);
1f92b0ac9c96 congruence rules for the relator
traytel
parents: 61241
diff changeset
  1728
          in
61334
8d40ddaa427f collect the names from goals in favor of fragile exports
traytel
parents: 61301
diff changeset
  1729
            fold (Variable.add_free_names lthy) (eq :: prem0 :: prem1 :: prems) []
8d40ddaa427f collect the names from goals in favor of fragile exports
traytel
parents: 61301
diff changeset
  1730
            |> (fn vars => Goal.prove_sorry lthy vars (prem0 :: prem1 :: prems) eq
61242
1f92b0ac9c96 congruence rules for the relator
traytel
parents: 61241
diff changeset
  1731
              (fn {context = ctxt, prems} =>
61334
8d40ddaa427f collect the names from goals in favor of fragile exports
traytel
parents: 61301
diff changeset
  1732
                mk_rel_cong_tac ctxt (chop 2 prems) (Lazy.force rel_mono_strong)))
70494
41108e3e9ca5 formal position for PThm nodes;
wenzelm
parents: 69593
diff changeset
  1733
            |> Thm.close_derivation \<^here>
61242
1f92b0ac9c96 congruence rules for the relator
traytel
parents: 61241
diff changeset
  1734
          end;
1f92b0ac9c96 congruence rules for the relator
traytel
parents: 61241
diff changeset
  1735
1f92b0ac9c96 congruence rules for the relator
traytel
parents: 61241
diff changeset
  1736
        val rel_cong = Lazy.lazy (mk_rel_cong Logic.mk_implies);
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67222
diff changeset
  1737
        val rel_cong_simp = Lazy.lazy (mk_rel_cong (fn (a, b) => \<^term>\<open>simp_implies\<close> $ a $ b));
61242
1f92b0ac9c96 congruence rules for the relator
traytel
parents: 61241
diff changeset
  1738
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1739
        fun mk_pred_prems f = map2 (HOLogic.mk_Trueprop oo f) Ps Ps_copy;
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1740
        fun mk_pred_concl f = HOLogic.mk_Trueprop
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1741
          (f (Term.list_comb (pred, Ps), Term.list_comb (pred, Ps_copy)));
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1742
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1743
        fun mk_pred_cong0 () =
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1744
          let
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1745
            val cong_prems = mk_pred_prems (curry HOLogic.mk_eq);
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1746
            val cong_concl = mk_pred_concl HOLogic.mk_eq;
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1747
          in
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1748
            Goal.prove_sorry lthy [] []
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1749
              (fold_rev Logic.all (Ps @ Ps_copy) (Logic.list_implies (cong_prems, cong_concl)))
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1750
              (fn {context = ctxt, prems = _} => (TRY o hyp_subst_tac ctxt THEN' rtac ctxt refl) 1)
70494
41108e3e9ca5 formal position for PThm nodes;
wenzelm
parents: 69593
diff changeset
  1751
            |> Thm.close_derivation \<^here>
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1752
          end;
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1753
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1754
        val pred_cong0 = Lazy.lazy mk_pred_cong0;
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1755
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1756
        fun mk_rel_eq_onp () =
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1757
          let
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1758
            val lhs = Term.list_comb (relAsAs, map mk_eq_onp Ps);
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1759
            val rhs = mk_eq_onp (Term.list_comb (pred, Ps));
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1760
          in
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1761
            Goal.prove_sorry lthy (map fst Ps') [] (mk_Trueprop_eq (lhs, rhs))
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1762
              (fn {context = ctxt, prems = _} =>
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1763
                mk_rel_eq_onp_tac ctxt (#pred_set axioms) (#map_id0 axioms) (Lazy.force rel_Grp))
70494
41108e3e9ca5 formal position for PThm nodes;
wenzelm
parents: 69593
diff changeset
  1764
            |> Thm.close_derivation \<^here>
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1765
          end;
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1766
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1767
        val rel_eq_onp = Lazy.lazy mk_rel_eq_onp;
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1768
        val pred_rel = Lazy.map (fn thm => thm RS sym RS @{thm eq_onp_eqD}) rel_eq_onp;
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1769
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1770
        fun mk_pred_mono_strong0 () =
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1771
          let
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1772
            fun mk_prem setA P Q a =
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1773
              HOLogic.mk_Trueprop
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1774
                (mk_Ball (setA $ x) (Term.absfree (dest_Free a) (HOLogic.mk_imp (P $ a, Q $ a))));
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1775
            val prems = HOLogic.mk_Trueprop (Term.list_comb (pred, Ps) $ x) ::
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1776
              @{map 4} mk_prem bnf_sets_As Ps Ps_copy zs;
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1777
            val concl = HOLogic.mk_Trueprop (Term.list_comb (pred, Ps_copy) $ x);
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1778
          in
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1779
            Goal.prove_sorry lthy [] []
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1780
              (fold_rev Logic.all (x :: Ps @ Ps_copy) (Logic.list_implies (prems, concl)))
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1781
              (fn {context = ctxt, prems = _} =>
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1782
                mk_pred_mono_strong0_tac ctxt (Lazy.force pred_rel) (Lazy.force rel_mono_strong0))
70494
41108e3e9ca5 formal position for PThm nodes;
wenzelm
parents: 69593
diff changeset
  1783
            |> Thm.close_derivation \<^here>
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1784
          end;
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1785
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1786
        val pred_mono_strong0 = Lazy.lazy mk_pred_mono_strong0;
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1787
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1788
        val pred_mono_strong = Lazy.map (Object_Logic.rulify lthy) pred_mono_strong0;
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1789
63714
b62f4f765353 derive pred_mono property for BNFs
traytel
parents: 63399
diff changeset
  1790
        fun mk_pred_mono () =
b62f4f765353 derive pred_mono property for BNFs
traytel
parents: 63399
diff changeset
  1791
          let
b62f4f765353 derive pred_mono property for BNFs
traytel
parents: 63399
diff changeset
  1792
            val mono_prems = mk_pred_prems mk_leq;
b62f4f765353 derive pred_mono property for BNFs
traytel
parents: 63399
diff changeset
  1793
            val mono_concl = mk_pred_concl (uncurry mk_leq);
b62f4f765353 derive pred_mono property for BNFs
traytel
parents: 63399
diff changeset
  1794
          in
b62f4f765353 derive pred_mono property for BNFs
traytel
parents: 63399
diff changeset
  1795
            Goal.prove_sorry lthy [] []
b62f4f765353 derive pred_mono property for BNFs
traytel
parents: 63399
diff changeset
  1796
              (fold_rev Logic.all (Ps @ Ps_copy) (Logic.list_implies (mono_prems, mono_concl)))
b62f4f765353 derive pred_mono property for BNFs
traytel
parents: 63399
diff changeset
  1797
              (fn {context = ctxt, prems = _} =>
b62f4f765353 derive pred_mono property for BNFs
traytel
parents: 63399
diff changeset
  1798
                mk_pred_mono_tac ctxt (Lazy.force rel_eq_onp) (Lazy.force rel_mono))
70494
41108e3e9ca5 formal position for PThm nodes;
wenzelm
parents: 69593
diff changeset
  1799
            |> Thm.close_derivation \<^here>
63714
b62f4f765353 derive pred_mono property for BNFs
traytel
parents: 63399
diff changeset
  1800
          end;
b62f4f765353 derive pred_mono property for BNFs
traytel
parents: 63399
diff changeset
  1801
b62f4f765353 derive pred_mono property for BNFs
traytel
parents: 63399
diff changeset
  1802
        val pred_mono = Lazy.lazy mk_pred_mono;
b62f4f765353 derive pred_mono property for BNFs
traytel
parents: 63399
diff changeset
  1803
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1804
        fun mk_pred_cong_prem mk_implies x z set P P_copy =
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1805
          Logic.all z
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1806
            (mk_implies (mk_Trueprop_mem (z, set $ x), mk_Trueprop_eq (P $ z, P_copy $ z)));
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1807
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1808
        fun mk_pred_cong mk_implies () =
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1809
          let
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1810
            val prem0 = mk_Trueprop_eq (x, x_copy);
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1811
            val prems = @{map 4} (mk_pred_cong_prem mk_implies x_copy) zs bnf_sets_As Ps Ps_copy;
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1812
            val eq = mk_Trueprop_eq (Term.list_comb (pred, Ps) $ x,
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1813
              Term.list_comb (pred, Ps_copy) $ x_copy);
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1814
          in
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1815
            fold (Variable.add_free_names lthy) (eq :: prem0 :: prems) []
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1816
            |> (fn vars => Goal.prove_sorry lthy vars (prem0 :: prems) eq
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1817
              (fn {context = ctxt, prems} =>
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1818
                mk_rel_cong_tac ctxt (chop 1 prems) (Lazy.force pred_mono_strong)))
70494
41108e3e9ca5 formal position for PThm nodes;
wenzelm
parents: 69593
diff changeset
  1819
            |> Thm.close_derivation \<^here>
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1820
          end;
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1821
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1822
        val pred_cong = Lazy.lazy (mk_pred_cong Logic.mk_implies);
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67222
diff changeset
  1823
        val pred_cong_simp = Lazy.lazy (mk_pred_cong (fn (a, b) => \<^term>\<open>simp_implies\<close> $ a $ b));
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1824
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1825
        fun mk_map_cong_pred () =
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1826
          let
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1827
            val prem0 = mk_Trueprop_eq (x, x_copy);
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1828
            fun mk_eq f g z = Term.absfree (dest_Free z) (HOLogic.mk_eq (f $ z, g $ z));
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1829
            val prem = HOLogic.mk_Trueprop
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1830
              (Term.list_comb (pred, @{map 3} mk_eq fs fs_copy zs) $ x_copy);
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1831
            val eq = mk_Trueprop_eq (Term.list_comb (bnf_map_AsBs, fs) $ x,
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1832
              Term.list_comb (bnf_map_AsBs, fs_copy) $ x_copy);
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1833
            val goal = fold_rev Logic.all (x :: x_copy :: fs @ fs_copy)
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1834
              (Logic.list_implies ([prem0, prem], eq));
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1835
          in
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1836
            Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} =>
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1837
              unfold_thms_tac ctxt [#pred_set axioms] THEN
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1838
              HEADGOAL (EVERY' [REPEAT_DETERM o etac ctxt conjE,
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1839
                etac ctxt (Lazy.force map_cong) THEN_ALL_NEW
63399
d1742d1b7f0f more antiquotations;
wenzelm
parents: 63170
diff changeset
  1840
                  (etac ctxt @{thm bspec} THEN' assume_tac ctxt)]))
70494
41108e3e9ca5 formal position for PThm nodes;
wenzelm
parents: 69593
diff changeset
  1841
            |> Thm.close_derivation \<^here>
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1842
          end;
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1843
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1844
        val map_cong_pred = Lazy.lazy mk_map_cong_pred;
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1845
57932
c29659f77f8d generate 'rel_map' theorem for BNFs
desharna
parents: 57668
diff changeset
  1846
        fun mk_rel_map () =
c29659f77f8d generate 'rel_map' theorem for BNFs
desharna
parents: 57668
diff changeset
  1847
          let
c29659f77f8d generate 'rel_map' theorem for BNFs
desharna
parents: 57668
diff changeset
  1848
            fun mk_goal lhs rhs =
c29659f77f8d generate 'rel_map' theorem for BNFs
desharna
parents: 57668
diff changeset
  1849
              fold_rev Logic.all ([x, y] @ S_CsBs @ S_AsCs @ is @ gs) (mk_Trueprop_eq (lhs, rhs));
c29659f77f8d generate 'rel_map' theorem for BNFs
desharna
parents: 57668
diff changeset
  1850
c29659f77f8d generate 'rel_map' theorem for BNFs
desharna
parents: 57668
diff changeset
  1851
            val lhss =
c29659f77f8d generate 'rel_map' theorem for BNFs
desharna
parents: 57668
diff changeset
  1852
              [Term.list_comb (relCsBs, S_CsBs) $ (Term.list_comb (bnf_map_AsCs, is) $ x) $ y,
c29659f77f8d generate 'rel_map' theorem for BNFs
desharna
parents: 57668
diff changeset
  1853
               Term.list_comb (relAsCs, S_AsCs) $ x $ (Term.list_comb (bnf_map_BsCs, gs) $ y)];
c29659f77f8d generate 'rel_map' theorem for BNFs
desharna
parents: 57668
diff changeset
  1854
            val rhss =
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58432
diff changeset
  1855
              [Term.list_comb (rel, @{map 3} (fn f => fn P => fn T =>
57932
c29659f77f8d generate 'rel_map' theorem for BNFs
desharna
parents: 57668
diff changeset
  1856
                 mk_vimage2p f (HOLogic.id_const T) $ P) is S_CsBs Bs') $ x $ y,
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58432
diff changeset
  1857
               Term.list_comb (rel, @{map 3} (fn f => fn P => fn T =>
57932
c29659f77f8d generate 'rel_map' theorem for BNFs
desharna
parents: 57668
diff changeset
  1858
                 mk_vimage2p (HOLogic.id_const T) f $ P) gs S_AsCs As') $ x $ y];
c29659f77f8d generate 'rel_map' theorem for BNFs
desharna
parents: 57668
diff changeset
  1859
            val goals = map2 mk_goal lhss rhss;
c29659f77f8d generate 'rel_map' theorem for BNFs
desharna
parents: 57668
diff changeset
  1860
          in
57970
eaa986cd285a generate 'inj_map_strong' for BNFs
desharna
parents: 57969
diff changeset
  1861
            goals
eaa986cd285a generate 'inj_map_strong' for BNFs
desharna
parents: 57969
diff changeset
  1862
            |> map (fn goal => Goal.prove_sorry lthy [] [] goal
57932
c29659f77f8d generate 'rel_map' theorem for BNFs
desharna
parents: 57668
diff changeset
  1863
              (fn {context = ctxt, prems = _} =>
c29659f77f8d generate 'rel_map' theorem for BNFs
desharna
parents: 57668
diff changeset
  1864
                 mk_rel_map0_tac ctxt live (Lazy.force rel_OO) (Lazy.force rel_conversep)
57970
eaa986cd285a generate 'inj_map_strong' for BNFs
desharna
parents: 57969
diff changeset
  1865
                  (Lazy.force rel_Grp) (Lazy.force map_id)))
63170
eae6549dbea2 tuned proofs, to allow unfold_abs_def;
wenzelm
parents: 63009
diff changeset
  1866
            |> map (unfold_thms lthy @{thms vimage2p_def[of id, simplified id_apply]
eae6549dbea2 tuned proofs, to allow unfold_abs_def;
wenzelm
parents: 63009
diff changeset
  1867
                 vimage2p_def[of _ id, simplified id_apply]})
70494
41108e3e9ca5 formal position for PThm nodes;
wenzelm
parents: 69593
diff changeset
  1868
            |> map (Thm.close_derivation \<^here>)
57932
c29659f77f8d generate 'rel_map' theorem for BNFs
desharna
parents: 57668
diff changeset
  1869
          end;
c29659f77f8d generate 'rel_map' theorem for BNFs
desharna
parents: 57668
diff changeset
  1870
c29659f77f8d generate 'rel_map' theorem for BNFs
desharna
parents: 57668
diff changeset
  1871
        val rel_map = Lazy.lazy mk_rel_map;
c29659f77f8d generate 'rel_map' theorem for BNFs
desharna
parents: 57668
diff changeset
  1872
59726
64c2bb331035 BNF relators preserve reflexivity
traytel
parents: 59663
diff changeset
  1873
        fun mk_rel_refl () = @{thm ge_eq_refl[OF ord_eq_le_trans]} OF
64c2bb331035 BNF relators preserve reflexivity
traytel
parents: 59663
diff changeset
  1874
          [Lazy.force rel_eq RS sym, Lazy.force rel_mono OF (replicate live @{thm refl_ge_eq})];
64c2bb331035 BNF relators preserve reflexivity
traytel
parents: 59663
diff changeset
  1875
64c2bb331035 BNF relators preserve reflexivity
traytel
parents: 59663
diff changeset
  1876
        val rel_refl = Lazy.lazy mk_rel_refl;
64c2bb331035 BNF relators preserve reflexivity
traytel
parents: 59663
diff changeset
  1877
61240
464c5da3f508 more useful properties of the relators
traytel
parents: 61116
diff changeset
  1878
        fun mk_rel_refl_strong () =
464c5da3f508 more useful properties of the relators
traytel
parents: 61116
diff changeset
  1879
          (rule_by_tactic lthy (ALLGOALS (Object_Logic.full_atomize_tac lthy))
464c5da3f508 more useful properties of the relators
traytel
parents: 61116
diff changeset
  1880
            ((Lazy.force rel_eq RS @{thm predicate2_eqD}) RS @{thm iffD2[OF _ refl]} RS
464c5da3f508 more useful properties of the relators
traytel
parents: 61116
diff changeset
  1881
              Lazy.force rel_mono_strong)) OF
464c5da3f508 more useful properties of the relators
traytel
parents: 61116
diff changeset
  1882
            (replicate live @{thm diag_imp_eq_le})
464c5da3f508 more useful properties of the relators
traytel
parents: 61116
diff changeset
  1883
464c5da3f508 more useful properties of the relators
traytel
parents: 61116
diff changeset
  1884
        val rel_refl_strong = Lazy.lazy mk_rel_refl_strong;
464c5da3f508 more useful properties of the relators
traytel
parents: 61116
diff changeset
  1885
464c5da3f508 more useful properties of the relators
traytel
parents: 61116
diff changeset
  1886
        fun mk_rel_preserves mk_prop prop_conv_thm thm () =
464c5da3f508 more useful properties of the relators
traytel
parents: 61116
diff changeset
  1887
          let
464c5da3f508 more useful properties of the relators
traytel
parents: 61116
diff changeset
  1888
            val Rs = map2 retype_const_or_free self_pred2RTs Rs;
464c5da3f508 more useful properties of the relators
traytel
parents: 61116
diff changeset
  1889
            val prems = map (HOLogic.mk_Trueprop o mk_prop) Rs;
464c5da3f508 more useful properties of the relators
traytel
parents: 61116
diff changeset
  1890
            val goal = HOLogic.mk_Trueprop (mk_prop (Term.list_comb (relAsAs, Rs)));
61334
8d40ddaa427f collect the names from goals in favor of fragile exports
traytel
parents: 61301
diff changeset
  1891
        val vars = fold (Variable.add_free_names lthy) (goal :: prems) [];
61240
464c5da3f508 more useful properties of the relators
traytel
parents: 61116
diff changeset
  1892
          in
61334
8d40ddaa427f collect the names from goals in favor of fragile exports
traytel
parents: 61301
diff changeset
  1893
            Goal.prove_sorry lthy vars [] (Logic.list_implies (prems, goal))
61240
464c5da3f508 more useful properties of the relators
traytel
parents: 61116
diff changeset
  1894
              (fn {context = ctxt, prems = _} =>
464c5da3f508 more useful properties of the relators
traytel
parents: 61116
diff changeset
  1895
                unfold_thms_tac ctxt [prop_conv_thm] THEN
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1896
                HEADGOAL (rtac ctxt (Lazy.force thm RS sym RS @{thm ord_eq_le_trans})
61240
464c5da3f508 more useful properties of the relators
traytel
parents: 61116
diff changeset
  1897
                  THEN' rtac ctxt (Lazy.force rel_mono) THEN_ALL_NEW assume_tac ctxt))
70494
41108e3e9ca5 formal position for PThm nodes;
wenzelm
parents: 69593
diff changeset
  1898
            |> Thm.close_derivation \<^here>
61240
464c5da3f508 more useful properties of the relators
traytel
parents: 61116
diff changeset
  1899
          end;
464c5da3f508 more useful properties of the relators
traytel
parents: 61116
diff changeset
  1900
464c5da3f508 more useful properties of the relators
traytel
parents: 61116
diff changeset
  1901
        val rel_reflp = Lazy.lazy (mk_rel_preserves mk_reflp @{thm reflp_eq} rel_eq);
464c5da3f508 more useful properties of the relators
traytel
parents: 61116
diff changeset
  1902
        val rel_symp = Lazy.lazy (mk_rel_preserves mk_symp @{thm symp_conversep} rel_conversep);
464c5da3f508 more useful properties of the relators
traytel
parents: 61116
diff changeset
  1903
        val rel_transp = Lazy.lazy (mk_rel_preserves mk_transp @{thm transp_relcompp} rel_OO);
464c5da3f508 more useful properties of the relators
traytel
parents: 61116
diff changeset
  1904
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1905
        fun mk_pred_True () =
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1906
          let
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67222
diff changeset
  1907
            val lhs = Term.list_comb (pred, map (fn T => absdummy T \<^term>\<open>True\<close>) As');
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67222
diff changeset
  1908
            val rhs = absdummy CA' \<^term>\<open>True\<close>;
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1909
            val goal = mk_Trueprop_eq (lhs, rhs);
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1910
          in
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1911
            Goal.prove_sorry lthy [] [] goal
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1912
              (fn {context = ctxt, prems = _} =>
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1913
                HEADGOAL (EVERY' (map (rtac ctxt) [ext, Lazy.force pred_rel RS trans,
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1914
                  Lazy.force rel_cong0 RS fun_cong RS fun_cong RS trans OF
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1915
                    replicate live @{thm eq_onp_True},
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1916
                  Lazy.force rel_eq RS fun_cong RS fun_cong RS trans, @{thm eqTrueI[OF refl]}])))
70494
41108e3e9ca5 formal position for PThm nodes;
wenzelm
parents: 69593
diff changeset
  1917
            |> Thm.close_derivation \<^here>
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1918
          end;
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1919
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1920
        val pred_True = Lazy.lazy mk_pred_True;
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1921
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1922
        fun mk_pred_map () =
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1923
          let
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1924
            val lhs = Term.list_comb (pred', Qs) $ (Term.list_comb (bnf_map_AsBs, fs) $ x);
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1925
            val rhs = Term.list_comb (pred, @{map 2} (curry HOLogic.mk_comp) Qs fs) $ x;
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1926
            val goal = mk_Trueprop_eq (lhs, rhs);
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1927
            val vars = Variable.add_free_names lthy goal [];
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1928
            val pred_set = #pred_set axioms RS fun_cong RS sym;
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1929
          in
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1930
            Goal.prove_sorry lthy vars [] goal
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1931
              (fn {context = ctxt, prems = _} =>
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1932
                HEADGOAL (rtac ctxt (pred_set RSN (2, pred_set RSN (2, box_equals)))) THEN
67222
19809bc9d7ff made tactics more robust
traytel
parents: 66272
diff changeset
  1933
                unfold_thms_tac ctxt
19809bc9d7ff made tactics more robust
traytel
parents: 66272
diff changeset
  1934
                  (@{thms Ball_image_comp ball_empty} @ map Lazy.force set_map) THEN
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1935
                HEADGOAL (rtac ctxt refl))
70494
41108e3e9ca5 formal position for PThm nodes;
wenzelm
parents: 69593
diff changeset
  1936
            |> Thm.close_derivation \<^here>
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1937
          end;
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1938
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1939
        val pred_map = Lazy.lazy mk_pred_map;
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  1940
52719
480a3479fa47 transfer rule for map (not yet registered as a transfer rule)
traytel
parents: 52635
diff changeset
  1941
        fun mk_map_transfer () =
480a3479fa47 transfer rule for map (not yet registered as a transfer rule)
traytel
parents: 52635
diff changeset
  1942
          let
55945
e96383acecf9 renamed 'fun_rel' to 'rel_fun'
blanchet
parents: 55854
diff changeset
  1943
            val rels = map2 mk_rel_fun transfer_domRs transfer_ranRs;
e96383acecf9 renamed 'fun_rel' to 'rel_fun'
blanchet
parents: 55854
diff changeset
  1944
            val rel = mk_rel_fun
52725
ba2bbe825a5e proper transfer rule format for map_transfer
traytel
parents: 52721
diff changeset
  1945
              (Term.list_comb (mk_bnf_rel transfer_domRTs CA' CB1, transfer_domRs))
ba2bbe825a5e proper transfer rule format for map_transfer
traytel
parents: 52721
diff changeset
  1946
              (Term.list_comb (mk_bnf_rel transfer_ranRTs CB' CB2, transfer_ranRs));
ba2bbe825a5e proper transfer rule format for map_transfer
traytel
parents: 52721
diff changeset
  1947
            val concl = HOLogic.mk_Trueprop
55945
e96383acecf9 renamed 'fun_rel' to 'rel_fun'
blanchet
parents: 55854
diff changeset
  1948
              (fold_rev mk_rel_fun rels rel $ bnf_map_AsBs $ mk_bnf_map B1Ts B2Ts);
52719
480a3479fa47 transfer rule for map (not yet registered as a transfer rule)
traytel
parents: 52635
diff changeset
  1949
          in
480a3479fa47 transfer rule for map (not yet registered as a transfer rule)
traytel
parents: 52635
diff changeset
  1950
            Goal.prove_sorry lthy [] []
52725
ba2bbe825a5e proper transfer rule format for map_transfer
traytel
parents: 52721
diff changeset
  1951
              (fold_rev Logic.all (transfer_domRs @ transfer_ranRs) concl)
55197
5a54ed681ba2 less hermetic tactics
traytel
parents: 55061
diff changeset
  1952
              (fn {context = ctxt, prems = _} => mk_map_transfer_tac ctxt (Lazy.force rel_mono)
5a54ed681ba2 less hermetic tactics
traytel
parents: 55061
diff changeset
  1953
                (Lazy.force in_rel) (map Lazy.force set_map) (#map_cong0 axioms)
5a54ed681ba2 less hermetic tactics
traytel
parents: 55061
diff changeset
  1954
                (Lazy.force map_comp))
70494
41108e3e9ca5 formal position for PThm nodes;
wenzelm
parents: 69593
diff changeset
  1955
            |> Thm.close_derivation \<^here>
52719
480a3479fa47 transfer rule for map (not yet registered as a transfer rule)
traytel
parents: 52635
diff changeset
  1956
          end;
480a3479fa47 transfer rule for map (not yet registered as a transfer rule)
traytel
parents: 52635
diff changeset
  1957
480a3479fa47 transfer rule for map (not yet registered as a transfer rule)
traytel
parents: 52635
diff changeset
  1958
        val map_transfer = Lazy.lazy mk_map_transfer;
480a3479fa47 transfer rule for map (not yet registered as a transfer rule)
traytel
parents: 52635
diff changeset
  1959
62329
9f7fa08d2307 derive transfer rule for predicator
traytel
parents: 62326
diff changeset
  1960
        fun mk_pred_transfer () =
9f7fa08d2307 derive transfer rule for predicator
traytel
parents: 62326
diff changeset
  1961
          let
9f7fa08d2307 derive transfer rule for predicator
traytel
parents: 62326
diff changeset
  1962
            val iff = HOLogic.eq_const HOLogic.boolT;
9f7fa08d2307 derive transfer rule for predicator
traytel
parents: 62326
diff changeset
  1963
            val prem_rels = map (fn T => mk_rel_fun T iff) Rs;
9f7fa08d2307 derive transfer rule for predicator
traytel
parents: 62326
diff changeset
  1964
            val prem_elems = mk_rel_fun (Term.list_comb (rel, Rs)) iff;
9f7fa08d2307 derive transfer rule for predicator
traytel
parents: 62326
diff changeset
  1965
            val goal = HOLogic.mk_Trueprop
9f7fa08d2307 derive transfer rule for predicator
traytel
parents: 62326
diff changeset
  1966
              (fold_rev mk_rel_fun prem_rels prem_elems $ pred $ pred');
9f7fa08d2307 derive transfer rule for predicator
traytel
parents: 62326
diff changeset
  1967
            val vars = Variable.add_free_names lthy goal [];
9f7fa08d2307 derive transfer rule for predicator
traytel
parents: 62326
diff changeset
  1968
          in
9f7fa08d2307 derive transfer rule for predicator
traytel
parents: 62326
diff changeset
  1969
            Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} =>
9f7fa08d2307 derive transfer rule for predicator
traytel
parents: 62326
diff changeset
  1970
              mk_pred_transfer_tac ctxt live (Lazy.force in_rel) (Lazy.force pred_map)
9f7fa08d2307 derive transfer rule for predicator
traytel
parents: 62326
diff changeset
  1971
                (Lazy.force pred_cong))
70494
41108e3e9ca5 formal position for PThm nodes;
wenzelm
parents: 69593
diff changeset
  1972
            |> Thm.close_derivation \<^here>
62329
9f7fa08d2307 derive transfer rule for predicator
traytel
parents: 62326
diff changeset
  1973
          end;
9f7fa08d2307 derive transfer rule for predicator
traytel
parents: 62326
diff changeset
  1974
9f7fa08d2307 derive transfer rule for predicator
traytel
parents: 62326
diff changeset
  1975
        val pred_transfer = Lazy.lazy mk_pred_transfer;
9f7fa08d2307 derive transfer rule for predicator
traytel
parents: 62326
diff changeset
  1976
58104
c5316f843f72 generate 'rel_transfer' for BNFs
desharna
parents: 58102
diff changeset
  1977
        fun mk_rel_transfer () =
c5316f843f72 generate 'rel_transfer' for BNFs
desharna
parents: 58102
diff changeset
  1978
          let
c5316f843f72 generate 'rel_transfer' for BNFs
desharna
parents: 58102
diff changeset
  1979
            val iff = HOLogic.eq_const HOLogic.boolT;
c5316f843f72 generate 'rel_transfer' for BNFs
desharna
parents: 58102
diff changeset
  1980
            val prem_rels =
c5316f843f72 generate 'rel_transfer' for BNFs
desharna
parents: 58102
diff changeset
  1981
              map2 (fn T1 => fn T2 => mk_rel_fun T1 (mk_rel_fun T2 iff)) S_AsCs S_BsEs;
c5316f843f72 generate 'rel_transfer' for BNFs
desharna
parents: 58102
diff changeset
  1982
            val prem_elems =
c5316f843f72 generate 'rel_transfer' for BNFs
desharna
parents: 58102
diff changeset
  1983
              mk_rel_fun (Term.list_comb (mk_bnf_rel pred2RTsAsCs CA' CC', S_AsCs))
c5316f843f72 generate 'rel_transfer' for BNFs
desharna
parents: 58102
diff changeset
  1984
                (mk_rel_fun (Term.list_comb (mk_bnf_rel pred2RTsBsEs CB' CE', S_BsEs)) iff);
c5316f843f72 generate 'rel_transfer' for BNFs
desharna
parents: 58102
diff changeset
  1985
            val goal =
c5316f843f72 generate 'rel_transfer' for BNFs
desharna
parents: 58102
diff changeset
  1986
              HOLogic.mk_Trueprop (fold_rev mk_rel_fun prem_rels prem_elems $ rel $ relCsEs);
61334
8d40ddaa427f collect the names from goals in favor of fragile exports
traytel
parents: 61301
diff changeset
  1987
            val vars = Variable.add_free_names lthy goal [];
58104
c5316f843f72 generate 'rel_transfer' for BNFs
desharna
parents: 58102
diff changeset
  1988
          in
61334
8d40ddaa427f collect the names from goals in favor of fragile exports
traytel
parents: 61301
diff changeset
  1989
            Goal.prove_sorry lthy vars [] goal
58104
c5316f843f72 generate 'rel_transfer' for BNFs
desharna
parents: 58102
diff changeset
  1990
              (fn {context = ctxt, prems = _} =>
c5316f843f72 generate 'rel_transfer' for BNFs
desharna
parents: 58102
diff changeset
  1991
                mk_rel_transfer_tac ctxt (Lazy.force in_rel) (Lazy.force rel_map)
c5316f843f72 generate 'rel_transfer' for BNFs
desharna
parents: 58102
diff changeset
  1992
                  (Lazy.force rel_mono_strong))
70494
41108e3e9ca5 formal position for PThm nodes;
wenzelm
parents: 69593
diff changeset
  1993
            |> Thm.close_derivation \<^here>
58104
c5316f843f72 generate 'rel_transfer' for BNFs
desharna
parents: 58102
diff changeset
  1994
          end;
c5316f843f72 generate 'rel_transfer' for BNFs
desharna
parents: 58102
diff changeset
  1995
c5316f843f72 generate 'rel_transfer' for BNFs
desharna
parents: 58102
diff changeset
  1996
        val rel_transfer = Lazy.lazy mk_rel_transfer;
c5316f843f72 generate 'rel_transfer' for BNFs
desharna
parents: 58102
diff changeset
  1997
58106
c8cba801c483 generate 'set_transfer' for BNFs
desharna
parents: 58104
diff changeset
  1998
        fun mk_set_transfer () =
c8cba801c483 generate 'set_transfer' for BNFs
desharna
parents: 58104
diff changeset
  1999
          let
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67222
diff changeset
  2000
            val rel_sets = map2 (fn A => fn B => mk_rel 1 [A] [B] \<^term>\<open>rel_set\<close>) As' Bs';
58106
c8cba801c483 generate 'set_transfer' for BNFs
desharna
parents: 58104
diff changeset
  2001
            val rel_Rs = Term.list_comb (rel, Rs);
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58432
diff changeset
  2002
            val goals = @{map 4} (fn R => fn rel_set => fn setA => fn setB => HOLogic.mk_Trueprop
58106
c8cba801c483 generate 'set_transfer' for BNFs
desharna
parents: 58104
diff changeset
  2003
              (mk_rel_fun rel_Rs (rel_set $ R) $ setA $ setB)) Rs rel_sets bnf_sets_As bnf_sets_Bs;
c8cba801c483 generate 'set_transfer' for BNFs
desharna
parents: 58104
diff changeset
  2004
          in
c8cba801c483 generate 'set_transfer' for BNFs
desharna
parents: 58104
diff changeset
  2005
            if null goals then []
c8cba801c483 generate 'set_transfer' for BNFs
desharna
parents: 58104
diff changeset
  2006
            else
61334
8d40ddaa427f collect the names from goals in favor of fragile exports
traytel
parents: 61301
diff changeset
  2007
              let
8d40ddaa427f collect the names from goals in favor of fragile exports
traytel
parents: 61301
diff changeset
  2008
                val goal = Logic.mk_conjunction_balanced goals;
8d40ddaa427f collect the names from goals in favor of fragile exports
traytel
parents: 61301
diff changeset
  2009
                val vars = Variable.add_free_names lthy goal [];
8d40ddaa427f collect the names from goals in favor of fragile exports
traytel
parents: 61301
diff changeset
  2010
              in
8d40ddaa427f collect the names from goals in favor of fragile exports
traytel
parents: 61301
diff changeset
  2011
                Goal.prove_sorry lthy vars [] goal
8d40ddaa427f collect the names from goals in favor of fragile exports
traytel
parents: 61301
diff changeset
  2012
                  (fn {context = ctxt, prems = _} =>
8d40ddaa427f collect the names from goals in favor of fragile exports
traytel
parents: 61301
diff changeset
  2013
                     mk_set_transfer_tac ctxt (Lazy.force in_rel) (map Lazy.force set_map))
70494
41108e3e9ca5 formal position for PThm nodes;
wenzelm
parents: 69593
diff changeset
  2014
                |> Thm.close_derivation \<^here>
61334
8d40ddaa427f collect the names from goals in favor of fragile exports
traytel
parents: 61301
diff changeset
  2015
                |> Conjunction.elim_balanced (length goals)
8d40ddaa427f collect the names from goals in favor of fragile exports
traytel
parents: 61301
diff changeset
  2016
              end
58106
c8cba801c483 generate 'set_transfer' for BNFs
desharna
parents: 58104
diff changeset
  2017
          end;
c8cba801c483 generate 'set_transfer' for BNFs
desharna
parents: 58104
diff changeset
  2018
c8cba801c483 generate 'set_transfer' for BNFs
desharna
parents: 58104
diff changeset
  2019
        val set_transfer = Lazy.lazy mk_set_transfer;
c8cba801c483 generate 'set_transfer' for BNFs
desharna
parents: 58104
diff changeset
  2020
57970
eaa986cd285a generate 'inj_map_strong' for BNFs
desharna
parents: 57969
diff changeset
  2021
        fun mk_inj_map_strong () =
eaa986cd285a generate 'inj_map_strong' for BNFs
desharna
parents: 57969
diff changeset
  2022
          let
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58432
diff changeset
  2023
            val assms = @{map 5} (fn setA => fn z => fn f => fn z' => fn f' =>
57970
eaa986cd285a generate 'inj_map_strong' for BNFs
desharna
parents: 57969
diff changeset
  2024
              fold_rev Logic.all [z, z']
eaa986cd285a generate 'inj_map_strong' for BNFs
desharna
parents: 57969
diff changeset
  2025
                (Logic.mk_implies (mk_Trueprop_mem (z, setA $ x),
eaa986cd285a generate 'inj_map_strong' for BNFs
desharna
parents: 57969
diff changeset
  2026
                   Logic.mk_implies (mk_Trueprop_mem (z', setA $ x'),
eaa986cd285a generate 'inj_map_strong' for BNFs
desharna
parents: 57969
diff changeset
  2027
                     Logic.mk_implies (mk_Trueprop_eq (f $ z, f' $ z'),
eaa986cd285a generate 'inj_map_strong' for BNFs
desharna
parents: 57969
diff changeset
  2028
                       mk_Trueprop_eq (z, z')))))) bnf_sets_As zs fs zs' fs';
eaa986cd285a generate 'inj_map_strong' for BNFs
desharna
parents: 57969
diff changeset
  2029
            val concl = Logic.mk_implies
eaa986cd285a generate 'inj_map_strong' for BNFs
desharna
parents: 57969
diff changeset
  2030
              (mk_Trueprop_eq
eaa986cd285a generate 'inj_map_strong' for BNFs
desharna
parents: 57969
diff changeset
  2031
                 (Term.list_comb (bnf_map_AsBs, fs) $ x,
eaa986cd285a generate 'inj_map_strong' for BNFs
desharna
parents: 57969
diff changeset
  2032
                  Term.list_comb (bnf_map_AsBs, fs') $ x'),
eaa986cd285a generate 'inj_map_strong' for BNFs
desharna
parents: 57969
diff changeset
  2033
               mk_Trueprop_eq (x, x'));
eaa986cd285a generate 'inj_map_strong' for BNFs
desharna
parents: 57969
diff changeset
  2034
            val goal = fold_rev Logic.all (x :: x' :: fs @ fs')
eaa986cd285a generate 'inj_map_strong' for BNFs
desharna
parents: 57969
diff changeset
  2035
              (fold_rev (curry Logic.mk_implies) assms concl);
eaa986cd285a generate 'inj_map_strong' for BNFs
desharna
parents: 57969
diff changeset
  2036
          in
eaa986cd285a generate 'inj_map_strong' for BNFs
desharna
parents: 57969
diff changeset
  2037
            Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} =>
eaa986cd285a generate 'inj_map_strong' for BNFs
desharna
parents: 57969
diff changeset
  2038
              mk_inj_map_strong_tac ctxt (Lazy.force rel_eq) (Lazy.force rel_map)
eaa986cd285a generate 'inj_map_strong' for BNFs
desharna
parents: 57969
diff changeset
  2039
                (Lazy.force rel_mono_strong))
70494
41108e3e9ca5 formal position for PThm nodes;
wenzelm
parents: 69593
diff changeset
  2040
            |> Thm.close_derivation \<^here>
57970
eaa986cd285a generate 'inj_map_strong' for BNFs
desharna
parents: 57969
diff changeset
  2041
          end;
eaa986cd285a generate 'inj_map_strong' for BNFs
desharna
parents: 57969
diff changeset
  2042
eaa986cd285a generate 'inj_map_strong' for BNFs
desharna
parents: 57969
diff changeset
  2043
        val inj_map_strong = Lazy.lazy mk_inj_map_strong;
eaa986cd285a generate 'inj_map_strong' for BNFs
desharna
parents: 57969
diff changeset
  2044
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  2045
        val defs = mk_defs bnf_map_def bnf_set_defs bnf_rel_def bnf_pred_def;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2046
52635
4f84b730c489 got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
traytel
parents: 51930
diff changeset
  2047
        val facts = mk_facts bd_Card_order bd_Cinfinite bd_Cnotzero collect_set_map in_bd in_cong
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  2048
          in_mono in_rel inj_map inj_map_strong map_comp map_cong map_cong_simp map_cong_pred map_id
75624
22d1c5f2b9f4 strict bounds for BNFs (by Jan van Brügge)
traytel
parents: 75276
diff changeset
  2049
          map_ident0 map_ident map_ident_strong map_transfer rel_eq rel_flip set_map rel_cong0 rel_cong
22d1c5f2b9f4 strict bounds for BNFs (by Jan van Brügge)
traytel
parents: 75276
diff changeset
  2050
          rel_cong_simp rel_map rel_mono rel_mono_strong0 rel_mono_strong set_transfer rel_Grp rel_conversep
22d1c5f2b9f4 strict bounds for BNFs (by Jan van Brügge)
traytel
parents: 75276
diff changeset
  2051
          rel_OO rel_refl rel_refl_strong rel_reflp rel_symp rel_transp rel_transfer rel_eq_onp
22d1c5f2b9f4 strict bounds for BNFs (by Jan van Brügge)
traytel
parents: 75276
diff changeset
  2052
          pred_transfer pred_True pred_map pred_rel pred_mono_strong0 pred_mono_strong pred_mono
22d1c5f2b9f4 strict bounds for BNFs (by Jan van Brügge)
traytel
parents: 75276
diff changeset
  2053
          pred_cong0 pred_cong pred_cong_simp;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2054
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2055
        val wits = map2 mk_witness bnf_wits wit_thms;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2056
49507
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49506
diff changeset
  2057
        val bnf_rel =
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49506
diff changeset
  2058
          Term.subst_atomic_types ((Ds ~~ deads) @ (As' ~~ alphas) @ (Bs' ~~ betas)) rel;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2059
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  2060
        val bnf_pred = Term.subst_atomic_types ((Ds ~~ deads) @ (As' ~~ alphas)) pred;
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  2061
54841
af71b753c459 express weak pullback property of bnfs only in terms of the relator
traytel
parents: 54740
diff changeset
  2062
        val bnf = mk_bnf bnf_b Calpha live alphas betas dead deads bnf_map bnf_sets bnf_bd axioms
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  2063
          defs facts wits bnf_rel bnf_pred;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2064
      in
57631
959caab43a3d use the noted theorem whenever possible, also in 'BNF_Def'
blanchet
parents: 57567
diff changeset
  2065
        note_bnf_thms fact_policy qualify bnf_b bnf lthy
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2066
      end;
49459
3f8e2b5249ec adapted FP code to new relator approach
blanchet
parents: 49458
diff changeset
  2067
3f8e2b5249ec adapted FP code to new relator approach
blanchet
parents: 49458
diff changeset
  2068
    val one_step_defs =
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  2069
      no_reflexive (bnf_map_def :: bnf_bd_def :: bnf_set_defs @ bnf_wit_defs @
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  2070
        [bnf_rel_def, bnf_pred_def]);
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2071
  in
49459
3f8e2b5249ec adapted FP code to new relator approach
blanchet
parents: 49458
diff changeset
  2072
    (key, goals, wit_goalss, after_qed, lthy, one_step_defs)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2073
  end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2074
58659
6c9821c32dd5 Local_Interpretation is superseded by Plugin with formal Plugin_Name management, avoiding undeclared strings;
wenzelm
parents: 58634
diff changeset
  2075
structure BNF_Plugin = Plugin(type T = bnf);
56346
42533f8f4729 added BNF interpretation hook
blanchet
parents: 56016
diff changeset
  2076
58659
6c9821c32dd5 Local_Interpretation is superseded by Plugin with formal Plugin_Name management, avoiding undeclared strings;
wenzelm
parents: 58634
diff changeset
  2077
fun bnf_interpretation name f =
6c9821c32dd5 Local_Interpretation is superseded by Plugin with formal Plugin_Name management, avoiding undeclared strings;
wenzelm
parents: 58634
diff changeset
  2078
  BNF_Plugin.interpretation name
6c9821c32dd5 Local_Interpretation is superseded by Plugin with formal Plugin_Name management, avoiding undeclared strings;
wenzelm
parents: 58634
diff changeset
  2079
    (fn bnf => fn lthy => f (transfer_bnf (Proof_Context.theory_of lthy) bnf) lthy);
56376
5a93b8f928a2 added same idiomatic handling of namings for Ctr_Sugar/BNF-related interpretation hooks as for typedef and (old-style) datatypes
blanchet
parents: 56346
diff changeset
  2080
58659
6c9821c32dd5 Local_Interpretation is superseded by Plugin with formal Plugin_Name management, avoiding undeclared strings;
wenzelm
parents: 58634
diff changeset
  2081
val interpret_bnf = BNF_Plugin.data;
56346
42533f8f4729 added BNF interpretation hook
blanchet
parents: 56016
diff changeset
  2082
58177
166131276380 introduced local interpretation mechanism for BNFs, to solve issues with datatypes in locales
blanchet
parents: 58175
diff changeset
  2083
fun register_bnf_raw key bnf =
78095
bc42c074e58f tuned signature: more position information;
wenzelm
parents: 76383
diff changeset
  2084
  Local_Theory.declaration {syntax = false, pervasive = true, pos = \<^here>}
58177
166131276380 introduced local interpretation mechanism for BNFs, to solve issues with datatypes in locales
blanchet
parents: 58175
diff changeset
  2085
    (fn phi => Data.map (Symtab.update (key, morph_bnf phi bnf)));
166131276380 introduced local interpretation mechanism for BNFs, to solve issues with datatypes in locales
blanchet
parents: 58175
diff changeset
  2086
58189
9d714be4f028 added 'plugins' option to control which hooks are enabled
blanchet
parents: 58188
diff changeset
  2087
fun register_bnf plugins key bnf =
9d714be4f028 added 'plugins' option to control which hooks are enabled
blanchet
parents: 58188
diff changeset
  2088
  register_bnf_raw key bnf #> interpret_bnf plugins bnf;
49434
433dc7e028c8 separated registration of BNFs from bnf_def (BNFs are now stored only for bnf_def and (co)data commands)
traytel
parents: 49430
diff changeset
  2089
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  2090
fun bnf_def const_policy fact_policy internal qualify tacs wit_tac Ds map_b rel_b pred_b set_bs
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  2091
    raw_csts =
54189
c0186a0d8cb3 define a trivial nonemptiness witness if none is provided
traytel
parents: 54158
diff changeset
  2092
  (fn (_, goals, (triv_tac_opt, wit_goalss), after_qed, lthy, one_step_defs) =>
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2093
  let
55197
5a54ed681ba2 less hermetic tactics
traytel
parents: 55061
diff changeset
  2094
    fun mk_wits_tac ctxt set_maps =
5a54ed681ba2 less hermetic tactics
traytel
parents: 55061
diff changeset
  2095
      TRYALL Goal.conjunction_tac THEN
54189
c0186a0d8cb3 define a trivial nonemptiness witness if none is provided
traytel
parents: 54158
diff changeset
  2096
      (case triv_tac_opt of
55197
5a54ed681ba2 less hermetic tactics
traytel
parents: 55061
diff changeset
  2097
        SOME tac => tac ctxt set_maps
5a54ed681ba2 less hermetic tactics
traytel
parents: 55061
diff changeset
  2098
      | NONE => unfold_thms_tac ctxt one_step_defs THEN wit_tac ctxt);
54189
c0186a0d8cb3 define a trivial nonemptiness witness if none is provided
traytel
parents: 54158
diff changeset
  2099
    val wit_goals = map Logic.mk_conjunction_balanced wit_goalss;
c0186a0d8cb3 define a trivial nonemptiness witness if none is provided
traytel
parents: 54158
diff changeset
  2100
    fun mk_wit_thms set_maps =
55197
5a54ed681ba2 less hermetic tactics
traytel
parents: 55061
diff changeset
  2101
      Goal.prove_sorry lthy [] [] (Logic.mk_conjunction_balanced wit_goals)
5a54ed681ba2 less hermetic tactics
traytel
parents: 55061
diff changeset
  2102
        (fn {context = ctxt, prems = _} => mk_wits_tac ctxt set_maps)
70494
41108e3e9ca5 formal position for PThm nodes;
wenzelm
parents: 69593
diff changeset
  2103
        |> Thm.close_derivation \<^here>
54189
c0186a0d8cb3 define a trivial nonemptiness witness if none is provided
traytel
parents: 54158
diff changeset
  2104
        |> Conjunction.elim_balanced (length wit_goals)
c0186a0d8cb3 define a trivial nonemptiness witness if none is provided
traytel
parents: 54158
diff changeset
  2105
        |> map2 (Conjunction.elim_balanced o length) wit_goalss
61116
6189d179c2b5 close derivation *before* splitting conjuncts, like Goal.prove_common (see also 757cad5a3fe9) -- potential improvement of performance;
wenzelm
parents: 61101
diff changeset
  2106
        |> (map o map) (Thm.forall_elim_vars 0);
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2107
  in
70494
41108e3e9ca5 formal position for PThm nodes;
wenzelm
parents: 69593
diff changeset
  2108
    map2 (Thm.close_derivation \<^here> oo Goal.prove_sorry lthy [] [])
55197
5a54ed681ba2 less hermetic tactics
traytel
parents: 55061
diff changeset
  2109
      goals (map (fn tac => fn {context = ctxt, prems = _} =>
5a54ed681ba2 less hermetic tactics
traytel
parents: 55061
diff changeset
  2110
        unfold_thms_tac ctxt one_step_defs THEN tac ctxt) tacs)
54189
c0186a0d8cb3 define a trivial nonemptiness witness if none is provided
traytel
parents: 54158
diff changeset
  2111
    |> (fn thms => after_qed mk_wit_thms (map single thms) lthy)
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  2112
  end) o prepare_def const_policy fact_policy internal qualify (K I) (K I) Ds map_b rel_b pred_b
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  2113
    set_bs raw_csts;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2114
58659
6c9821c32dd5 Local_Interpretation is superseded by Plugin with formal Plugin_Name management, avoiding undeclared strings;
wenzelm
parents: 58634
diff changeset
  2115
fun bnf_cmd (raw_csts, raw_plugins) =
58189
9d714be4f028 added 'plugins' option to control which hooks are enabled
blanchet
parents: 58188
diff changeset
  2116
  (fn (key, goals, (triv_tac_opt, wit_goalss), after_qed, lthy, defs) =>
54189
c0186a0d8cb3 define a trivial nonemptiness witness if none is provided
traytel
parents: 54158
diff changeset
  2117
  let
58659
6c9821c32dd5 Local_Interpretation is superseded by Plugin with formal Plugin_Name management, avoiding undeclared strings;
wenzelm
parents: 58634
diff changeset
  2118
    val plugins = raw_plugins lthy;
54189
c0186a0d8cb3 define a trivial nonemptiness witness if none is provided
traytel
parents: 54158
diff changeset
  2119
    val wit_goals = map Logic.mk_conjunction_balanced wit_goalss;
c0186a0d8cb3 define a trivial nonemptiness witness if none is provided
traytel
parents: 54158
diff changeset
  2120
    fun mk_triv_wit_thms tac set_maps =
c0186a0d8cb3 define a trivial nonemptiness witness if none is provided
traytel
parents: 54158
diff changeset
  2121
      Goal.prove_sorry lthy [] [] (Logic.mk_conjunction_balanced wit_goals)
55197
5a54ed681ba2 less hermetic tactics
traytel
parents: 55061
diff changeset
  2122
        (fn {context = ctxt, prems = _} => TRYALL Goal.conjunction_tac THEN tac ctxt set_maps)
70494
41108e3e9ca5 formal position for PThm nodes;
wenzelm
parents: 69593
diff changeset
  2123
        |> Thm.close_derivation \<^here>
54189
c0186a0d8cb3 define a trivial nonemptiness witness if none is provided
traytel
parents: 54158
diff changeset
  2124
        |> Conjunction.elim_balanced (length wit_goals)
c0186a0d8cb3 define a trivial nonemptiness witness if none is provided
traytel
parents: 54158
diff changeset
  2125
        |> map2 (Conjunction.elim_balanced o length) wit_goalss
61116
6189d179c2b5 close derivation *before* splitting conjuncts, like Goal.prove_common (see also 757cad5a3fe9) -- potential improvement of performance;
wenzelm
parents: 61101
diff changeset
  2126
        |> (map o map) (Thm.forall_elim_vars 0);
57301
7b997028aaac generate 'rel_coinduct0' theorem for codatatypes
desharna
parents: 56954
diff changeset
  2127
    val (mk_wit_thms, nontriv_wit_goals) =
54189
c0186a0d8cb3 define a trivial nonemptiness witness if none is provided
traytel
parents: 54158
diff changeset
  2128
      (case triv_tac_opt of
c0186a0d8cb3 define a trivial nonemptiness witness if none is provided
traytel
parents: 54158
diff changeset
  2129
        NONE => (fn _ => [], map (map (rpair [])) wit_goalss)
c0186a0d8cb3 define a trivial nonemptiness witness if none is provided
traytel
parents: 54158
diff changeset
  2130
      | SOME tac => (mk_triv_wit_thms tac, []));
c0186a0d8cb3 define a trivial nonemptiness witness if none is provided
traytel
parents: 54158
diff changeset
  2131
  in
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  2132
    lthy
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  2133
    |> Proof.theorem NONE (uncurry (register_bnf plugins key) oo after_qed mk_wit_thms)
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  2134
      (map (single o rpair []) goals @ nontriv_wit_goals)
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  2135
    |> Proof.unfolding ([[(@{thm OO_Grp_alt} :: @{thm mem_Collect_eq} :: defs, [])]])
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  2136
    |> Proof.refine_singleton (Method.Basic (fn ctxt =>
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  2137
      Method.SIMPLE_METHOD (TRYALL (rtac ctxt refl))))
58189
9d714be4f028 added 'plugins' option to control which hooks are enabled
blanchet
parents: 58188
diff changeset
  2138
  end) o prepare_def Do_Inline (user_policy Note_Some) false I Syntax.read_typ Syntax.read_term
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62093
diff changeset
  2139
    NONE Binding.empty Binding.empty Binding.empty [] raw_csts;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2140
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2141
fun print_bnfs ctxt =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2142
  let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2143
    fun pretty_set sets i = Pretty.block
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2144
      [Pretty.str (mk_setN (i + 1) ^ ":"), Pretty.brk 1,
59663
fb544855e3b1 split into subgoals
blanchet
parents: 59621
diff changeset
  2145
        Pretty.quote (Syntax.pretty_term ctxt (nth sets i))];
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2146
56903
69b6369a98fa generate 'map_ident' theorem for BNFs
desharna
parents: 56766
diff changeset
  2147
    fun pretty_bnf (key, BNF {T, map, sets, bd, live, lives, dead, deads, ...}) =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2148
      Pretty.big_list
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2149
        (Pretty.string_of (Pretty.block [Pretty.str key, Pretty.str ":", Pretty.brk 1,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2150
          Pretty.quote (Syntax.pretty_typ ctxt T)]))
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2151
        ([Pretty.block [Pretty.str "live:", Pretty.brk 1, Pretty.str (string_of_int live),
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2152
            Pretty.brk 3, Pretty.list "[" "]" (List.map (Syntax.pretty_typ ctxt) lives)],
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2153
          Pretty.block [Pretty.str "dead:", Pretty.brk 1, Pretty.str (string_of_int dead),
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2154
            Pretty.brk 3, Pretty.list "[" "]" (List.map (Syntax.pretty_typ ctxt) deads)],
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2155
          Pretty.block [Pretty.str (mapN ^ ":"), Pretty.brk 1,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2156
            Pretty.quote (Syntax.pretty_term ctxt map)]] @
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2157
          List.map (pretty_set sets) (0 upto length sets - 1) @
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2158
          [Pretty.block [Pretty.str (bdN ^ ":"), Pretty.brk 1,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2159
            Pretty.quote (Syntax.pretty_term ctxt bd)]]);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2160
  in
59822
f014a2dc0ac8 sort BNFs in output
blanchet
parents: 59726
diff changeset
  2161
    Pretty.big_list "Registered bounded natural functors:"
60924
610794dff23c tuned signature, in accordance to sortBy in Scala;
wenzelm
parents: 60801
diff changeset
  2162
      (map pretty_bnf (sort_by fst (Symtab.dest (Data.get (Context.Proof ctxt)))))
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2163
    |> Pretty.writeln
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2164
  end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2165
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2166
val _ =
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67222
diff changeset
  2167
  Outer_Syntax.command \<^command_keyword>\<open>print_bnfs\<close>
53289
5e0623448bdb renamed BNF axiom
blanchet
parents: 53288
diff changeset
  2168
    "print all bounded natural functors"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2169
    (Scan.succeed (Toplevel.keep (print_bnfs o Toplevel.context_of)));
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2170
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2171
val _ =
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67222
diff changeset
  2172
  Outer_Syntax.local_theory_to_proof \<^command_keyword>\<open>bnf\<close>
53289
5e0623448bdb renamed BNF axiom
blanchet
parents: 53288
diff changeset
  2173
    "register a type as a bounded natural functor"
54421
632be352a5a3 more explicit syntax for defining a bnf
traytel
parents: 54285
diff changeset
  2174
    (parse_opt_binding_colon -- Parse.typ --|
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67222
diff changeset
  2175
       (Parse.reserved "map" -- \<^keyword>\<open>:\<close>) -- Parse.term --
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67222
diff changeset
  2176
       Scan.optional ((Parse.reserved "sets" -- \<^keyword>\<open>:\<close>) |--
58189
9d714be4f028 added 'plugins' option to control which hooks are enabled
blanchet
parents: 58188
diff changeset
  2177
         Scan.repeat1 (Scan.unless (Parse.reserved "bd") Parse.term)) [] --|
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67222
diff changeset
  2178
       (Parse.reserved "bd" -- \<^keyword>\<open>:\<close>) -- Parse.term --
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67222
diff changeset
  2179
       Scan.optional ((Parse.reserved "wits" -- \<^keyword>\<open>:\<close>) |--
58432
121d5e3319ee improved 'bnf' parser
blanchet
parents: 58352
diff changeset
  2180
         Scan.repeat1 (Scan.unless (Parse.reserved "rel" ||
121d5e3319ee improved 'bnf' parser
blanchet
parents: 58352
diff changeset
  2181
           Parse.reserved "plugins") Parse.term)) [] --
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67222
diff changeset
  2182
       Scan.option ((Parse.reserved "rel" -- \<^keyword>\<open>:\<close>) |-- Parse.term) --
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67222
diff changeset
  2183
       Scan.option ((Parse.reserved "pred" -- \<^keyword>\<open>:\<close>) |-- Parse.term) --
58660
8d4aebb9e327 clarified load order;
wenzelm
parents: 58659
diff changeset
  2184
       Scan.optional Plugin_Name.parse_filter (K Plugin_Name.default_filter)
51836
4d6dcd51dd52 renamed "bnf_def" keyword to "bnf" (since it's not a definition, but rather a registration)
blanchet
parents: 51823
diff changeset
  2185
       >> bnf_cmd);
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2186
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  2187
end;