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