src/HOL/BNF/Tools/bnf_def.ML
author traytel
Thu, 12 Sep 2013 11:23:49 +0200
changeset 53561 92bcac4f9ac9
parent 53290 b6c3be868217
child 54045 369a4a14583a
permissions -rw-r--r--
simplified derivation of in_rel
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
49509
163914705f8d renamed top-level theory from "Codatatype" to "BNF"
blanchet
parents: 49507
diff changeset
     1
(*  Title:      HOL/BNF/Tools/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
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     4
    Copyright   2012
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     5
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     6
Definition of bounded natural functors.
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
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     9
signature BNF_DEF =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    10
sig
51837
087498724486 lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
blanchet
parents: 51836
diff changeset
    11
  type bnf
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    12
  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
    13
51837
087498724486 lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
blanchet
parents: 51836
diff changeset
    14
  val morph_bnf: morphism -> bnf -> bnf
087498724486 lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
blanchet
parents: 51836
diff changeset
    15
  val eq_bnf: bnf * bnf -> bool
087498724486 lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
blanchet
parents: 51836
diff changeset
    16
  val bnf_of: Proof.context -> string -> bnf option
087498724486 lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
blanchet
parents: 51836
diff changeset
    17
  val register_bnf: string -> (bnf * local_theory) -> (bnf * 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
    18
51837
087498724486 lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
blanchet
parents: 51836
diff changeset
    19
  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
    20
  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
    21
  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
    22
  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
    23
  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
    24
  val deads_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
    25
  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
    26
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    27
  val mapN: string
49507
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49506
diff changeset
    28
  val relN: string
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    29
  val setN: string
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    30
  val mk_setN: int -> string
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    31
51837
087498724486 lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
blanchet
parents: 51836
diff changeset
    32
  val map_of_bnf: bnf -> term
087498724486 lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
blanchet
parents: 51836
diff changeset
    33
  val sets_of_bnf: bnf -> term list
087498724486 lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
blanchet
parents: 51836
diff changeset
    34
  val rel_of_bnf: bnf -> term
49214
2a3cb4c71b87 construct the right iterator theorem in the recursive case
blanchet
parents: 49123
diff changeset
    35
51837
087498724486 lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
blanchet
parents: 51836
diff changeset
    36
  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
    37
  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
    38
  val mk_map_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
    39
  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
    40
  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
    41
  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
    42
51837
087498724486 lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
blanchet
parents: 51836
diff changeset
    43
  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
    44
  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
    45
  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
    46
  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
    47
  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
    48
  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
    49
  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
    50
  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
    51
  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
    52
  val in_rel_of_bnf: bnf -> thm
53287
271b34513bfb renamed BNF axiom
blanchet
parents: 53285
diff changeset
    53
  val map_comp0_of_bnf: bnf -> thm
53288
050d0bc9afa5 renamed BNF fact
blanchet
parents: 53287
diff changeset
    54
  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
    55
  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
    56
  val map_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
    57
  val map_def_of_bnf: bnf -> thm
53270
c8628119d18e renamed BNF axiom
blanchet
parents: 53265
diff changeset
    58
  val map_id0_of_bnf: bnf -> thm
53285
f09645642984 renamed BNF fact
blanchet
parents: 53270
diff changeset
    59
  val map_id_of_bnf: bnf -> thm
52719
480a3479fa47 transfer rule for map (not yet registered as a transfer rule)
traytel
parents: 52635
diff changeset
    60
  val map_transfer_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
    61
  val map_wppull_of_bnf: bnf -> thm
087498724486 lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
blanchet
parents: 51836
diff changeset
    62
  val map_wpull_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 rel_def_of_bnf: bnf -> thm
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
    64
  val rel_Grp_of_bnf: bnf -> thm
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
    65
  val rel_OO_of_bnf: bnf -> thm
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
    66
  val rel_OO_Grp_of_bnf: bnf -> thm
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
    67
  val rel_cong_of_bnf: bnf -> thm
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
    68
  val rel_conversep_of_bnf: bnf -> thm
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
    69
  val rel_mono_of_bnf: bnf -> thm
51916
eac9e9a45bf5 stronger monotonicity property for relators
traytel
parents: 51915
diff changeset
    70
  val rel_mono_strong_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
    71
  val rel_eq_of_bnf: bnf -> thm
087498724486 lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
blanchet
parents: 51836
diff changeset
    72
  val rel_flip_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 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
    74
  val set_defs_of_bnf: bnf -> thm list
53289
5e0623448bdb renamed BNF axiom
blanchet
parents: 53288
diff changeset
    75
  val set_map0_of_bnf: bnf -> thm list
53290
b6c3be868217 renamed BNF fact
blanchet
parents: 53289
diff changeset
    76
  val set_map_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
    77
  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
    78
  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
    79
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    80
  val mk_witness: int list * term -> thm list -> nonemptiness_witness
49103
3caaa80f53a4 generalized signature
traytel
parents: 49021
diff changeset
    81
  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
    82
  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
    83
52635
4f84b730c489 got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
traytel
parents: 51930
diff changeset
    84
  val zip_axioms: 'a -> 'a -> 'a -> 'a list -> 'a -> 'a -> 'a list -> 'a -> 'a -> 'a list
49456
fa8302c8dea1 adapted BNF composition to new relator approach
blanchet
parents: 49455
diff changeset
    85
53031
83cbe188855a moved useful library functions upstream
blanchet
parents: 52957
diff changeset
    86
  val flatten_type_args_of_bnf: bnf -> 'a -> 'a list -> 'a list
83cbe188855a moved useful library functions upstream
blanchet
parents: 52957
diff changeset
    87
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    88
  datatype const_policy = Dont_Inline | Hardly_Inline | Smart_Inline | Do_Inline
49538
c90818b63599 simplified fact policies
blanchet
parents: 49537
diff changeset
    89
  datatype fact_policy = Dont_Note | Note_Some | Note_All
c90818b63599 simplified fact policies
blanchet
parents: 49537
diff changeset
    90
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    91
  val bnf_note_all: bool Config.T
53143
ba80154a1118 configuration option to control timing output for (co)datatypes
traytel
parents: 53126
diff changeset
    92
  val bnf_timing: bool Config.T
49435
483007ddbdc2 bnf_note_all mode for "pre_"-BNFs
traytel
parents: 49434
diff changeset
    93
  val user_policy: fact_policy -> Proof.context -> fact_policy
52720
fdc04f9bf906 separate ML interface to note facts of a bnf
traytel
parents: 52719
diff changeset
    94
  val note_bnf_thms: fact_policy -> (binding -> binding) -> binding -> bnf -> Proof.context ->
fdc04f9bf906 separate ML interface to note facts of a bnf
traytel
parents: 52719
diff changeset
    95
    Proof.context
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    96
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    97
  val print_bnfs: Proof.context -> unit
49018
b2884253b421 renamed ML function for consistency
blanchet
parents: 49016
diff changeset
    98
  val bnf_def: const_policy -> (Proof.context -> fact_policy) -> (binding -> binding) ->
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    99
    ({prems: thm list, context: Proof.context} -> tactic) list ->
51758
55963309557b honor user-specified name for map function
blanchet
parents: 51757
diff changeset
   100
    ({prems: thm list, context: Proof.context} -> tactic) -> typ list option -> binding ->
51767
bbcdd8519253 honor user-specified name for relator + generalize syntax
blanchet
parents: 51766
diff changeset
   101
    binding -> binding list ->
bbcdd8519253 honor user-specified name for relator + generalize syntax
blanchet
parents: 51766
diff changeset
   102
    ((((binding * term) * term list) * term) * term list) * term option ->
51837
087498724486 lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
blanchet
parents: 51836
diff changeset
   103
    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
   104
end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   105
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   106
structure BNF_Def : BNF_DEF =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   107
struct
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   108
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   109
open BNF_Util
49463
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49462
diff changeset
   110
open BNF_Tactics
49284
5f39b7940b49 spin off "bnf_def_tactics.ML"
blanchet
parents: 49279
diff changeset
   111
open BNF_Def_Tactics
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   112
51765
224b6eb2313a added "fundef_cong" attribute to "map_cong"
blanchet
parents: 51762
diff changeset
   113
val fundef_cong_attrs = @{attributes [fundef_cong]};
224b6eb2313a added "fundef_cong" attribute to "map_cong"
blanchet
parents: 51762
diff changeset
   114
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   115
type axioms = {
53270
c8628119d18e renamed BNF axiom
blanchet
parents: 53265
diff changeset
   116
  map_id0: thm,
53287
271b34513bfb renamed BNF axiom
blanchet
parents: 53285
diff changeset
   117
  map_comp0: thm,
51761
4c9f08836d87 renamed "map_cong" axiom to "map_cong0" in preparation for real "map_cong"
blanchet
parents: 51758
diff changeset
   118
  map_cong0: thm,
53289
5e0623448bdb renamed BNF axiom
blanchet
parents: 53288
diff changeset
   119
  set_map0: thm list,
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   120
  bd_card_order: thm,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   121
  bd_cinfinite: thm,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   122
  set_bd: thm list,
49453
ff0e540d8758 add rel as first-class citizen of BNF
blanchet
parents: 49452
diff changeset
   123
  map_wpull: thm,
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
   124
  rel_OO_Grp: thm
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   125
};
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   126
53289
5e0623448bdb renamed BNF axiom
blanchet
parents: 53288
diff changeset
   127
fun mk_axioms' ((((((((id, comp), cong), map), c_o), cinf), set_bd), wpull), rel) =
5e0623448bdb renamed BNF axiom
blanchet
parents: 53288
diff changeset
   128
  {map_id0 = id, map_comp0 = comp, map_cong0 = cong, set_map0 = map, bd_card_order = c_o,
52635
4f84b730c489 got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
traytel
parents: 51930
diff changeset
   129
   bd_cinfinite = cinf, set_bd = set_bd, map_wpull = wpull, rel_OO_Grp = rel};
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   130
51930
52fd62618631 prefer explicitly qualified exceptions, which is particular important for robust handlers;
wenzelm
parents: 51917
diff changeset
   131
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
   132
  | 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
   133
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   134
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
   135
  |> map the_single
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   136
  |> dest_cons
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   137
  ||>> dest_cons
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   138
  ||>> dest_cons
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   139
  ||>> chop n
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   140
  ||>> dest_cons
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   141
  ||>> dest_cons
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   142
  ||>> chop n
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   143
  ||>> dest_cons
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   144
  ||> the_single
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   145
  |> mk_axioms';
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   146
53289
5e0623448bdb renamed BNF axiom
blanchet
parents: 53288
diff changeset
   147
fun zip_axioms mid mcomp mcong smap bdco bdinf sbd wpull rel =
5e0623448bdb renamed BNF axiom
blanchet
parents: 53288
diff changeset
   148
  [mid, mcomp, mcong] @ smap @ [bdco, bdinf] @ sbd @ [wpull, rel];
49460
4dd451a075ce fixed infinite loop with trivial rel_O_Gr + tuning
blanchet
parents: 49459
diff changeset
   149
53289
5e0623448bdb renamed BNF axiom
blanchet
parents: 53288
diff changeset
   150
fun dest_axioms {map_id0, map_comp0, map_cong0, set_map0, bd_card_order, bd_cinfinite, set_bd,
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
   151
  map_wpull, rel_OO_Grp} =
53289
5e0623448bdb renamed BNF axiom
blanchet
parents: 53288
diff changeset
   152
  zip_axioms map_id0 map_comp0 map_cong0 set_map0 bd_card_order bd_cinfinite set_bd map_wpull
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
   153
    rel_OO_Grp;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   154
53289
5e0623448bdb renamed BNF axiom
blanchet
parents: 53288
diff changeset
   155
fun map_axioms f {map_id0, map_comp0, map_cong0, set_map0, bd_card_order, bd_cinfinite, set_bd,
52635
4f84b730c489 got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
traytel
parents: 51930
diff changeset
   156
  map_wpull, rel_OO_Grp} =
53270
c8628119d18e renamed BNF axiom
blanchet
parents: 53265
diff changeset
   157
  {map_id0 = f map_id0,
53287
271b34513bfb renamed BNF axiom
blanchet
parents: 53285
diff changeset
   158
    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
   159
    map_cong0 = f map_cong0,
53289
5e0623448bdb renamed BNF axiom
blanchet
parents: 53288
diff changeset
   160
    set_map0 = map f set_map0,
49463
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49462
diff changeset
   161
    bd_card_order = f bd_card_order,
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49462
diff changeset
   162
    bd_cinfinite = f bd_cinfinite,
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49462
diff changeset
   163
    set_bd = map f set_bd,
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49462
diff changeset
   164
    map_wpull = f map_wpull,
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
   165
    rel_OO_Grp = f rel_OO_Grp};
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   166
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   167
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
   168
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   169
type defs = {
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   170
  map_def: thm,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   171
  set_defs: thm list,
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
   172
  rel_def: thm
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   173
}
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   174
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
   175
fun mk_defs map sets rel = {map_def = map, set_defs = sets, rel_def = rel};
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   176
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
   177
fun map_defs f {map_def, set_defs, rel_def} =
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
   178
  {map_def = f map_def, set_defs = map f set_defs, rel_def = f rel_def};
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   179
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   180
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
   181
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   182
type facts = {
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   183
  bd_Card_order: thm,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   184
  bd_Cinfinite: thm,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   185
  bd_Cnotzero: thm,
51766
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51765
diff changeset
   186
  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
   187
  in_bd: thm lazy,
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   188
  in_cong: thm lazy,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   189
  in_mono: thm lazy,
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
   190
  in_rel: thm lazy,
53288
050d0bc9afa5 renamed BNF fact
blanchet
parents: 53287
diff changeset
   191
  map_comp: thm lazy,
51762
219a3063ed29 derive "map_cong"
blanchet
parents: 51761
diff changeset
   192
  map_cong: thm lazy,
53285
f09645642984 renamed BNF fact
blanchet
parents: 53270
diff changeset
   193
  map_id: thm lazy,
52719
480a3479fa47 transfer rule for map (not yet registered as a transfer rule)
traytel
parents: 52635
diff changeset
   194
  map_transfer: thm lazy,
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   195
  map_wppull: thm lazy,
49591
91b228e26348 generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents: 49585
diff changeset
   196
  rel_eq: thm lazy,
49537
fe1deee434b6 generate "rel_as_srel" and "rel_flip" properties
blanchet
parents: 49536
diff changeset
   197
  rel_flip: thm lazy,
53290
b6c3be868217 renamed BNF fact
blanchet
parents: 53289
diff changeset
   198
  set_map: thm lazy list,
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
   199
  rel_cong: thm lazy,
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
   200
  rel_mono: thm lazy,
51916
eac9e9a45bf5 stronger monotonicity property for relators
traytel
parents: 51915
diff changeset
   201
  rel_mono_strong: thm lazy,
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
   202
  rel_Grp: thm lazy,
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
   203
  rel_conversep: thm lazy,
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
   204
  rel_OO: thm lazy
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   205
};
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   206
52635
4f84b730c489 got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
traytel
parents: 51930
diff changeset
   207
fun mk_facts bd_Card_order bd_Cinfinite bd_Cnotzero collect_set_map in_bd in_cong in_mono in_rel
53290
b6c3be868217 renamed BNF fact
blanchet
parents: 53289
diff changeset
   208
    map_comp map_cong map_id map_transfer map_wppull rel_eq rel_flip set_map rel_cong rel_mono
52719
480a3479fa47 transfer rule for map (not yet registered as a transfer rule)
traytel
parents: 52635
diff changeset
   209
    rel_mono_strong rel_Grp rel_conversep rel_OO = {
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   210
  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
   211
  bd_Cinfinite = bd_Cinfinite,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   212
  bd_Cnotzero = bd_Cnotzero,
51766
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51765
diff changeset
   213
  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
   214
  in_bd = in_bd,
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   215
  in_cong = in_cong,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   216
  in_mono = in_mono,
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
   217
  in_rel = in_rel,
53288
050d0bc9afa5 renamed BNF fact
blanchet
parents: 53287
diff changeset
   218
  map_comp = map_comp,
51762
219a3063ed29 derive "map_cong"
blanchet
parents: 51761
diff changeset
   219
  map_cong = map_cong,
53285
f09645642984 renamed BNF fact
blanchet
parents: 53270
diff changeset
   220
  map_id = map_id,
52719
480a3479fa47 transfer rule for map (not yet registered as a transfer rule)
traytel
parents: 52635
diff changeset
   221
  map_transfer = map_transfer,
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   222
  map_wppull = map_wppull,
49591
91b228e26348 generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents: 49585
diff changeset
   223
  rel_eq = rel_eq,
49537
fe1deee434b6 generate "rel_as_srel" and "rel_flip" properties
blanchet
parents: 49536
diff changeset
   224
  rel_flip = rel_flip,
53290
b6c3be868217 renamed BNF fact
blanchet
parents: 53289
diff changeset
   225
  set_map = set_map,
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
   226
  rel_cong = rel_cong,
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
   227
  rel_mono = rel_mono,
51916
eac9e9a45bf5 stronger monotonicity property for relators
traytel
parents: 51915
diff changeset
   228
  rel_mono_strong = rel_mono_strong,
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
   229
  rel_Grp = rel_Grp,
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
   230
  rel_conversep = rel_conversep,
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
   231
  rel_OO = rel_OO};
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   232
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   233
fun map_facts f {
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   234
  bd_Card_order,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   235
  bd_Cinfinite,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   236
  bd_Cnotzero,
51766
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51765
diff changeset
   237
  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
   238
  in_bd,
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   239
  in_cong,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   240
  in_mono,
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
   241
  in_rel,
53288
050d0bc9afa5 renamed BNF fact
blanchet
parents: 53287
diff changeset
   242
  map_comp,
51762
219a3063ed29 derive "map_cong"
blanchet
parents: 51761
diff changeset
   243
  map_cong,
53285
f09645642984 renamed BNF fact
blanchet
parents: 53270
diff changeset
   244
  map_id,
52719
480a3479fa47 transfer rule for map (not yet registered as a transfer rule)
traytel
parents: 52635
diff changeset
   245
  map_transfer,
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   246
  map_wppull,
49591
91b228e26348 generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents: 49585
diff changeset
   247
  rel_eq,
49537
fe1deee434b6 generate "rel_as_srel" and "rel_flip" properties
blanchet
parents: 49536
diff changeset
   248
  rel_flip,
53290
b6c3be868217 renamed BNF fact
blanchet
parents: 53289
diff changeset
   249
  set_map,
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
   250
  rel_cong,
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
   251
  rel_mono,
51916
eac9e9a45bf5 stronger monotonicity property for relators
traytel
parents: 51915
diff changeset
   252
  rel_mono_strong,
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
   253
  rel_Grp,
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
   254
  rel_conversep,
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
   255
  rel_OO} =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   256
  {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
   257
    bd_Cinfinite = f bd_Cinfinite,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   258
    bd_Cnotzero = f bd_Cnotzero,
51766
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51765
diff changeset
   259
    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
   260
    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
   261
    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
   262
    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
   263
    in_rel = Lazy.map f in_rel,
53288
050d0bc9afa5 renamed BNF fact
blanchet
parents: 53287
diff changeset
   264
    map_comp = Lazy.map f map_comp,
51762
219a3063ed29 derive "map_cong"
blanchet
parents: 51761
diff changeset
   265
    map_cong = Lazy.map f map_cong,
53285
f09645642984 renamed BNF fact
blanchet
parents: 53270
diff changeset
   266
    map_id = Lazy.map f map_id,
52719
480a3479fa47 transfer rule for map (not yet registered as a transfer rule)
traytel
parents: 52635
diff changeset
   267
    map_transfer = Lazy.map f map_transfer,
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   268
    map_wppull = Lazy.map f map_wppull,
49591
91b228e26348 generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents: 49585
diff changeset
   269
    rel_eq = Lazy.map f rel_eq,
49537
fe1deee434b6 generate "rel_as_srel" and "rel_flip" properties
blanchet
parents: 49536
diff changeset
   270
    rel_flip = Lazy.map f rel_flip,
53290
b6c3be868217 renamed BNF fact
blanchet
parents: 53289
diff changeset
   271
    set_map = map (Lazy.map f) set_map,
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
   272
    rel_cong = Lazy.map f rel_cong,
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
   273
    rel_mono = Lazy.map f rel_mono,
51916
eac9e9a45bf5 stronger monotonicity property for relators
traytel
parents: 51915
diff changeset
   274
    rel_mono_strong = Lazy.map f rel_mono_strong,
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
   275
    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
   276
    rel_conversep = Lazy.map f rel_conversep,
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
   277
    rel_OO = Lazy.map f rel_OO};
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   278
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   279
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
   280
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   281
type nonemptiness_witness = {
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   282
  I: int list,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   283
  wit: term,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   284
  prop: thm list
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   285
};
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   286
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   287
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
   288
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
   289
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
   290
51837
087498724486 lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
blanchet
parents: 51836
diff changeset
   291
datatype bnf = BNF of {
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   292
  name: binding,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   293
  T: typ,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   294
  live: int,
53261
3c26a7042d8e removed outdated comments
panny
parents: 53143
diff changeset
   295
  lives: typ list, (*source type variables of map*)
3c26a7042d8e removed outdated comments
panny
parents: 53143
diff changeset
   296
  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
   297
  dead: int,
53261
3c26a7042d8e removed outdated comments
panny
parents: 53143
diff changeset
   298
  deads: typ list,
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   299
  map: term,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   300
  sets: term list,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   301
  bd: term,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   302
  axioms: axioms,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   303
  defs: defs,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   304
  facts: facts,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   305
  nwits: int,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   306
  wits: nonemptiness_witness list,
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
   307
  rel: term
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   308
};
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   309
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   310
(* getters *)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   311
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   312
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
   313
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
   314
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
   315
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
   316
  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
   317
  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
   318
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
   319
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
   320
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
   321
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
   322
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
   323
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
   324
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
   325
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
   326
53031
83cbe188855a moved useful library functions upstream
blanchet
parents: 52957
diff changeset
   327
fun flatten_type_args_of_bnf bnf dead_x xs =
83cbe188855a moved useful library functions upstream
blanchet
parents: 52957
diff changeset
   328
  let
83cbe188855a moved useful library functions upstream
blanchet
parents: 52957
diff changeset
   329
    val Type (_, Ts) = T_of_bnf bnf;
83cbe188855a moved useful library functions upstream
blanchet
parents: 52957
diff changeset
   330
    val lives = lives_of_bnf bnf;
83cbe188855a moved useful library functions upstream
blanchet
parents: 52957
diff changeset
   331
    val deads = deads_of_bnf bnf;
83cbe188855a moved useful library functions upstream
blanchet
parents: 52957
diff changeset
   332
  in
53040
e6edd7abc4ce renamed function
blanchet
parents: 53039
diff changeset
   333
    permute_like (op =) (deads @ lives) Ts (replicate (length deads) dead_x @ xs)
53031
83cbe188855a moved useful library functions upstream
blanchet
parents: 52957
diff changeset
   334
  end;
83cbe188855a moved useful library functions upstream
blanchet
parents: 52957
diff changeset
   335
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   336
(*terms*)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   337
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
   338
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
   339
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
   340
  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
   341
  in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   342
    Term.subst_atomic_types
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   343
      ((#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
   344
  end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   345
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
   346
  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
   347
  in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   348
    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
   349
      ((#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
   350
  end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   351
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
   352
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
   353
  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
   354
  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
   355
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
   356
  let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   357
    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
   358
    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
   359
  in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   360
    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
   361
      ((#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
   362
  end;
49507
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49506
diff changeset
   363
val rel_of_bnf = #rel o rep_bnf;
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49506
diff changeset
   364
fun mk_rel_of_bnf Ds Ts Us bnf =
49462
blanchet
parents: 49461
diff changeset
   365
  let val bnf_rep = rep_bnf bnf;
blanchet
parents: 49461
diff changeset
   366
  in
blanchet
parents: 49461
diff changeset
   367
    Term.subst_atomic_types
49507
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49506
diff changeset
   368
      ((#deads bnf_rep ~~ Ds) @ (#lives bnf_rep ~~ Ts) @ (#lives' bnf_rep ~~ Us)) (#rel bnf_rep)
49462
blanchet
parents: 49461
diff changeset
   369
  end;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   370
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   371
(*thms*)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   372
val bd_card_order_of_bnf = #bd_card_order o #axioms o rep_bnf;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   373
val bd_cinfinite_of_bnf = #bd_cinfinite o #axioms o rep_bnf;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   374
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
   375
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
   376
val bd_Cnotzero_of_bnf = #bd_Cnotzero o #facts o rep_bnf;
51766
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51765
diff changeset
   377
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
   378
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
   379
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
   380
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
   381
val in_rel_of_bnf = Lazy.force o #in_rel 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
   382
val map_def_of_bnf = #map_def o #defs o rep_bnf;
53270
c8628119d18e renamed BNF axiom
blanchet
parents: 53265
diff changeset
   383
val map_id0_of_bnf = #map_id0 o #axioms o rep_bnf;
53285
f09645642984 renamed BNF fact
blanchet
parents: 53270
diff changeset
   384
val map_id_of_bnf = Lazy.force o #map_id o #facts o rep_bnf;
53287
271b34513bfb renamed BNF axiom
blanchet
parents: 53285
diff changeset
   385
val map_comp0_of_bnf = #map_comp0 o #axioms o rep_bnf;
53288
050d0bc9afa5 renamed BNF fact
blanchet
parents: 53287
diff changeset
   386
val map_comp_of_bnf = Lazy.force o #map_comp o #facts o rep_bnf;
51761
4c9f08836d87 renamed "map_cong" axiom to "map_cong0" in preparation for real "map_cong"
blanchet
parents: 51758
diff changeset
   387
val map_cong0_of_bnf = #map_cong0 o #axioms o rep_bnf;
51762
219a3063ed29 derive "map_cong"
blanchet
parents: 51761
diff changeset
   388
val map_cong_of_bnf = Lazy.force o #map_cong o #facts o rep_bnf;
52731
dacd47a0633f transfer rule for {c,d}tor_{,un}fold
traytel
parents: 52725
diff changeset
   389
val map_transfer_of_bnf = Lazy.force o #map_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
   390
val map_wppull_of_bnf = Lazy.force o #map_wppull o #facts o rep_bnf;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   391
val map_wpull_of_bnf = #map_wpull o #axioms o rep_bnf;
49507
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49506
diff changeset
   392
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
   393
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
   394
val rel_flip_of_bnf = Lazy.force o #rel_flip 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
   395
val set_bd_of_bnf = #set_bd o #axioms o rep_bnf;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   396
val set_defs_of_bnf = #set_defs o #defs o rep_bnf;
53289
5e0623448bdb renamed BNF axiom
blanchet
parents: 53288
diff changeset
   397
val set_map0_of_bnf = #set_map0 o #axioms o rep_bnf;
53290
b6c3be868217 renamed BNF fact
blanchet
parents: 53289
diff changeset
   398
val set_map_of_bnf = map Lazy.force o #set_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
   399
val rel_cong_of_bnf = Lazy.force o #rel_cong o #facts o rep_bnf;
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
   400
val rel_mono_of_bnf = Lazy.force o #rel_mono o #facts o rep_bnf;
51916
eac9e9a45bf5 stronger monotonicity property for relators
traytel
parents: 51915
diff changeset
   401
val rel_mono_strong_of_bnf = Lazy.force o #rel_mono_strong 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
   402
val rel_Grp_of_bnf = Lazy.force o #rel_Grp o #facts o rep_bnf;
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
   403
val rel_conversep_of_bnf = Lazy.force o #rel_conversep o #facts o rep_bnf;
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
   404
val rel_OO_of_bnf = Lazy.force o #rel_OO o #facts o rep_bnf;
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
   405
val rel_OO_Grp_of_bnf = #rel_OO_Grp o #axioms o rep_bnf;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   406
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
   407
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
   408
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
   409
fun mk_bnf name T live lives lives' dead deads map sets bd axioms defs facts wits rel =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   410
  BNF {name = name, T = T,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   411
       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
   412
       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
   413
       axioms = axioms, defs = defs, facts = facts,
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
   414
       nwits = length wits, wits = wits, rel = rel};
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   415
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   416
fun morph_bnf phi (BNF {name = name, T = T, live = live, lives = lives, lives' = lives',
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   417
  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
   418
  axioms = axioms, defs = defs, facts = facts,
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
   419
  nwits = nwits, wits = wits, rel = rel}) =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   420
  BNF {name = Morphism.binding phi name, T = Morphism.typ phi T,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   421
    live = live, lives = List.map (Morphism.typ phi) lives,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   422
    lives' = List.map (Morphism.typ phi) lives',
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   423
    dead = dead, deads = List.map (Morphism.typ phi) deads,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   424
    map = Morphism.term phi map, sets = List.map (Morphism.term phi) sets,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   425
    bd = Morphism.term phi bd,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   426
    axioms = morph_axioms phi axioms,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   427
    defs = morph_defs phi defs,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   428
    facts = morph_facts phi facts,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   429
    nwits = nwits,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   430
    wits = List.map (morph_witness phi) wits,
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
   431
    rel = Morphism.term phi rel};
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   432
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   433
fun eq_bnf (BNF {T = T1, live = live1, dead = dead1, ...},
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   434
  BNF {T = T2, live = live2, dead = dead2, ...}) =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   435
  Type.could_unify (T1, T2) andalso live1 = live2 andalso dead1 = dead2;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   436
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   437
structure Data = Generic_Data
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   438
(
51837
087498724486 lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
blanchet
parents: 51836
diff changeset
   439
  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
   440
  val empty = Symtab.empty;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   441
  val extend = I;
49462
blanchet
parents: 49461
diff changeset
   442
  val merge = Symtab.merge eq_bnf;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   443
);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   444
53126
f4d2c64c7aa8 transfer stored fp_sugar theorems into the "current" theory when retrieving an fp_sugar (avoids non-trivial merges)
traytel
parents: 53040
diff changeset
   445
fun bnf_of ctxt =
f4d2c64c7aa8 transfer stored fp_sugar theorems into the "current" theory when retrieving an fp_sugar (avoids non-trivial merges)
traytel
parents: 53040
diff changeset
   446
  Symtab.lookup (Data.get (Context.Proof ctxt))
f4d2c64c7aa8 transfer stored fp_sugar theorems into the "current" theory when retrieving an fp_sugar (avoids non-trivial merges)
traytel
parents: 53040
diff changeset
   447
  #> Option.map (morph_bnf (Morphism.thm_morphism (Thm.transfer (Proof_Context.theory_of ctxt))));
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   448
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   449
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   450
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   451
(* Utilities *)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   452
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   453
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
   454
  let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   455
    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
   456
    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
   457
    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
   458
  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
   459
49507
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49506
diff changeset
   460
fun normalize_rel ctxt instTs instA instB rel =
49462
blanchet
parents: 49461
diff changeset
   461
  let
blanchet
parents: 49461
diff changeset
   462
    val thy = Proof_Context.theory_of ctxt;
blanchet
parents: 49461
diff changeset
   463
    val tyenv =
49507
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49506
diff changeset
   464
      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
   465
        Vartab.empty;
49507
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49506
diff changeset
   466
  in Envir.subst_term (tyenv, Vartab.empty) rel end
49453
ff0e540d8758 add rel as first-class citizen of BNF
blanchet
parents: 49452
diff changeset
   467
  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
   468
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   469
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
   470
  let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   471
    fun strip_param (Ts, T as Type (@{type_name fun}, [T1, T2])) =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   472
        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
   473
      | strip_param x = x;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   474
    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
   475
    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
   476
    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
   477
  in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   478
    (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
   479
  end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   480
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   481
fun minimize_wits wits =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   482
 let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   483
   fun minimize done [] = done
49103
3caaa80f53a4 generalized signature
traytel
parents: 49021
diff changeset
   484
     | 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
   485
       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
   486
       then minimize done todo
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   487
       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
   488
 in minimize [] wits end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   489
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   490
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   491
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   492
(* Names *)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   493
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   494
val mapN = "map";
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   495
val setN = "set";
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   496
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
   497
val bdN = "bd";
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   498
val witN = "wit";
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   499
fun mk_witN i = witN ^ nonzero_string_of_int i;
49507
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49506
diff changeset
   500
val relN = "rel";
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   501
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   502
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
   503
val bd_cinfiniteN = "bd_cinfinite";
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   504
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
   505
val bd_CinfiniteN = "bd_Cinfinite";
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   506
val bd_CnotzeroN = "bd_Cnotzero";
51766
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51765
diff changeset
   507
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
   508
val in_bdN = "in_bd";
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   509
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
   510
val in_relN = "in_rel";
53270
c8628119d18e renamed BNF axiom
blanchet
parents: 53265
diff changeset
   511
val map_id0N = "map_id0";
53285
f09645642984 renamed BNF fact
blanchet
parents: 53270
diff changeset
   512
val map_idN = "map_id";
53287
271b34513bfb renamed BNF axiom
blanchet
parents: 53285
diff changeset
   513
val map_comp0N = "map_comp0";
53288
050d0bc9afa5 renamed BNF fact
blanchet
parents: 53287
diff changeset
   514
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
   515
val map_cong0N = "map_cong0";
51762
219a3063ed29 derive "map_cong"
blanchet
parents: 51761
diff changeset
   516
val map_congN = "map_cong";
52719
480a3479fa47 transfer rule for map (not yet registered as a transfer rule)
traytel
parents: 52635
diff changeset
   517
val map_transferN = "map_transfer";
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   518
val map_wpullN = "map_wpull";
49591
91b228e26348 generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents: 49585
diff changeset
   519
val rel_eqN = "rel_eq";
49537
fe1deee434b6 generate "rel_as_srel" and "rel_flip" properties
blanchet
parents: 49536
diff changeset
   520
val rel_flipN = "rel_flip";
53289
5e0623448bdb renamed BNF axiom
blanchet
parents: 53288
diff changeset
   521
val set_map0N = "set_map0";
53290
b6c3be868217 renamed BNF fact
blanchet
parents: 53289
diff changeset
   522
val set_mapN = "set_map";
49537
fe1deee434b6 generate "rel_as_srel" and "rel_flip" properties
blanchet
parents: 49536
diff changeset
   523
val set_bdN = "set_bd";
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
   524
val rel_GrpN = "rel_Grp";
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
   525
val rel_conversepN = "rel_conversep";
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
   526
val rel_monoN = "rel_mono"
51916
eac9e9a45bf5 stronger monotonicity property for relators
traytel
parents: 51915
diff changeset
   527
val rel_mono_strongN = "rel_mono_strong"
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
   528
val rel_OON = "rel_compp";
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
   529
val rel_OO_GrpN = "rel_compp_Grp";
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   530
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   531
datatype const_policy = Dont_Inline | Hardly_Inline | Smart_Inline | Do_Inline;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   532
49538
c90818b63599 simplified fact policies
blanchet
parents: 49537
diff changeset
   533
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
   534
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   535
val bnf_note_all = Attrib.setup_config_bool @{binding bnf_note_all} (K false);
53143
ba80154a1118 configuration option to control timing output for (co)datatypes
traytel
parents: 53126
diff changeset
   536
val bnf_timing = Attrib.setup_config_bool @{binding bnf_timing} (K false);
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   537
49538
c90818b63599 simplified fact policies
blanchet
parents: 49537
diff changeset
   538
fun user_policy policy ctxt = if Config.get ctxt bnf_note_all 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
   539
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   540
val smart_max_inline_size = 25; (*FUDGE*)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   541
53265
cc9a2976f836 qualify BNF constants properly
blanchet
parents: 53261
diff changeset
   542
fun note_bnf_thms fact_policy qualify bnf_b bnf =
52720
fdc04f9bf906 separate ML interface to note facts of a bnf
traytel
parents: 52719
diff changeset
   543
  let
fdc04f9bf906 separate ML interface to note facts of a bnf
traytel
parents: 52719
diff changeset
   544
    val axioms = axioms_of_bnf bnf;
fdc04f9bf906 separate ML interface to note facts of a bnf
traytel
parents: 52719
diff changeset
   545
    val facts = facts_of_bnf bnf;
fdc04f9bf906 separate ML interface to note facts of a bnf
traytel
parents: 52719
diff changeset
   546
    val wits = wits_of_bnf bnf;
fdc04f9bf906 separate ML interface to note facts of a bnf
traytel
parents: 52719
diff changeset
   547
  in
fdc04f9bf906 separate ML interface to note facts of a bnf
traytel
parents: 52719
diff changeset
   548
    (if fact_policy = Note_All then
fdc04f9bf906 separate ML interface to note facts of a bnf
traytel
parents: 52719
diff changeset
   549
      let
fdc04f9bf906 separate ML interface to note facts of a bnf
traytel
parents: 52719
diff changeset
   550
        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
   551
        val notes =
fdc04f9bf906 separate ML interface to note facts of a bnf
traytel
parents: 52719
diff changeset
   552
          [(bd_card_orderN, [#bd_card_order axioms]),
fdc04f9bf906 separate ML interface to note facts of a bnf
traytel
parents: 52719
diff changeset
   553
            (bd_cinfiniteN, [#bd_cinfinite axioms]),
fdc04f9bf906 separate ML interface to note facts of a bnf
traytel
parents: 52719
diff changeset
   554
            (bd_Card_orderN, [#bd_Card_order facts]),
fdc04f9bf906 separate ML interface to note facts of a bnf
traytel
parents: 52719
diff changeset
   555
            (bd_CinfiniteN, [#bd_Cinfinite facts]),
fdc04f9bf906 separate ML interface to note facts of a bnf
traytel
parents: 52719
diff changeset
   556
            (bd_CnotzeroN, [#bd_Cnotzero facts]),
fdc04f9bf906 separate ML interface to note facts of a bnf
traytel
parents: 52719
diff changeset
   557
            (collect_set_mapN, [Lazy.force (#collect_set_map facts)]),
fdc04f9bf906 separate ML interface to note facts of a bnf
traytel
parents: 52719
diff changeset
   558
            (in_bdN, [Lazy.force (#in_bd facts)]),
fdc04f9bf906 separate ML interface to note facts of a bnf
traytel
parents: 52719
diff changeset
   559
            (in_monoN, [Lazy.force (#in_mono facts)]),
fdc04f9bf906 separate ML interface to note facts of a bnf
traytel
parents: 52719
diff changeset
   560
            (in_relN, [Lazy.force (#in_rel facts)]),
53287
271b34513bfb renamed BNF axiom
blanchet
parents: 53285
diff changeset
   561
            (map_comp0N, [#map_comp0 axioms]),
53270
c8628119d18e renamed BNF axiom
blanchet
parents: 53265
diff changeset
   562
            (map_id0N, [#map_id0 axioms]),
52720
fdc04f9bf906 separate ML interface to note facts of a bnf
traytel
parents: 52719
diff changeset
   563
            (map_transferN, [Lazy.force (#map_transfer facts)]),
fdc04f9bf906 separate ML interface to note facts of a bnf
traytel
parents: 52719
diff changeset
   564
            (map_wpullN, [#map_wpull axioms]),
fdc04f9bf906 separate ML interface to note facts of a bnf
traytel
parents: 52719
diff changeset
   565
            (rel_mono_strongN, [Lazy.force (#rel_mono_strong facts)]),
53289
5e0623448bdb renamed BNF axiom
blanchet
parents: 53288
diff changeset
   566
            (set_map0N, #set_map0 axioms),
52720
fdc04f9bf906 separate ML interface to note facts of a bnf
traytel
parents: 52719
diff changeset
   567
            (set_bdN, #set_bd axioms)] @
fdc04f9bf906 separate ML interface to note facts of a bnf
traytel
parents: 52719
diff changeset
   568
            (witNs ~~ wit_thmss_of_bnf bnf)
fdc04f9bf906 separate ML interface to note facts of a bnf
traytel
parents: 52719
diff changeset
   569
            |> map (fn (thmN, thms) =>
53265
cc9a2976f836 qualify BNF constants properly
blanchet
parents: 53261
diff changeset
   570
              ((qualify (Binding.qualify true (Binding.name_of bnf_b) (Binding.name thmN)), []),
52720
fdc04f9bf906 separate ML interface to note facts of a bnf
traytel
parents: 52719
diff changeset
   571
              [(thms, [])]));
fdc04f9bf906 separate ML interface to note facts of a bnf
traytel
parents: 52719
diff changeset
   572
        in
fdc04f9bf906 separate ML interface to note facts of a bnf
traytel
parents: 52719
diff changeset
   573
          Local_Theory.notes notes #> snd
fdc04f9bf906 separate ML interface to note facts of a bnf
traytel
parents: 52719
diff changeset
   574
        end
fdc04f9bf906 separate ML interface to note facts of a bnf
traytel
parents: 52719
diff changeset
   575
      else
fdc04f9bf906 separate ML interface to note facts of a bnf
traytel
parents: 52719
diff changeset
   576
        I)
fdc04f9bf906 separate ML interface to note facts of a bnf
traytel
parents: 52719
diff changeset
   577
    #> (if fact_policy <> Dont_Note then
fdc04f9bf906 separate ML interface to note facts of a bnf
traytel
parents: 52719
diff changeset
   578
        let
fdc04f9bf906 separate ML interface to note facts of a bnf
traytel
parents: 52719
diff changeset
   579
          val notes =
53288
050d0bc9afa5 renamed BNF fact
blanchet
parents: 53287
diff changeset
   580
            [(map_compN, [Lazy.force (#map_comp facts)], []),
52720
fdc04f9bf906 separate ML interface to note facts of a bnf
traytel
parents: 52719
diff changeset
   581
            (map_cong0N, [#map_cong0 axioms], []),
fdc04f9bf906 separate ML interface to note facts of a bnf
traytel
parents: 52719
diff changeset
   582
            (map_congN, [Lazy.force (#map_cong facts)], fundef_cong_attrs),
53285
f09645642984 renamed BNF fact
blanchet
parents: 53270
diff changeset
   583
            (map_idN, [Lazy.force (#map_id facts)], []),
52720
fdc04f9bf906 separate ML interface to note facts of a bnf
traytel
parents: 52719
diff changeset
   584
            (rel_eqN, [Lazy.force (#rel_eq facts)], []),
fdc04f9bf906 separate ML interface to note facts of a bnf
traytel
parents: 52719
diff changeset
   585
            (rel_flipN, [Lazy.force (#rel_flip facts)], []),
53290
b6c3be868217 renamed BNF fact
blanchet
parents: 53289
diff changeset
   586
            (set_mapN, map Lazy.force (#set_map facts), []),
52720
fdc04f9bf906 separate ML interface to note facts of a bnf
traytel
parents: 52719
diff changeset
   587
            (rel_OO_GrpN, no_refl [#rel_OO_Grp axioms], []),
fdc04f9bf906 separate ML interface to note facts of a bnf
traytel
parents: 52719
diff changeset
   588
            (rel_GrpN, [Lazy.force (#rel_Grp facts)], []),
fdc04f9bf906 separate ML interface to note facts of a bnf
traytel
parents: 52719
diff changeset
   589
            (rel_conversepN, [Lazy.force (#rel_conversep facts)], []),
fdc04f9bf906 separate ML interface to note facts of a bnf
traytel
parents: 52719
diff changeset
   590
            (rel_monoN, [Lazy.force (#rel_mono facts)], []),
fdc04f9bf906 separate ML interface to note facts of a bnf
traytel
parents: 52719
diff changeset
   591
            (rel_OON, [Lazy.force (#rel_OO facts)], [])]
fdc04f9bf906 separate ML interface to note facts of a bnf
traytel
parents: 52719
diff changeset
   592
            |> filter_out (null o #2)
fdc04f9bf906 separate ML interface to note facts of a bnf
traytel
parents: 52719
diff changeset
   593
            |> map (fn (thmN, thms, attrs) =>
53265
cc9a2976f836 qualify BNF constants properly
blanchet
parents: 53261
diff changeset
   594
              ((qualify (Binding.qualify true (Binding.name_of bnf_b) (Binding.name thmN)),
52720
fdc04f9bf906 separate ML interface to note facts of a bnf
traytel
parents: 52719
diff changeset
   595
                attrs), [(thms, [])]));
fdc04f9bf906 separate ML interface to note facts of a bnf
traytel
parents: 52719
diff changeset
   596
        in
fdc04f9bf906 separate ML interface to note facts of a bnf
traytel
parents: 52719
diff changeset
   597
          Local_Theory.notes notes #> snd
fdc04f9bf906 separate ML interface to note facts of a bnf
traytel
parents: 52719
diff changeset
   598
        end
fdc04f9bf906 separate ML interface to note facts of a bnf
traytel
parents: 52719
diff changeset
   599
      else
fdc04f9bf906 separate ML interface to note facts of a bnf
traytel
parents: 52719
diff changeset
   600
        I)
fdc04f9bf906 separate ML interface to note facts of a bnf
traytel
parents: 52719
diff changeset
   601
  end;
fdc04f9bf906 separate ML interface to note facts of a bnf
traytel
parents: 52719
diff changeset
   602
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   603
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   604
(* Define new BNFs *)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   605
51767
bbcdd8519253 honor user-specified name for relator + generalize syntax
blanchet
parents: 51766
diff changeset
   606
fun prepare_def const_policy mk_fact_policy qualify prep_term Ds_opt map_b rel_b set_bs
53265
cc9a2976f836 qualify BNF constants properly
blanchet
parents: 53261
diff changeset
   607
  (((((raw_bnf_b, raw_map), raw_sets), raw_bd_Abs), raw_wits), raw_rel_opt) no_defs_lthy =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   608
  let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   609
    val fact_policy = mk_fact_policy no_defs_lthy;
53265
cc9a2976f836 qualify BNF constants properly
blanchet
parents: 53261
diff changeset
   610
    val bnf_b = qualify raw_bnf_b;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   611
    val live = length raw_sets;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   612
    val nwits = length raw_wits;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   613
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   614
    val map_rhs = prep_term no_defs_lthy raw_map;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   615
    val set_rhss = map (prep_term no_defs_lthy) raw_sets;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   616
    val (bd_rhsT, bd_rhs) = (case prep_term no_defs_lthy raw_bd_Abs of
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   617
      Abs (_, T, t) => (T, t)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   618
    | _ => error "Bad bound constant");
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   619
    val wit_rhss = map (prep_term no_defs_lthy) raw_wits;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   620
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
   621
    fun err T =
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
   622
      error ("Trying to register the type " ^ quote (Syntax.string_of_typ no_defs_lthy T) ^
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
   623
        " as unnamed BNF");
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
   624
53265
cc9a2976f836 qualify BNF constants properly
blanchet
parents: 53261
diff changeset
   625
    val (bnf_b, key) =
cc9a2976f836 qualify BNF constants properly
blanchet
parents: 53261
diff changeset
   626
      if Binding.eq_name (bnf_b, Binding.empty) then
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
   627
        (case bd_rhsT of
53265
cc9a2976f836 qualify BNF constants properly
blanchet
parents: 53261
diff changeset
   628
          Type (C, Ts) => if forall (can dest_TFree) Ts
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
   629
            then (Binding.qualified_name C, C) else err bd_rhsT
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
   630
        | T => err T)
53265
cc9a2976f836 qualify BNF constants properly
blanchet
parents: 53261
diff changeset
   631
      else (bnf_b, Local_Theory.full_name no_defs_lthy bnf_b);
cc9a2976f836 qualify BNF constants properly
blanchet
parents: 53261
diff changeset
   632
cc9a2976f836 qualify BNF constants properly
blanchet
parents: 53261
diff changeset
   633
    val def_qualify = Binding.qualify false (Binding.name_of bnf_b);
cc9a2976f836 qualify BNF constants properly
blanchet
parents: 53261
diff changeset
   634
cc9a2976f836 qualify BNF constants properly
blanchet
parents: 53261
diff changeset
   635
    fun mk_suffix_binding suf = Binding.suffix_name ("_" ^ suf) 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
   636
49463
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49462
diff changeset
   637
    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
   638
      let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   639
        val inline =
49538
c90818b63599 simplified fact policies
blanchet
parents: 49537
diff changeset
   640
          (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
   641
          (case const_policy of
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   642
            Dont_Inline => false
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   643
          | Hardly_Inline => Term.is_Free rhs orelse Term.is_Const rhs
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   644
          | Smart_Inline => Term.size_of_term rhs <= smart_max_inline_size
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   645
          | Do_Inline => true)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   646
      in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   647
        if inline then
49460
4dd451a075ce fixed infinite loop with trivial rel_O_Gr + tuning
blanchet
parents: 49459
diff changeset
   648
          ((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
   649
        else
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   650
          let val b = b () in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   651
            apfst (apsnd snd) (Local_Theory.define ((b, NoSyn), ((Thm.def_binding b, []), rhs))
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   652
              lthy)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   653
          end
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   654
      end;
49459
3f8e2b5249ec adapted FP code to new relator approach
blanchet
parents: 49458
diff changeset
   655
3f8e2b5249ec adapted FP code to new relator approach
blanchet
parents: 49458
diff changeset
   656
    fun maybe_restore lthy_old lthy =
3f8e2b5249ec adapted FP code to new relator approach
blanchet
parents: 49458
diff changeset
   657
      lthy |> not (pointer_eq (lthy_old, lthy)) ? Local_Theory.restore;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   658
51758
55963309557b honor user-specified name for map function
blanchet
parents: 51757
diff changeset
   659
    val map_bind_def =
53265
cc9a2976f836 qualify BNF constants properly
blanchet
parents: 53261
diff changeset
   660
      (fn () => def_qualify (if Binding.is_empty map_b then mk_suffix_binding mapN else map_b),
cc9a2976f836 qualify BNF constants properly
blanchet
parents: 53261
diff changeset
   661
         map_rhs);
49459
3f8e2b5249ec adapted FP code to new relator approach
blanchet
parents: 49458
diff changeset
   662
    val set_binds_defs =
3f8e2b5249ec adapted FP code to new relator approach
blanchet
parents: 49458
diff changeset
   663
      let
51757
7babcb61aa5c honor user-specified set function names
blanchet
parents: 51551
diff changeset
   664
        fun set_name i get_b =
7babcb61aa5c honor user-specified set function names
blanchet
parents: 51551
diff changeset
   665
          (case try (nth set_bs) (i - 1) of
7babcb61aa5c honor user-specified set function names
blanchet
parents: 51551
diff changeset
   666
            SOME b => if Binding.is_empty b then get_b else K b
53265
cc9a2976f836 qualify BNF constants properly
blanchet
parents: 53261
diff changeset
   667
          | NONE => get_b) #> def_qualify;
cc9a2976f836 qualify BNF constants properly
blanchet
parents: 53261
diff changeset
   668
        val bs = if live = 1 then [set_name 1 (fn () => mk_suffix_binding setN)]
cc9a2976f836 qualify BNF constants properly
blanchet
parents: 53261
diff changeset
   669
          else map (fn i => set_name i (fn () => mk_suffix_binding (mk_setN i))) (1 upto live);
51757
7babcb61aa5c honor user-specified set function names
blanchet
parents: 51551
diff changeset
   670
      in bs ~~ set_rhss end;
53265
cc9a2976f836 qualify BNF constants properly
blanchet
parents: 53261
diff changeset
   671
    val bd_bind_def = (fn () => def_qualify (mk_suffix_binding bdN), bd_rhs);
49459
3f8e2b5249ec adapted FP code to new relator approach
blanchet
parents: 49458
diff changeset
   672
    val wit_binds_defs =
3f8e2b5249ec adapted FP code to new relator approach
blanchet
parents: 49458
diff changeset
   673
      let
53265
cc9a2976f836 qualify BNF constants properly
blanchet
parents: 53261
diff changeset
   674
        val bs = if nwits = 1 then [fn () => def_qualify (mk_suffix_binding witN)]
cc9a2976f836 qualify BNF constants properly
blanchet
parents: 53261
diff changeset
   675
          else map (fn i => fn () => def_qualify (mk_suffix_binding (mk_witN i))) (1 upto nwits);
51757
7babcb61aa5c honor user-specified set function names
blanchet
parents: 51551
diff changeset
   676
      in bs ~~ wit_rhss end;
49459
3f8e2b5249ec adapted FP code to new relator approach
blanchet
parents: 49458
diff changeset
   677
3f8e2b5249ec adapted FP code to new relator approach
blanchet
parents: 49458
diff changeset
   678
    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
   679
      (bnf_set_terms, raw_set_defs)),
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   680
      (bnf_bd_term, raw_bd_def)),
49459
3f8e2b5249ec adapted FP code to new relator approach
blanchet
parents: 49458
diff changeset
   681
      (bnf_wit_terms, raw_wit_defs)), (lthy, lthy_old)) =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   682
        no_defs_lthy
49463
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49462
diff changeset
   683
        |> maybe_define true map_bind_def
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49462
diff changeset
   684
        ||>> apfst split_list o fold_map (maybe_define true) set_binds_defs
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49462
diff changeset
   685
        ||>> maybe_define true bd_bind_def
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49462
diff changeset
   686
        ||>> apfst split_list o fold_map (maybe_define true) wit_binds_defs
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   687
        ||> `(maybe_restore no_defs_lthy);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   688
49459
3f8e2b5249ec adapted FP code to new relator approach
blanchet
parents: 49458
diff changeset
   689
    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
   690
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   691
    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
   692
    val bnf_set_defs = map (Morphism.thm phi) raw_set_defs;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   693
    val bnf_bd_def = Morphism.thm phi raw_bd_def;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   694
    val bnf_wit_defs = map (Morphism.thm phi) raw_wit_defs;
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
    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
   697
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   698
    (*TODO: handle errors*)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   699
    (*simple shape analysis of a map function*)
49395
323414474c1f use strip_typeN in bnf_def (instead of repairing strip_type)
traytel
parents: 49386
diff changeset
   700
    val ((alphas, betas), (CA, _)) =
323414474c1f use strip_typeN in bnf_def (instead of repairing strip_type)
traytel
parents: 49386
diff changeset
   701
      fastype_of bnf_map
323414474c1f use strip_typeN in bnf_def (instead of repairing strip_type)
traytel
parents: 49386
diff changeset
   702
      |> strip_typeN live
323414474c1f use strip_typeN in bnf_def (instead of repairing strip_type)
traytel
parents: 49386
diff changeset
   703
      |>> map_split dest_funT
323414474c1f use strip_typeN in bnf_def (instead of repairing strip_type)
traytel
parents: 49386
diff changeset
   704
      ||> dest_funT
323414474c1f use strip_typeN in bnf_def (instead of repairing strip_type)
traytel
parents: 49386
diff changeset
   705
      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
   706
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   707
    val CA_params = map TVar (Term.add_tvarsT CA []);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   708
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   709
    val bnf_sets = map2 (normalize_set CA_params) alphas (map (Morphism.term phi) bnf_set_terms);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   710
    val bdT = Morphism.typ phi bd_rhsT;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   711
    val bnf_bd =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   712
      Term.subst_TVars (Term.add_tvar_namesT bdT [] ~~ CA_params) (Morphism.term phi bnf_bd_term);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   713
    val bnf_wits = map (normalize_wit CA_params CA alphas o Morphism.term phi) bnf_wit_terms;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   714
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   715
    (*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
   716
    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
   717
      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
   718
    | 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
   719
    val dead = length deads;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   720
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   721
    (*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
   722
    (*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
   723
    (*TODO: check type of bnf_bd*)
49507
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49506
diff changeset
   724
    (*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
   725
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   726
    val ((((((((((As', Bs'), Cs), Ds), B1Ts), B2Ts), domTs), ranTs), ranTs'), ranTs''),
49459
3f8e2b5249ec adapted FP code to new relator approach
blanchet
parents: 49458
diff changeset
   727
      (Ts, T)) = lthy
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   728
      |> mk_TFrees live
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   729
      ||>> mk_TFrees live
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   730
      ||>> mk_TFrees live
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   731
      ||>> mk_TFrees dead
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   732
      ||>> mk_TFrees live
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   733
      ||>> mk_TFrees live
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   734
      ||>> mk_TFrees live
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   735
      ||>> mk_TFrees live
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   736
      ||>> mk_TFrees live
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   737
      ||>> mk_TFrees live
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   738
      ||> fst o mk_TFrees 1
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   739
      ||> the_single
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   740
      ||> `(replicate live);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   741
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   742
    fun mk_bnf_map As' Bs' =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   743
      Term.subst_atomic_types ((deads ~~ Ds) @ (alphas ~~ As') @ (betas ~~ Bs')) bnf_map;
49453
ff0e540d8758 add rel as first-class citizen of BNF
blanchet
parents: 49452
diff changeset
   744
    fun mk_bnf_t As' = Term.subst_atomic_types ((deads ~~ Ds) @ (alphas ~~ As'));
ff0e540d8758 add rel as first-class citizen of BNF
blanchet
parents: 49452
diff changeset
   745
    fun mk_bnf_T As' = Term.typ_subst_atomic ((deads ~~ Ds) @ (alphas ~~ As'));
ff0e540d8758 add rel as first-class citizen of BNF
blanchet
parents: 49452
diff changeset
   746
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
   747
    val RTs = map HOLogic.mk_prodT (As' ~~ Bs');
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
   748
    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
   749
    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
   750
    val pred2RTsBsCs = map2 mk_pred2T Bs' Cs;
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
   751
    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
   752
    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
   753
    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
   754
    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
   755
49453
ff0e540d8758 add rel as first-class citizen of BNF
blanchet
parents: 49452
diff changeset
   756
    val CA' = mk_bnf_T As' CA;
ff0e540d8758 add rel as first-class citizen of BNF
blanchet
parents: 49452
diff changeset
   757
    val CB' = mk_bnf_T Bs' CA;
ff0e540d8758 add rel as first-class citizen of BNF
blanchet
parents: 49452
diff changeset
   758
    val CC' = mk_bnf_T Cs CA;
ff0e540d8758 add rel as first-class citizen of BNF
blanchet
parents: 49452
diff changeset
   759
    val CRs' = mk_bnf_T RTs CA;
52719
480a3479fa47 transfer rule for map (not yet registered as a transfer rule)
traytel
parents: 52635
diff changeset
   760
    val CB1 = mk_bnf_T B1Ts CA;
480a3479fa47 transfer rule for map (not yet registered as a transfer rule)
traytel
parents: 52635
diff changeset
   761
    val CB2 = mk_bnf_T B2Ts CA;
49453
ff0e540d8758 add rel as first-class citizen of BNF
blanchet
parents: 49452
diff changeset
   762
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   763
    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
   764
    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
   765
    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
   766
    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
   767
    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
   768
    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
   769
    val bnf_bd_As = mk_bnf_t As' bnf_bd;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   770
    val bnf_wit_As = map (apsnd (mk_bnf_t As')) bnf_wits;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   771
49595
e8c57e59cbf8 got rid of other instance of shaky "Thm.generalize"
blanchet
parents: 49594
diff changeset
   772
    val pre_names_lthy = lthy;
53561
92bcac4f9ac9 simplified derivation of in_rel
traytel
parents: 53290
diff changeset
   773
    val ((((((((((((((((((((((((fs, gs), hs), x), y), zs), ys), As),
49459
3f8e2b5249ec adapted FP code to new relator approach
blanchet
parents: 49458
diff changeset
   774
      As_copy), Xs), B1s), B2s), f1s), f2s), e1s), e2s), p1s), p2s), bs), (Rs, Rs')), Rs_copy), Ss),
52719
480a3479fa47 transfer rule for map (not yet registered as a transfer rule)
traytel
parents: 52635
diff changeset
   775
      transfer_domRs), transfer_ranRs), names_lthy) = pre_names_lthy
52923
traytel
parents: 52844
diff changeset
   776
      |> mk_Frees "f" (map2 (curry op -->) As' Bs')
traytel
parents: 52844
diff changeset
   777
      ||>> mk_Frees "g" (map2 (curry op -->) Bs' Cs)
traytel
parents: 52844
diff changeset
   778
      ||>> mk_Frees "h" (map2 (curry op -->) As' Ts)
51894
traytel
parents: 51893
diff changeset
   779
      ||>> yield_singleton (mk_Frees "x") CA'
traytel
parents: 51893
diff changeset
   780
      ||>> 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
   781
      ||>> mk_Frees "z" As'
51916
eac9e9a45bf5 stronger monotonicity property for relators
traytel
parents: 51915
diff changeset
   782
      ||>> mk_Frees "y" Bs'
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   783
      ||>> 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
   784
      ||>> 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
   785
      ||>> mk_Frees "A" (map HOLogic.mk_setT domTs)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   786
      ||>> mk_Frees "B1" (map HOLogic.mk_setT B1Ts)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   787
      ||>> mk_Frees "B2" (map HOLogic.mk_setT B2Ts)
52923
traytel
parents: 52844
diff changeset
   788
      ||>> mk_Frees "f1" (map2 (curry op -->) B1Ts ranTs)
traytel
parents: 52844
diff changeset
   789
      ||>> mk_Frees "f2" (map2 (curry op -->) B2Ts ranTs)
traytel
parents: 52844
diff changeset
   790
      ||>> mk_Frees "e1" (map2 (curry op -->) B1Ts ranTs')
traytel
parents: 52844
diff changeset
   791
      ||>> mk_Frees "e2" (map2 (curry op -->) B2Ts ranTs'')
traytel
parents: 52844
diff changeset
   792
      ||>> mk_Frees "p1" (map2 (curry op -->) domTs B1Ts)
traytel
parents: 52844
diff changeset
   793
      ||>> mk_Frees "p2" (map2 (curry op -->) domTs B2Ts)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   794
      ||>> mk_Frees "b" As'
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
   795
      ||>> mk_Frees' "R" pred2RTs
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
   796
      ||>> mk_Frees "R" pred2RTs
52719
480a3479fa47 transfer rule for map (not yet registered as a transfer rule)
traytel
parents: 52635
diff changeset
   797
      ||>> mk_Frees "S" pred2RTsBsCs
480a3479fa47 transfer rule for map (not yet registered as a transfer rule)
traytel
parents: 52635
diff changeset
   798
      ||>> mk_Frees "R" transfer_domRTs
480a3479fa47 transfer rule for map (not yet registered as a transfer rule)
traytel
parents: 52635
diff changeset
   799
      ||>> 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
   800
51762
219a3063ed29 derive "map_cong"
blanchet
parents: 51761
diff changeset
   801
    val fs_copy = map2 (retype_free o fastype_of) fs gs;
219a3063ed29 derive "map_cong"
blanchet
parents: 51761
diff changeset
   802
    val x_copy = retype_free CA' y;
219a3063ed29 derive "map_cong"
blanchet
parents: 51761
diff changeset
   803
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
   804
    val setRs =
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
   805
      map3 (fn R => fn T => fn U =>
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
   806
          HOLogic.Collect_const (HOLogic.mk_prodT (T, U)) $ HOLogic.mk_split R) Rs As' Bs';
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
   807
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
   808
    (*Grp (in (Collect (split R1) .. Collect (split Rn))) (map fst .. fst)^--1 OO
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
   809
      Grp (in (Collect (split R1) .. Collect (split Rn))) (map snd .. snd)*)
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
   810
    val OO_Grp =
49459
3f8e2b5249ec adapted FP code to new relator approach
blanchet
parents: 49458
diff changeset
   811
      let
3f8e2b5249ec adapted FP code to new relator approach
blanchet
parents: 49458
diff changeset
   812
        val map1 = Term.list_comb (mk_bnf_map RTs As', map fst_const RTs);
3f8e2b5249ec adapted FP code to new relator approach
blanchet
parents: 49458
diff changeset
   813
        val map2 = Term.list_comb (mk_bnf_map RTs Bs', map snd_const RTs);
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
   814
        val bnf_in = mk_in setRs (map (mk_bnf_t RTs) bnf_sets) CRs';
49459
3f8e2b5249ec adapted FP code to new relator approach
blanchet
parents: 49458
diff changeset
   815
      in
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
   816
        mk_rel_compp (mk_conversep (mk_Grp bnf_in map1), mk_Grp bnf_in map2)
49459
3f8e2b5249ec adapted FP code to new relator approach
blanchet
parents: 49458
diff changeset
   817
      end;
3f8e2b5249ec adapted FP code to new relator approach
blanchet
parents: 49458
diff changeset
   818
49507
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49506
diff changeset
   819
    val rel_rhs = (case raw_rel_opt of
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
   820
        NONE => fold_rev Term.absfree Rs' OO_Grp
49507
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49506
diff changeset
   821
      | SOME raw_rel => prep_term no_defs_lthy raw_rel);
49463
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49462
diff changeset
   822
51767
bbcdd8519253 honor user-specified name for relator + generalize syntax
blanchet
parents: 51766
diff changeset
   823
    val rel_bind_def =
53265
cc9a2976f836 qualify BNF constants properly
blanchet
parents: 53261
diff changeset
   824
      (fn () => def_qualify (if Binding.is_empty rel_b then mk_suffix_binding relN else rel_b),
cc9a2976f836 qualify BNF constants properly
blanchet
parents: 53261
diff changeset
   825
         rel_rhs);
49463
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49462
diff changeset
   826
49507
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49506
diff changeset
   827
    val ((bnf_rel_term, raw_rel_def), (lthy, lthy_old)) =
49463
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49462
diff changeset
   828
      lthy
49507
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49506
diff changeset
   829
      |> maybe_define (is_some raw_rel_opt) rel_bind_def
49463
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49462
diff changeset
   830
      ||> `(maybe_restore lthy);
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49462
diff changeset
   831
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49462
diff changeset
   832
    val phi = Proof_Context.export_morphism lthy_old lthy;
49507
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49506
diff changeset
   833
    val bnf_rel_def = Morphism.thm phi raw_rel_def;
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49506
diff changeset
   834
    val bnf_rel = Morphism.term phi bnf_rel_term;
49463
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49462
diff changeset
   835
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
   836
    fun mk_bnf_rel RTs CA' CB' = normalize_rel lthy RTs CA' CB' bnf_rel;
49459
3f8e2b5249ec adapted FP code to new relator approach
blanchet
parents: 49458
diff changeset
   837
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
   838
    val rel = mk_bnf_rel pred2RTs CA' CB';
52719
480a3479fa47 transfer rule for map (not yet registered as a transfer rule)
traytel
parents: 52635
diff changeset
   839
    val relAsAs = mk_bnf_rel self_pred2RTs CA' CA';
49459
3f8e2b5249ec adapted FP code to new relator approach
blanchet
parents: 49458
diff changeset
   840
49460
4dd451a075ce fixed infinite loop with trivial rel_O_Gr + tuning
blanchet
parents: 49459
diff changeset
   841
    val _ = case no_reflexive (raw_map_def :: raw_set_defs @ [raw_bd_def] @
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
   842
        raw_wit_defs @ [raw_rel_def]) of
49459
3f8e2b5249ec adapted FP code to new relator approach
blanchet
parents: 49458
diff changeset
   843
        [] => ()
3f8e2b5249ec adapted FP code to new relator approach
blanchet
parents: 49458
diff changeset
   844
      | defs => Proof_Display.print_consts true lthy_old (K false)
3f8e2b5249ec adapted FP code to new relator approach
blanchet
parents: 49458
diff changeset
   845
          (map (dest_Free o fst o Logic.dest_equals o prop_of) defs);
3f8e2b5249ec adapted FP code to new relator approach
blanchet
parents: 49458
diff changeset
   846
53270
c8628119d18e renamed BNF axiom
blanchet
parents: 53265
diff changeset
   847
    val map_id0_goal =
51762
219a3063ed29 derive "map_cong"
blanchet
parents: 51761
diff changeset
   848
      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
   849
        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
   850
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   851
53287
271b34513bfb renamed BNF axiom
blanchet
parents: 53285
diff changeset
   852
    val map_comp0_goal =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   853
      let
49018
b2884253b421 renamed ML function for consistency
blanchet
parents: 49016
diff changeset
   854
        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
   855
        val comp_bnf_map_app = HOLogic.mk_comp
51762
219a3063ed29 derive "map_cong"
blanchet
parents: 51761
diff changeset
   856
          (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
   857
      in
49123
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents: 49111
diff changeset
   858
        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
   859
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   860
51762
219a3063ed29 derive "map_cong"
blanchet
parents: 51761
diff changeset
   861
    fun mk_map_cong_prem x z set f f_copy =
219a3063ed29 derive "map_cong"
blanchet
parents: 51761
diff changeset
   862
      Logic.all z (Logic.mk_implies
219a3063ed29 derive "map_cong"
blanchet
parents: 51761
diff changeset
   863
        (HOLogic.mk_Trueprop (HOLogic.mk_mem (z, set $ x)),
219a3063ed29 derive "map_cong"
blanchet
parents: 51761
diff changeset
   864
        mk_Trueprop_eq (f $ z, f_copy $ z)));
219a3063ed29 derive "map_cong"
blanchet
parents: 51761
diff changeset
   865
51761
4c9f08836d87 renamed "map_cong" axiom to "map_cong0" in preparation for real "map_cong"
blanchet
parents: 51758
diff changeset
   866
    val map_cong0_goal =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   867
      let
51762
219a3063ed29 derive "map_cong"
blanchet
parents: 51761
diff changeset
   868
        val prems = map4 (mk_map_cong_prem x) zs bnf_sets_As fs fs_copy;
219a3063ed29 derive "map_cong"
blanchet
parents: 51761
diff changeset
   869
        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
   870
          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
   871
      in
51762
219a3063ed29 derive "map_cong"
blanchet
parents: 51761
diff changeset
   872
        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
   873
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   874
53289
5e0623448bdb renamed BNF axiom
blanchet
parents: 53288
diff changeset
   875
    val set_map0s_goal =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   876
      let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   877
        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
   878
          let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   879
            val set_comp_map =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   880
              HOLogic.mk_comp (setB, Term.list_comb (bnf_map_AsBs, fs));
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   881
            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
   882
          in
49123
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents: 49111
diff changeset
   883
            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
   884
          end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   885
      in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   886
        map3 mk_goal bnf_sets_As bnf_sets_Bs fs
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   887
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   888
49458
blanchet
parents: 49456
diff changeset
   889
    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
   890
49458
blanchet
parents: 49456
diff changeset
   891
    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
   892
49458
blanchet
parents: 49456
diff changeset
   893
    val set_bds_goal =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   894
      let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   895
        fun mk_goal set =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   896
          Logic.all x (HOLogic.mk_Trueprop (mk_ordLeq (mk_card_of (set $ x)) bnf_bd_As));
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   897
      in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   898
        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
   899
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   900
49458
blanchet
parents: 49456
diff changeset
   901
    val map_wpull_goal =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   902
      let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   903
        val prems = map HOLogic.mk_Trueprop
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   904
          (map8 mk_wpull Xs B1s B2s f1s f2s (replicate live NONE) p1s p2s);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   905
        val CX = mk_bnf_T domTs CA;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   906
        val bnf_sets_CX = map2 (normalize_set (map (mk_bnf_T domTs) CA_params)) domTs bnf_sets;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   907
        val bnf_sets_CB1 = map2 (normalize_set (map (mk_bnf_T B1Ts) CA_params)) B1Ts bnf_sets;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   908
        val bnf_sets_CB2 = map2 (normalize_set (map (mk_bnf_T B2Ts) CA_params)) B2Ts bnf_sets;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   909
        val bnf_map_app_f1 = Term.list_comb (mk_bnf_map B1Ts ranTs, f1s);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   910
        val bnf_map_app_f2 = Term.list_comb (mk_bnf_map B2Ts ranTs, f2s);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   911
        val bnf_map_app_p1 = Term.list_comb (mk_bnf_map domTs B1Ts, p1s);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   912
        val bnf_map_app_p2 = Term.list_comb (mk_bnf_map domTs B2Ts, p2s);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   913
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   914
        val map_wpull = mk_wpull (mk_in Xs bnf_sets_CX CX)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   915
          (mk_in B1s bnf_sets_CB1 CB1) (mk_in B2s bnf_sets_CB2 CB2)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   916
          bnf_map_app_f1 bnf_map_app_f2 NONE bnf_map_app_p1 bnf_map_app_p2;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   917
      in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   918
        fold_rev Logic.all (Xs @ B1s @ B2s @ f1s @ f2s @ p1s @ p2s)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   919
          (Logic.list_implies (prems, HOLogic.mk_Trueprop map_wpull))
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   920
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   921
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
   922
    val rel_OO_Grp_goal = fold_rev Logic.all Rs (mk_Trueprop_eq (Term.list_comb (rel, Rs), OO_Grp));
49453
ff0e540d8758 add rel as first-class citizen of BNF
blanchet
parents: 49452
diff changeset
   923
53289
5e0623448bdb renamed BNF axiom
blanchet
parents: 53288
diff changeset
   924
    val goals = zip_axioms map_id0_goal map_comp0_goal map_cong0_goal set_map0s_goal
53270
c8628119d18e renamed BNF axiom
blanchet
parents: 53265
diff changeset
   925
      card_order_bd_goal cinfinite_bd_goal set_bds_goal map_wpull_goal rel_OO_Grp_goal;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   926
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   927
    fun mk_wit_goals (I, wit) =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   928
      let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   929
        val xs = map (nth bs) I;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   930
        fun wit_goal i =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   931
          let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   932
            val z = nth zs i;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   933
            val set_wit = nth bnf_sets_As i $ Term.list_comb (wit, xs);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   934
            val concl = HOLogic.mk_Trueprop
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   935
              (if member (op =) I i then HOLogic.mk_eq (z, nth bs i)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   936
              else @{term False});
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   937
          in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   938
            fold_rev Logic.all (z :: xs)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   939
              (Logic.mk_implies (HOLogic.mk_Trueprop (HOLogic.mk_mem (z, set_wit)), concl))
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   940
          end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   941
      in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   942
        map wit_goal (0 upto live - 1)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   943
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   944
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   945
    val wit_goalss = map mk_wit_goals bnf_wit_As;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   946
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   947
    fun after_qed thms lthy =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   948
      let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   949
        val (axioms, wit_thms) = apfst (mk_axioms live) (chop (length goals) thms);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   950
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49103
diff changeset
   951
        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
   952
        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
   953
        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
   954
51766
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51765
diff changeset
   955
        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
   956
          let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   957
            val defT = mk_bnf_T Ts CA --> HOLogic.mk_setT T;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   958
            val collect_map = HOLogic.mk_comp
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   959
              (mk_collect (map (mk_bnf_t Ts) bnf_sets) defT,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   960
              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
   961
            val image_collect = mk_collect
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   962
              (map2 (fn h => fn set => HOLogic.mk_comp (mk_image h, set)) hs bnf_sets_As)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   963
              defT;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   964
            (*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
   965
            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
   966
          in
53289
5e0623448bdb renamed BNF axiom
blanchet
parents: 53288
diff changeset
   967
            Goal.prove_sorry lthy [] [] goal (K (mk_collect_set_map_tac (#set_map0 axioms)))
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49103
diff changeset
   968
            |> Thm.close_derivation
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   969
          end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   970
51766
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51765
diff changeset
   971
        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
   972
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   973
        fun mk_in_mono () =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   974
          let
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
   975
            val prems_mono = map2 (HOLogic.mk_Trueprop oo mk_leq) As As_copy;
49458
blanchet
parents: 49456
diff changeset
   976
            val in_mono_goal =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   977
              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
   978
                (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
   979
                  (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
   980
          in
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 49631
diff changeset
   981
            Goal.prove_sorry lthy [] [] in_mono_goal (K (mk_in_mono_tac live))
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49103
diff changeset
   982
            |> Thm.close_derivation
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   983
          end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   984
49538
c90818b63599 simplified fact policies
blanchet
parents: 49537
diff changeset
   985
        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
   986
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   987
        fun mk_in_cong () =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   988
          let
51762
219a3063ed29 derive "map_cong"
blanchet
parents: 51761
diff changeset
   989
            val prems_cong = map2 (curry mk_Trueprop_eq) As As_copy;
49458
blanchet
parents: 49456
diff changeset
   990
            val in_cong_goal =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   991
              fold_rev Logic.all (As @ As_copy)
51762
219a3063ed29 derive "map_cong"
blanchet
parents: 51761
diff changeset
   992
                (Logic.list_implies (prems_cong,
219a3063ed29 derive "map_cong"
blanchet
parents: 51761
diff changeset
   993
                  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
   994
          in
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 51790
diff changeset
   995
            Goal.prove_sorry lthy [] [] in_cong_goal
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 51790
diff changeset
   996
              (K ((TRY o hyp_subst_tac lthy THEN' rtac refl) 1))
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49103
diff changeset
   997
            |> Thm.close_derivation
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   998
          end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   999
49538
c90818b63599 simplified fact policies
blanchet
parents: 49537
diff changeset
  1000
        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
  1001
53285
f09645642984 renamed BNF fact
blanchet
parents: 53270
diff changeset
  1002
        val map_id = Lazy.lazy (fn () => mk_map_id (#map_id0 axioms));
53288
050d0bc9afa5 renamed BNF fact
blanchet
parents: 53287
diff changeset
  1003
        val map_comp = Lazy.lazy (fn () => mk_map_comp (#map_comp0 axioms));
51762
219a3063ed29 derive "map_cong"
blanchet
parents: 51761
diff changeset
  1004
219a3063ed29 derive "map_cong"
blanchet
parents: 51761
diff changeset
  1005
        fun mk_map_cong () =
219a3063ed29 derive "map_cong"
blanchet
parents: 51761
diff changeset
  1006
          let
219a3063ed29 derive "map_cong"
blanchet
parents: 51761
diff changeset
  1007
            val prem0 = mk_Trueprop_eq (x, x_copy);
219a3063ed29 derive "map_cong"
blanchet
parents: 51761
diff changeset
  1008
            val prems = map4 (mk_map_cong_prem x_copy) zs bnf_sets_As fs fs_copy;
219a3063ed29 derive "map_cong"
blanchet
parents: 51761
diff changeset
  1009
            val eq = mk_Trueprop_eq (Term.list_comb (bnf_map_AsBs, fs) $ x,
219a3063ed29 derive "map_cong"
blanchet
parents: 51761
diff changeset
  1010
              Term.list_comb (bnf_map_AsBs, fs_copy) $ x_copy);
219a3063ed29 derive "map_cong"
blanchet
parents: 51761
diff changeset
  1011
            val goal = fold_rev Logic.all (x :: x_copy :: fs @ fs_copy)
219a3063ed29 derive "map_cong"
blanchet
parents: 51761
diff changeset
  1012
              (Logic.list_implies (prem0 :: prems, eq));
219a3063ed29 derive "map_cong"
blanchet
parents: 51761
diff changeset
  1013
          in
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 51790
diff changeset
  1014
            Goal.prove_sorry lthy [] [] goal (fn _ => mk_map_cong_tac lthy (#map_cong0 axioms))
51762
219a3063ed29 derive "map_cong"
blanchet
parents: 51761
diff changeset
  1015
            |> Thm.close_derivation
219a3063ed29 derive "map_cong"
blanchet
parents: 51761
diff changeset
  1016
          end;
219a3063ed29 derive "map_cong"
blanchet
parents: 51761
diff changeset
  1017
219a3063ed29 derive "map_cong"
blanchet
parents: 51761
diff changeset
  1018
        val map_cong = Lazy.lazy mk_map_cong;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1019
53290
b6c3be868217 renamed BNF fact
blanchet
parents: 53289
diff changeset
  1020
        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
  1021
52635
4f84b730c489 got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
traytel
parents: 51930
diff changeset
  1022
        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
  1023
          let
52813
963297a24206 more robust tactic
traytel
parents: 52731
diff changeset
  1024
            val bdT = fst (dest_relT (fastype_of bnf_bd_As));
963297a24206 more robust tactic
traytel
parents: 52731
diff changeset
  1025
            val bdTs = replicate live bdT;
963297a24206 more robust tactic
traytel
parents: 52731
diff changeset
  1026
            val bd_bnfT = mk_bnf_T bdTs CA;
963297a24206 more robust tactic
traytel
parents: 52731
diff changeset
  1027
            val surj_imp_ordLeq_inst = (if live = 0 then TrueI else
963297a24206 more robust tactic
traytel
parents: 52731
diff changeset
  1028
              let
963297a24206 more robust tactic
traytel
parents: 52731
diff changeset
  1029
                val ranTs = map (fn AT => mk_sumT (AT, HOLogic.unitT)) As';
963297a24206 more robust tactic
traytel
parents: 52731
diff changeset
  1030
                val funTs = map (fn T => bdT --> T) ranTs;
963297a24206 more robust tactic
traytel
parents: 52731
diff changeset
  1031
                val ran_bnfT = mk_bnf_T ranTs CA;
963297a24206 more robust tactic
traytel
parents: 52731
diff changeset
  1032
                val (revTs, Ts) = `rev (bd_bnfT :: funTs);
963297a24206 more robust tactic
traytel
parents: 52731
diff changeset
  1033
                val cTs = map (SOME o certifyT lthy) [ran_bnfT, Library.foldr1 HOLogic.mk_prodT Ts];
963297a24206 more robust tactic
traytel
parents: 52731
diff changeset
  1034
                val tinst = fold (fn T => fn t => HOLogic.mk_split (Term.absdummy T t)) (tl revTs)
963297a24206 more robust tactic
traytel
parents: 52731
diff changeset
  1035
                  (Term.absdummy (hd revTs) (Term.list_comb (mk_bnf_map bdTs ranTs,
963297a24206 more robust tactic
traytel
parents: 52731
diff changeset
  1036
                    map Bound (live - 1 downto 0)) $ Bound live));
963297a24206 more robust tactic
traytel
parents: 52731
diff changeset
  1037
                val cts = [NONE, SOME (certify lthy tinst)];
963297a24206 more robust tactic
traytel
parents: 52731
diff changeset
  1038
              in
963297a24206 more robust tactic
traytel
parents: 52731
diff changeset
  1039
                Drule.instantiate' cTs cts @{thm surj_imp_ordLeq}
963297a24206 more robust tactic
traytel
parents: 52731
diff changeset
  1040
              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
  1041
            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
  1042
              (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
  1043
                else mk_csum (Library.foldr1 (uncurry mk_csum) (map mk_card_of As)) ctwo)
52813
963297a24206 more robust tactic
traytel
parents: 52731
diff changeset
  1044
              (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
  1045
            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
  1046
              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
  1047
                (HOLogic.mk_Trueprop (mk_ordLeq (mk_card_of (mk_in As bnf_sets_As CA')) bd));
4f84b730c489 got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
traytel
parents: 51930
diff changeset
  1048
          in
4f84b730c489 got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
traytel
parents: 51930
diff changeset
  1049
            Goal.prove_sorry lthy [] [] in_bd_goal
52813
963297a24206 more robust tactic
traytel
parents: 52731
diff changeset
  1050
              (mk_in_bd_tac live surj_imp_ordLeq_inst
53288
050d0bc9afa5 renamed BNF fact
blanchet
parents: 53287
diff changeset
  1051
                (Lazy.force map_comp) (Lazy.force map_id) (#map_cong0 axioms)
53290
b6c3be868217 renamed BNF fact
blanchet
parents: 53289
diff changeset
  1052
                (map Lazy.force set_map) (#set_bd axioms) (#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
  1053
                bd_Card_order bd_Cinfinite bd_Cnotzero)
4f84b730c489 got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
traytel
parents: 51930
diff changeset
  1054
            |> Thm.close_derivation
4f84b730c489 got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
traytel
parents: 51930
diff changeset
  1055
          end;
4f84b730c489 got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
traytel
parents: 51930
diff changeset
  1056
4f84b730c489 got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
traytel
parents: 51930
diff changeset
  1057
        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
  1058
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1059
        fun mk_map_wppull () =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1060
          let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1061
            val prems = if live = 0 then [] else
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1062
              [HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1063
                (map8 mk_wpull Xs B1s B2s f1s f2s (map SOME (e1s ~~ e2s)) p1s p2s))];
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1064
            val CX = mk_bnf_T domTs CA;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1065
            val bnf_sets_CX =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1066
              map2 (normalize_set (map (mk_bnf_T domTs) CA_params)) domTs bnf_sets;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1067
            val bnf_sets_CB1 =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1068
              map2 (normalize_set (map (mk_bnf_T B1Ts) CA_params)) B1Ts bnf_sets;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1069
            val bnf_sets_CB2 =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1070
              map2 (normalize_set (map (mk_bnf_T B2Ts) CA_params)) B2Ts bnf_sets;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1071
            val bnf_map_app_f1 = Term.list_comb (mk_bnf_map B1Ts ranTs, f1s);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1072
            val bnf_map_app_f2 = Term.list_comb (mk_bnf_map B2Ts ranTs, f2s);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1073
            val bnf_map_app_e1 = Term.list_comb (mk_bnf_map B1Ts ranTs', e1s);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1074
            val bnf_map_app_e2 = Term.list_comb (mk_bnf_map B2Ts ranTs'', e2s);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1075
            val bnf_map_app_p1 = Term.list_comb (mk_bnf_map domTs B1Ts, p1s);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1076
            val bnf_map_app_p2 = Term.list_comb (mk_bnf_map domTs B2Ts, p2s);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1077
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1078
            val concl = mk_wpull (mk_in Xs bnf_sets_CX CX)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1079
              (mk_in B1s bnf_sets_CB1 CB1) (mk_in B2s bnf_sets_CB2 CB2)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1080
              bnf_map_app_f1 bnf_map_app_f2 (SOME (bnf_map_app_e1, bnf_map_app_e2))
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1081
              bnf_map_app_p1 bnf_map_app_p2;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1082
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1083
            val goal =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1084
              fold_rev Logic.all (Xs @ B1s @ B2s @ f1s @ f2s @ e1s @ e2s @ p1s @ p2s)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1085
                (Logic.list_implies (prems, HOLogic.mk_Trueprop concl))
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1086
          in
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 49631
diff changeset
  1087
            Goal.prove_sorry lthy [] [] goal
53270
c8628119d18e renamed BNF axiom
blanchet
parents: 53265
diff changeset
  1088
              (fn _ => mk_map_wppull_tac (#map_id0 axioms) (#map_cong0 axioms)
53290
b6c3be868217 renamed BNF fact
blanchet
parents: 53289
diff changeset
  1089
                (#map_wpull axioms) (Lazy.force map_comp) (map Lazy.force set_map))
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49103
diff changeset
  1090
            |> Thm.close_derivation
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1091
          end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1092
51762
219a3063ed29 derive "map_cong"
blanchet
parents: 51761
diff changeset
  1093
        val map_wppull = Lazy.lazy mk_map_wppull;
49453
ff0e540d8758 add rel as first-class citizen of BNF
blanchet
parents: 49452
diff changeset
  1094
53561
92bcac4f9ac9 simplified derivation of in_rel
traytel
parents: 53290
diff changeset
  1095
        val rel_OO_Grp = #rel_OO_Grp axioms;
92bcac4f9ac9 simplified derivation of in_rel
traytel
parents: 53290
diff changeset
  1096
        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
  1097
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
  1098
        fun mk_rel_Grp () =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1099
          let
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
  1100
            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
  1101
            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
  1102
            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
  1103
          in
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 49631
diff changeset
  1104
            Goal.prove_sorry lthy [] [] goal
53285
f09645642984 renamed BNF fact
blanchet
parents: 53270
diff changeset
  1105
              (mk_rel_Grp_tac rel_OO_Grps (#map_id0 axioms) (#map_cong0 axioms) (Lazy.force map_id)
53290
b6c3be868217 renamed BNF fact
blanchet
parents: 53289
diff changeset
  1106
                (Lazy.force map_comp) (map Lazy.force set_map))
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49103
diff changeset
  1107
            |> Thm.close_derivation
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1108
          end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1109
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
  1110
        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
  1111
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
  1112
        fun mk_rel_prems f = map2 (HOLogic.mk_Trueprop oo f) Rs Rs_copy
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
  1113
        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
  1114
          (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
  1115
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
  1116
        fun mk_rel_mono () =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1117
          let
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
  1118
            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
  1119
            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
  1120
          in
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 49631
diff changeset
  1121
            Goal.prove_sorry lthy [] []
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1122
              (fold_rev Logic.all (Rs @ Rs_copy) (Logic.list_implies (mono_prems, mono_concl)))
52844
66fa0f602cc4 more robust tactics (don't use unfolding when RHS might contain schematics not contained on the LHS)
traytel
parents: 52813
diff changeset
  1123
              (K (mk_rel_mono_tac rel_OO_Grps (Lazy.force in_mono)))
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49103
diff changeset
  1124
            |> Thm.close_derivation
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1125
          end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1126
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
  1127
        fun mk_rel_cong () =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1128
          let
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
  1129
            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
  1130
            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
  1131
          in
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 49631
diff changeset
  1132
            Goal.prove_sorry lthy [] []
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1133
              (fold_rev Logic.all (Rs @ Rs_copy) (Logic.list_implies (cong_prems, cong_concl)))
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 51790
diff changeset
  1134
              (fn _ => (TRY o hyp_subst_tac lthy THEN' rtac refl) 1)
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49103
diff changeset
  1135
            |> Thm.close_derivation
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1136
          end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1137
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
  1138
        val rel_mono = Lazy.lazy mk_rel_mono;
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
  1139
        val rel_cong = Lazy.lazy mk_rel_cong;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1140
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
  1141
        fun mk_rel_eq () =
52719
480a3479fa47 transfer rule for map (not yet registered as a transfer rule)
traytel
parents: 52635
diff changeset
  1142
          Goal.prove_sorry lthy [] []
480a3479fa47 transfer rule for map (not yet registered as a transfer rule)
traytel
parents: 52635
diff changeset
  1143
            (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
  1144
              HOLogic.eq_const CA'))
53270
c8628119d18e renamed BNF axiom
blanchet
parents: 53265
diff changeset
  1145
            (K (mk_rel_eq_tac live (Lazy.force rel_Grp) (Lazy.force rel_cong) (#map_id0 axioms)))
52719
480a3479fa47 transfer rule for map (not yet registered as a transfer rule)
traytel
parents: 52635
diff changeset
  1146
          |> Thm.close_derivation;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1147
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
  1148
        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
  1149
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
  1150
        fun mk_rel_conversep () =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1151
          let
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
  1152
            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
  1153
            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
  1154
            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
  1155
            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
  1156
            val le_thm = Goal.prove_sorry lthy [] [] le_goal
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
  1157
              (mk_rel_conversep_le_tac rel_OO_Grps (Lazy.force rel_eq) (#map_cong0 axioms)
53290
b6c3be868217 renamed BNF fact
blanchet
parents: 53289
diff changeset
  1158
                (Lazy.force map_comp) (map Lazy.force set_map))
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49103
diff changeset
  1159
              |> Thm.close_derivation
49123
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents: 49111
diff changeset
  1160
            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
  1161
          in
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
  1162
            Goal.prove_sorry lthy [] [] goal
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
  1163
              (K (mk_rel_conversep_tac le_thm (Lazy.force rel_mono)))
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49103
diff changeset
  1164
            |> Thm.close_derivation
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1165
          end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1166
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
  1167
        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
  1168
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
  1169
        fun mk_rel_OO () =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1170
          let
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
  1171
            val relAsCs = mk_bnf_rel pred2RTsAsCs CA' CC';
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
  1172
            val relBsCs = mk_bnf_rel pred2RTsBsCs CB' CC';
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
  1173
            val lhs = Term.list_comb (relAsCs, map2 (curry mk_rel_compp) Rs Ss);
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
  1174
            val rhs = mk_rel_compp (Term.list_comb (rel, Rs), Term.list_comb (relBsCs, Ss));
49123
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents: 49111
diff changeset
  1175
            val goal = fold_rev Logic.all (Rs @ Ss) (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
  1176
          in
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 49631
diff changeset
  1177
            Goal.prove_sorry lthy [] [] goal
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
  1178
              (mk_rel_OO_tac rel_OO_Grps (Lazy.force rel_eq) (#map_cong0 axioms)
53290
b6c3be868217 renamed BNF fact
blanchet
parents: 53289
diff changeset
  1179
                (Lazy.force map_wppull) (Lazy.force map_comp) (map Lazy.force set_map))
49109
0e5b859e1c91 no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents: 49103
diff changeset
  1180
            |> Thm.close_derivation
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1181
          end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1182
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
  1183
        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
  1184
53561
92bcac4f9ac9 simplified derivation of in_rel
traytel
parents: 53290
diff changeset
  1185
        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
  1186
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
  1187
        val in_rel = Lazy.lazy mk_in_rel;
49537
fe1deee434b6 generate "rel_as_srel" and "rel_flip" properties
blanchet
parents: 49536
diff changeset
  1188
fe1deee434b6 generate "rel_as_srel" and "rel_flip" properties
blanchet
parents: 49536
diff changeset
  1189
        fun mk_rel_flip () =
fe1deee434b6 generate "rel_as_srel" and "rel_flip" properties
blanchet
parents: 49536
diff changeset
  1190
          let
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
  1191
            val rel_conversep_thm = Lazy.force rel_conversep;
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
  1192
            val cts = map (SOME o certify lthy) Rs;
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
  1193
            val rel_conversep_thm' = cterm_instantiate_pos cts rel_conversep_thm;
49537
fe1deee434b6 generate "rel_as_srel" and "rel_flip" properties
blanchet
parents: 49536
diff changeset
  1194
          in
51917
f964a9887713 store proper theorems even for fixed points that have no passive live variables
traytel
parents: 51916
diff changeset
  1195
            unfold_thms lthy @{thms conversep_iff} (rel_conversep_thm' RS @{thm predicate2_eqD})
49595
e8c57e59cbf8 got rid of other instance of shaky "Thm.generalize"
blanchet
parents: 49594
diff changeset
  1196
            |> singleton (Proof_Context.export names_lthy pre_names_lthy)
49537
fe1deee434b6 generate "rel_as_srel" and "rel_flip" properties
blanchet
parents: 49536
diff changeset
  1197
          end;
fe1deee434b6 generate "rel_as_srel" and "rel_flip" properties
blanchet
parents: 49536
diff changeset
  1198
49538
c90818b63599 simplified fact policies
blanchet
parents: 49537
diff changeset
  1199
        val rel_flip = Lazy.lazy mk_rel_flip;
49537
fe1deee434b6 generate "rel_as_srel" and "rel_flip" properties
blanchet
parents: 49536
diff changeset
  1200
51916
eac9e9a45bf5 stronger monotonicity property for relators
traytel
parents: 51915
diff changeset
  1201
        fun mk_rel_mono_strong () =
eac9e9a45bf5 stronger monotonicity property for relators
traytel
parents: 51915
diff changeset
  1202
          let
eac9e9a45bf5 stronger monotonicity property for relators
traytel
parents: 51915
diff changeset
  1203
            fun mk_prem setA setB R S a b =
eac9e9a45bf5 stronger monotonicity property for relators
traytel
parents: 51915
diff changeset
  1204
              HOLogic.mk_Trueprop
eac9e9a45bf5 stronger monotonicity property for relators
traytel
parents: 51915
diff changeset
  1205
                (mk_Ball (setA $ x) (Term.absfree (dest_Free a)
eac9e9a45bf5 stronger monotonicity property for relators
traytel
parents: 51915
diff changeset
  1206
                  (mk_Ball (setB $ y) (Term.absfree (dest_Free b)
eac9e9a45bf5 stronger monotonicity property for relators
traytel
parents: 51915
diff changeset
  1207
                    (HOLogic.mk_imp (R $ a $ b, S $ a $ b))))));
eac9e9a45bf5 stronger monotonicity property for relators
traytel
parents: 51915
diff changeset
  1208
            val prems = HOLogic.mk_Trueprop (Term.list_comb (rel, Rs) $ x $ y) :: 
eac9e9a45bf5 stronger monotonicity property for relators
traytel
parents: 51915
diff changeset
  1209
              map6 mk_prem bnf_sets_As bnf_sets_Bs Rs Rs_copy zs ys;
eac9e9a45bf5 stronger monotonicity property for relators
traytel
parents: 51915
diff changeset
  1210
            val concl = HOLogic.mk_Trueprop (Term.list_comb (rel, Rs_copy) $ x $ y);
eac9e9a45bf5 stronger monotonicity property for relators
traytel
parents: 51915
diff changeset
  1211
          in
eac9e9a45bf5 stronger monotonicity property for relators
traytel
parents: 51915
diff changeset
  1212
            Goal.prove_sorry lthy [] []
eac9e9a45bf5 stronger monotonicity property for relators
traytel
parents: 51915
diff changeset
  1213
              (fold_rev Logic.all (x :: y :: Rs @ Rs_copy) (Logic.list_implies (prems, concl)))
53290
b6c3be868217 renamed BNF fact
blanchet
parents: 53289
diff changeset
  1214
              (mk_rel_mono_strong_tac (Lazy.force in_rel) (map Lazy.force set_map))
51916
eac9e9a45bf5 stronger monotonicity property for relators
traytel
parents: 51915
diff changeset
  1215
            |> Thm.close_derivation
eac9e9a45bf5 stronger monotonicity property for relators
traytel
parents: 51915
diff changeset
  1216
          end;
eac9e9a45bf5 stronger monotonicity property for relators
traytel
parents: 51915
diff changeset
  1217
eac9e9a45bf5 stronger monotonicity property for relators
traytel
parents: 51915
diff changeset
  1218
        val rel_mono_strong = Lazy.lazy mk_rel_mono_strong;
eac9e9a45bf5 stronger monotonicity property for relators
traytel
parents: 51915
diff changeset
  1219
52719
480a3479fa47 transfer rule for map (not yet registered as a transfer rule)
traytel
parents: 52635
diff changeset
  1220
        fun mk_map_transfer () =
480a3479fa47 transfer rule for map (not yet registered as a transfer rule)
traytel
parents: 52635
diff changeset
  1221
          let
52725
ba2bbe825a5e proper transfer rule format for map_transfer
traytel
parents: 52721
diff changeset
  1222
            val rels = map2 mk_fun_rel transfer_domRs transfer_ranRs;
ba2bbe825a5e proper transfer rule format for map_transfer
traytel
parents: 52721
diff changeset
  1223
            val rel = mk_fun_rel
ba2bbe825a5e proper transfer rule format for map_transfer
traytel
parents: 52721
diff changeset
  1224
              (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
  1225
              (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
  1226
            val concl = HOLogic.mk_Trueprop
ba2bbe825a5e proper transfer rule format for map_transfer
traytel
parents: 52721
diff changeset
  1227
              (fold_rev mk_fun_rel 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
  1228
          in
480a3479fa47 transfer rule for map (not yet registered as a transfer rule)
traytel
parents: 52635
diff changeset
  1229
            Goal.prove_sorry lthy [] []
52725
ba2bbe825a5e proper transfer rule format for map_transfer
traytel
parents: 52721
diff changeset
  1230
              (fold_rev Logic.all (transfer_domRs @ transfer_ranRs) concl)
52719
480a3479fa47 transfer rule for map (not yet registered as a transfer rule)
traytel
parents: 52635
diff changeset
  1231
              (mk_map_transfer_tac (Lazy.force rel_mono) (Lazy.force in_rel)
53290
b6c3be868217 renamed BNF fact
blanchet
parents: 53289
diff changeset
  1232
                (map Lazy.force set_map) (#map_cong0 axioms) (Lazy.force map_comp))
52719
480a3479fa47 transfer rule for map (not yet registered as a transfer rule)
traytel
parents: 52635
diff changeset
  1233
            |> Thm.close_derivation
480a3479fa47 transfer rule for map (not yet registered as a transfer rule)
traytel
parents: 52635
diff changeset
  1234
          end;
480a3479fa47 transfer rule for map (not yet registered as a transfer rule)
traytel
parents: 52635
diff changeset
  1235
480a3479fa47 transfer rule for map (not yet registered as a transfer rule)
traytel
parents: 52635
diff changeset
  1236
        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
  1237
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
  1238
        val defs = mk_defs bnf_map_def bnf_set_defs bnf_rel_def;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1239
52635
4f84b730c489 got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
traytel
parents: 51930
diff changeset
  1240
        val facts = mk_facts bd_Card_order bd_Cinfinite bd_Cnotzero collect_set_map in_bd in_cong
53290
b6c3be868217 renamed BNF fact
blanchet
parents: 53289
diff changeset
  1241
          in_mono in_rel map_comp map_cong map_id map_transfer map_wppull rel_eq rel_flip set_map
51916
eac9e9a45bf5 stronger monotonicity property for relators
traytel
parents: 51915
diff changeset
  1242
          rel_cong rel_mono rel_mono_strong rel_Grp rel_conversep rel_OO;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1243
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1244
        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
  1245
49507
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49506
diff changeset
  1246
        val bnf_rel =
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49506
diff changeset
  1247
          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
  1248
53265
cc9a2976f836 qualify BNF constants properly
blanchet
parents: 53261
diff changeset
  1249
        val bnf = mk_bnf bnf_b CA live alphas betas dead deads bnf_map bnf_sets bnf_bd axioms defs
cc9a2976f836 qualify BNF constants properly
blanchet
parents: 53261
diff changeset
  1250
          facts wits bnf_rel;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1251
      in
53265
cc9a2976f836 qualify BNF constants properly
blanchet
parents: 53261
diff changeset
  1252
        (bnf, lthy |> note_bnf_thms fact_policy qualify bnf_b bnf)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1253
      end;
49459
3f8e2b5249ec adapted FP code to new relator approach
blanchet
parents: 49458
diff changeset
  1254
3f8e2b5249ec adapted FP code to new relator approach
blanchet
parents: 49458
diff changeset
  1255
    val one_step_defs =
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51837
diff changeset
  1256
      no_reflexive (bnf_map_def :: bnf_bd_def :: bnf_set_defs @ bnf_wit_defs @ [bnf_rel_def]);
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1257
  in
49459
3f8e2b5249ec adapted FP code to new relator approach
blanchet
parents: 49458
diff changeset
  1258
    (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
  1259
  end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1260
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
  1261
fun register_bnf key (bnf, lthy) =
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
  1262
  (bnf, Local_Theory.declaration {syntax = false, pervasive = true}
52957
717a23e14586 avoid DUP exception in local context (cf. 062aa11e98e1)
blanchet
parents: 52923
diff changeset
  1263
    (fn phi => Data.map (Symtab.default (key, morph_bnf phi bnf))) lthy);
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
  1264
49456
fa8302c8dea1 adapted BNF composition to new relator approach
blanchet
parents: 49455
diff changeset
  1265
(* TODO: Once the invariant "nwits > 0" holds, remove "mk_conjunction_balanced'" and "rtac TrueI"
fa8302c8dea1 adapted BNF composition to new relator approach
blanchet
parents: 49455
diff changeset
  1266
   below *)
fa8302c8dea1 adapted BNF composition to new relator approach
blanchet
parents: 49455
diff changeset
  1267
fun mk_conjunction_balanced' [] = @{prop True}
fa8302c8dea1 adapted BNF composition to new relator approach
blanchet
parents: 49455
diff changeset
  1268
  | mk_conjunction_balanced' ts = Logic.mk_conjunction_balanced ts;
fa8302c8dea1 adapted BNF composition to new relator approach
blanchet
parents: 49455
diff changeset
  1269
51767
bbcdd8519253 honor user-specified name for relator + generalize syntax
blanchet
parents: 51766
diff changeset
  1270
fun bnf_def const_policy fact_policy qualify tacs wit_tac Ds map_b rel_b set_bs =
49463
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49462
diff changeset
  1271
  (fn (_, 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
  1272
  let
49456
fa8302c8dea1 adapted BNF composition to new relator approach
blanchet
parents: 49455
diff changeset
  1273
    val wits_tac =
fa8302c8dea1 adapted BNF composition to new relator approach
blanchet
parents: 49455
diff changeset
  1274
      K (TRYALL Goal.conjunction_tac) THEN' K (TRYALL (rtac TrueI)) THEN'
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49495
diff changeset
  1275
      mk_unfold_thms_then_tac lthy one_step_defs wit_tac;
49456
fa8302c8dea1 adapted BNF composition to new relator approach
blanchet
parents: 49455
diff changeset
  1276
    val wit_goals = map mk_conjunction_balanced' wit_goalss;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1277
    val wit_thms =
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 49631
diff changeset
  1278
      Goal.prove_sorry lthy [] [] (mk_conjunction_balanced' wit_goals) wits_tac
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1279
      |> Conjunction.elim_balanced (length wit_goals)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1280
      |> map2 (Conjunction.elim_balanced o length) wit_goalss
49456
fa8302c8dea1 adapted BNF composition to new relator approach
blanchet
parents: 49455
diff changeset
  1281
      |> map (map (Thm.close_derivation o 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
  1282
  in
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 49631
diff changeset
  1283
    map2 (Thm.close_derivation oo Goal.prove_sorry lthy [] [])
49504
df9b897fb254 renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents: 49495
diff changeset
  1284
      goals (map (mk_unfold_thms_then_tac lthy one_step_defs) tacs)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1285
    |> (fn thms => after_qed (map single thms @ wit_thms) lthy)
51767
bbcdd8519253 honor user-specified name for relator + generalize syntax
blanchet
parents: 51766
diff changeset
  1286
  end) oo prepare_def const_policy fact_policy qualify (K I) Ds map_b rel_b set_bs;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1287
51836
4d6dcd51dd52 renamed "bnf_def" keyword to "bnf" (since it's not a definition, but rather a registration)
blanchet
parents: 51823
diff changeset
  1288
val bnf_cmd = (fn (key, goals, wit_goals, after_qed, lthy, defs) =>
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1289
  Proof.unfolding ([[(defs, [])]])
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
  1290
    (Proof.theorem NONE (snd o register_bnf key oo after_qed)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1291
      (map (single o rpair []) goals @ map (map (rpair [])) wit_goals) lthy)) oo
51767
bbcdd8519253 honor user-specified name for relator + generalize syntax
blanchet
parents: 51766
diff changeset
  1292
  prepare_def Do_Inline (user_policy Note_Some) I Syntax.read_term NONE Binding.empty Binding.empty
bbcdd8519253 honor user-specified name for relator + generalize syntax
blanchet
parents: 51766
diff changeset
  1293
    [];
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1294
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1295
fun print_bnfs ctxt =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1296
  let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1297
    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
  1298
      [Pretty.str (mk_setN (i + 1) ^ ":"), Pretty.brk 1,
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1299
          Pretty.quote (Syntax.pretty_term ctxt (nth sets i))];
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1300
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1301
    fun pretty_bnf (key, BNF {T = T, 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
  1302
      live = live, 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
  1303
      Pretty.big_list
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1304
        (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
  1305
          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
  1306
        ([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
  1307
            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
  1308
          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
  1309
            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
  1310
          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
  1311
            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
  1312
          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
  1313
          [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
  1314
            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
  1315
  in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1316
    Pretty.big_list "BNFs:" (map pretty_bnf (Symtab.dest (Data.get (Context.Proof ctxt))))
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1317
    |> Pretty.writeln
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1318
  end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1319
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1320
val _ =
51836
4d6dcd51dd52 renamed "bnf_def" keyword to "bnf" (since it's not a definition, but rather a registration)
blanchet
parents: 51823
diff changeset
  1321
  Outer_Syntax.improper_command @{command_spec "print_bnfs"}
53289
5e0623448bdb renamed BNF axiom
blanchet
parents: 53288
diff changeset
  1322
    "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
  1323
    (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
  1324
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1325
val _ =
51836
4d6dcd51dd52 renamed "bnf_def" keyword to "bnf" (since it's not a definition, but rather a registration)
blanchet
parents: 51823
diff changeset
  1326
  Outer_Syntax.local_theory_to_proof @{command_spec "bnf"}
53289
5e0623448bdb renamed BNF axiom
blanchet
parents: 53288
diff changeset
  1327
    "register a type as a bounded natural functor"
51790
22517d04d20b more intuitive syntax for equality-style discriminators of nullary constructors
blanchet
parents: 51787
diff changeset
  1328
    ((parse_opt_binding_colon -- Parse.term --
49277
blanchet
parents: 49236
diff changeset
  1329
       (@{keyword "["} |-- Parse.list Parse.term --| @{keyword "]"}) -- Parse.term --
49459
3f8e2b5249ec adapted FP code to new relator approach
blanchet
parents: 49458
diff changeset
  1330
       (@{keyword "["} |-- Parse.list Parse.term --| @{keyword "]"}) -- Scan.option Parse.term)
51836
4d6dcd51dd52 renamed "bnf_def" keyword to "bnf" (since it's not a definition, but rather a registration)
blanchet
parents: 51823
diff changeset
  1331
       >> bnf_cmd);
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1332
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
  1333
end;