src/HOL/Codatatype/Tools/bnf_fp_util.ML
author blanchet
Sat, 08 Sep 2012 22:54:37 +0200
changeset 49226 510c6d4a73ec
parent 49225 a9295b1f401b
child 49228 e43910ccee74
permissions -rw-r--r--
fixed and enabled iterator/recursor theorems
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     1
(*  Title:      HOL/Codatatype/Tools/bnf_fp_util.ML
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
    Copyright   2012
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     4
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     5
Shared library for the datatype and the codatatype construction.
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     6
*)
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
signature BNF_FP_UTIL =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     9
sig
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    10
  val time: Timer.real_timer -> string -> Timer.real_timer
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    11
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    12
  val IITN: string
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    13
  val LevN: string
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    14
  val algN: string
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    15
  val behN: string
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    16
  val bisN: string
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    17
  val carTN: string
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    18
  val coN: string
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    19
  val coinductN: string
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    20
  val coiterN: string
49128
1a86ef0a0210 renamed low-level (co)iterators and (co)recursors with "fld_"/"unf_" prefix
blanchet
parents: 49125
diff changeset
    21
  val unf_coiter_uniqueN: string
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    22
  val corecN: string
49020
f379cf5d71bd more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents: 49019
diff changeset
    23
  val exhaustN: string
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    24
  val fldN: string
49225
a9295b1f401b renamed for consistency
blanchet
parents: 49223
diff changeset
    25
  val fld_unf_coitersN: string
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    26
  val fld_exhaustN: string
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    27
  val fld_induct2N: string
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    28
  val fld_inductN: string
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    29
  val fld_injectN: string
49128
1a86ef0a0210 renamed low-level (co)iterators and (co)recursors with "fld_"/"unf_" prefix
blanchet
parents: 49125
diff changeset
    30
  val fld_iterN: string
49222
cbe8c859817c for compatiblity with old datatype package: not only "recs" with "s", but also "iters" and their "fld_"/"unf_" variants
blanchet
parents: 49218
diff changeset
    31
  val fld_itersN: string
49128
1a86ef0a0210 renamed low-level (co)iterators and (co)recursors with "fld_"/"unf_" prefix
blanchet
parents: 49125
diff changeset
    32
  val fld_recN: string
49222
cbe8c859817c for compatiblity with old datatype package: not only "recs" with "s", but also "iters" and their "fld_"/"unf_" variants
blanchet
parents: 49218
diff changeset
    33
  val fld_recsN: string
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    34
  val fld_unfN: string
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    35
  val hsetN: string
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    36
  val hset_recN: string
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    37
  val inductN: string
49019
fc4decdba5ce more work on BNF sugar
blanchet
parents: 48975
diff changeset
    38
  val injectN: string
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    39
  val isNodeN: string
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    40
  val iterN: string
49128
1a86ef0a0210 renamed low-level (co)iterators and (co)recursors with "fld_"/"unf_" prefix
blanchet
parents: 49125
diff changeset
    41
  val fld_iter_uniqueN: string
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    42
  val lsbisN: string
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    43
  val map_simpsN: string
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    44
  val map_uniqueN: string
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    45
  val min_algN: string
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    46
  val morN: string
49020
f379cf5d71bd more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents: 49019
diff changeset
    47
  val nchotomyN: string
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    48
  val pred_coinductN: string
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    49
  val pred_coinduct_uptoN: string
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    50
  val recN: string
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    51
  val rel_coinductN: string
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    52
  val rel_coinduct_uptoN: string
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    53
  val rvN: string
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    54
  val set_inclN: string
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    55
  val set_set_inclN: string
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    56
  val strTN: string
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    57
  val str_initN: string
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    58
  val sum_bdN: string
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    59
  val sum_bdTN: string
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    60
  val unfN: string
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    61
  val unf_coinductN: string
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    62
  val unf_coinduct_uptoN: string
49128
1a86ef0a0210 renamed low-level (co)iterators and (co)recursors with "fld_"/"unf_" prefix
blanchet
parents: 49125
diff changeset
    63
  val unf_coiterN: string
49222
cbe8c859817c for compatiblity with old datatype package: not only "recs" with "s", but also "iters" and their "fld_"/"unf_" variants
blanchet
parents: 49218
diff changeset
    64
  val unf_coitersN: string
49128
1a86ef0a0210 renamed low-level (co)iterators and (co)recursors with "fld_"/"unf_" prefix
blanchet
parents: 49125
diff changeset
    65
  val unf_corecN: string
49222
cbe8c859817c for compatiblity with old datatype package: not only "recs" with "s", but also "iters" and their "fld_"/"unf_" variants
blanchet
parents: 49218
diff changeset
    66
  val unf_corecsN: string
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    67
  val unf_exhaustN: string
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    68
  val unf_fldN: string
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    69
  val unf_injectN: string
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    70
  val uniqueN: string
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    71
  val uptoN: string
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    72
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    73
  val mk_exhaustN: string -> string
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    74
  val mk_injectN: string -> string
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    75
  val mk_nchotomyN: string -> string
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    76
  val mk_set_simpsN: int -> string
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    77
  val mk_set_minimalN: int -> string
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    78
  val mk_set_inductN: int -> string
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    79
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    80
  val typedef: bool -> binding option -> binding * (string * sort) list * mixfix -> term ->
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    81
    (binding * binding) option -> tactic -> local_theory -> (string * Typedef.info) * local_theory
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    82
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    83
  val split_conj_thm: thm -> thm list
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    84
  val split_conj_prems: int -> thm -> thm
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    85
49121
9e0acaa470ab more work on FP sugar
blanchet
parents: 49119
diff changeset
    86
  val Inl_const: typ -> typ -> term
9e0acaa470ab more work on FP sugar
blanchet
parents: 49119
diff changeset
    87
  val Inr_const: typ -> typ -> term
9e0acaa470ab more work on FP sugar
blanchet
parents: 49119
diff changeset
    88
9e0acaa470ab more work on FP sugar
blanchet
parents: 49119
diff changeset
    89
  val mk_Inl: term -> typ -> term
9e0acaa470ab more work on FP sugar
blanchet
parents: 49119
diff changeset
    90
  val mk_Inr: term -> typ -> term
9e0acaa470ab more work on FP sugar
blanchet
parents: 49119
diff changeset
    91
  val mk_InN: typ list -> term -> int -> term
49129
b5413cb7d860 define "case" constant
blanchet
parents: 49128
diff changeset
    92
  val mk_sum_case: term -> term -> term
b5413cb7d860 define "case" constant
blanchet
parents: 49128
diff changeset
    93
  val mk_sum_caseN: term list -> term
49121
9e0acaa470ab more work on FP sugar
blanchet
parents: 49119
diff changeset
    94
49176
6d29d2db5f88 construct high-level iterator RHS
blanchet
parents: 49169
diff changeset
    95
  val dest_sumTN: int -> typ -> typ list
6d29d2db5f88 construct high-level iterator RHS
blanchet
parents: 49169
diff changeset
    96
  val dest_tupleT: int -> typ -> typ list
6d29d2db5f88 construct high-level iterator RHS
blanchet
parents: 49169
diff changeset
    97
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    98
  val mk_Field: term -> term
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    99
  val mk_union: term * term -> term
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   100
49125
5fc5211cf104 implemented "mk_exhaust_tac"
blanchet
parents: 49121
diff changeset
   101
  val mk_sumEN: int -> thm
49130
3c26e17b2849 implemented "mk_case_tac" -- and got rid of "cheat_tac"
blanchet
parents: 49129
diff changeset
   102
  val mk_sum_casesN: int -> int -> thm
49125
5fc5211cf104 implemented "mk_exhaust_tac"
blanchet
parents: 49121
diff changeset
   103
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   104
  val mk_tactics: 'a -> 'a -> 'a -> 'a list -> 'a -> 'a -> 'a list -> 'a -> 'a -> 'a list
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   105
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   106
  val fixpoint: ('a * 'a -> bool) -> ('a list -> 'a list) -> 'a list -> 'a list
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   107
49185
073d7d1b7488 respect order of/additional type variables supplied by the user in fixed point constructions;
traytel
parents: 49176
diff changeset
   108
  val fp_bnf: (mixfix list -> (string * sort) list option -> binding list ->
073d7d1b7488 respect order of/additional type variables supplied by the user in fixed point constructions;
traytel
parents: 49176
diff changeset
   109
    typ list * typ list list -> BNF_Def.BNF list -> local_theory -> 'a) ->
49169
937a0fadddfb honor mixfix specifications
blanchet
parents: 49156
diff changeset
   110
    binding list -> mixfix list -> (string * sort) list -> ((string * sort) * typ) list ->
49226
510c6d4a73ec fixed and enabled iterator/recursor theorems
blanchet
parents: 49225
diff changeset
   111
    local_theory -> BNF_Def.BNF list * 'a
49185
073d7d1b7488 respect order of/additional type variables supplied by the user in fixed point constructions;
traytel
parents: 49176
diff changeset
   112
  val fp_bnf_cmd: (mixfix list -> (string * sort) list option -> binding list ->
073d7d1b7488 respect order of/additional type variables supplied by the user in fixed point constructions;
traytel
parents: 49176
diff changeset
   113
    typ list * typ list list -> BNF_Def.BNF list -> local_theory -> 'a) ->
49134
846264f80f16 optionally provide extra dead variables to the FP constructions
blanchet
parents: 49132
diff changeset
   114
    binding list * (string list * string list) -> local_theory -> 'a
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   115
end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   116
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   117
structure BNF_FP_Util : BNF_FP_UTIL =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   118
struct
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
open BNF_Comp
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   121
open BNF_Def
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   122
open BNF_Util
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   123
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   124
val timing = true;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   125
fun time timer msg = (if timing
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   126
  then warning (msg ^ ": " ^ ATP_Util.string_from_time (Timer.checkRealTimer timer))
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   127
  else (); Timer.startRealTimer ());
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   128
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   129
(*TODO: is this really different from Typedef.add_typedef_global?*)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   130
fun typedef def opt_name typ set opt_morphs tac lthy =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   131
  let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   132
    val ((name, info), (lthy, lthy_old)) =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   133
      lthy
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   134
      |> Typedef.add_typedef def opt_name typ set opt_morphs tac
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   135
      ||> `Local_Theory.restore;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   136
    val phi = Proof_Context.export_morphism lthy_old lthy;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   137
  in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   138
    ((name, Typedef.transform_info phi info), lthy)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   139
  end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   140
49223
blanchet
parents: 49222
diff changeset
   141
val preN = "pre_"
blanchet
parents: 49222
diff changeset
   142
val rawN = "raw_"
49218
d01a5c918298 renamed xxxBNF to pre_xxx
blanchet
parents: 49207
diff changeset
   143
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   144
val coN = "co"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   145
val algN = "alg"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   146
val IITN = "IITN"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   147
val iterN = "iter"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   148
val coiterN = coN ^ iterN
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   149
val uniqueN = "_unique"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   150
val fldN = "fld"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   151
val unfN = "unf"
49128
1a86ef0a0210 renamed low-level (co)iterators and (co)recursors with "fld_"/"unf_" prefix
blanchet
parents: 49125
diff changeset
   152
val fld_iterN = fldN ^ "_" ^ iterN
49222
cbe8c859817c for compatiblity with old datatype package: not only "recs" with "s", but also "iters" and their "fld_"/"unf_" variants
blanchet
parents: 49218
diff changeset
   153
val fld_itersN = fld_iterN ^ "s"
49128
1a86ef0a0210 renamed low-level (co)iterators and (co)recursors with "fld_"/"unf_" prefix
blanchet
parents: 49125
diff changeset
   154
val unf_coiterN = unfN ^ "_" ^ coiterN
49222
cbe8c859817c for compatiblity with old datatype package: not only "recs" with "s", but also "iters" and their "fld_"/"unf_" variants
blanchet
parents: 49218
diff changeset
   155
val unf_coitersN = unf_coiterN ^ "s"
49128
1a86ef0a0210 renamed low-level (co)iterators and (co)recursors with "fld_"/"unf_" prefix
blanchet
parents: 49125
diff changeset
   156
val fld_iter_uniqueN = fld_iterN ^ uniqueN
1a86ef0a0210 renamed low-level (co)iterators and (co)recursors with "fld_"/"unf_" prefix
blanchet
parents: 49125
diff changeset
   157
val unf_coiter_uniqueN = unf_coiterN ^ uniqueN
49225
a9295b1f401b renamed for consistency
blanchet
parents: 49223
diff changeset
   158
val fld_unf_coitersN = fldN ^ "_" ^ unf_coiterN ^ "s"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   159
val map_simpsN = mapN ^ "_simps"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   160
val map_uniqueN = mapN ^ uniqueN
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   161
val min_algN = "min_alg"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   162
val morN = "mor"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   163
val bisN = "bis"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   164
val lsbisN = "lsbis"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   165
val sum_bdTN = "sbdT"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   166
val sum_bdN = "sbd"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   167
val carTN = "carT"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   168
val strTN = "strT"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   169
val isNodeN = "isNode"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   170
val LevN = "Lev"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   171
val rvN = "recover"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   172
val behN = "beh"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   173
fun mk_set_simpsN i = mk_setN i ^ "_simps"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   174
fun mk_set_minimalN i = mk_setN i ^ "_minimal"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   175
fun mk_set_inductN i = mk_setN i ^ "_induct"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   176
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   177
val str_initN = "str_init"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   178
val recN = "rec"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   179
val corecN = coN ^ recN
49128
1a86ef0a0210 renamed low-level (co)iterators and (co)recursors with "fld_"/"unf_" prefix
blanchet
parents: 49125
diff changeset
   180
val fld_recN = fldN ^ "_" ^ recN
49222
cbe8c859817c for compatiblity with old datatype package: not only "recs" with "s", but also "iters" and their "fld_"/"unf_" variants
blanchet
parents: 49218
diff changeset
   181
val fld_recsN = fld_recN ^ "s"
49128
1a86ef0a0210 renamed low-level (co)iterators and (co)recursors with "fld_"/"unf_" prefix
blanchet
parents: 49125
diff changeset
   182
val unf_corecN = unfN ^ "_" ^ corecN
49222
cbe8c859817c for compatiblity with old datatype package: not only "recs" with "s", but also "iters" and their "fld_"/"unf_" variants
blanchet
parents: 49218
diff changeset
   183
val unf_corecsN = unf_corecN ^ "s"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   184
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   185
val fld_unfN = fldN ^ "_" ^ unfN
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   186
val unf_fldN = unfN ^ "_" ^ fldN
49020
f379cf5d71bd more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents: 49019
diff changeset
   187
val nchotomyN = "nchotomy"
f379cf5d71bd more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents: 49019
diff changeset
   188
fun mk_nchotomyN s = s ^ "_" ^ nchotomyN
49019
fc4decdba5ce more work on BNF sugar
blanchet
parents: 48975
diff changeset
   189
val injectN = "inject"
fc4decdba5ce more work on BNF sugar
blanchet
parents: 48975
diff changeset
   190
fun mk_injectN s = s ^ "_" ^ injectN
49020
f379cf5d71bd more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents: 49019
diff changeset
   191
val exhaustN = "exhaust"
f379cf5d71bd more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents: 49019
diff changeset
   192
fun mk_exhaustN s = s ^ "_" ^ exhaustN
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   193
val fld_injectN = mk_injectN fldN
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   194
val fld_exhaustN = mk_exhaustN fldN
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   195
val unf_injectN = mk_injectN unfN
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   196
val unf_exhaustN = mk_exhaustN unfN
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   197
val inductN = "induct"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   198
val coinductN = coN ^ inductN
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   199
val fld_inductN = fldN ^ "_" ^ inductN
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   200
val fld_induct2N = fld_inductN ^ "2"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   201
val unf_coinductN = unfN ^ "_" ^ coinductN
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   202
val rel_coinductN = relN ^ "_" ^ coinductN
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   203
val pred_coinductN = predN ^ "_" ^ coinductN
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   204
val uptoN = "upto"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   205
val unf_coinduct_uptoN = unf_coinductN ^ "_" ^ uptoN
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   206
val rel_coinduct_uptoN = rel_coinductN ^ "_" ^ uptoN
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   207
val pred_coinduct_uptoN = pred_coinductN ^ "_" ^ uptoN
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   208
val hsetN = "Hset"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   209
val hset_recN = hsetN ^ "_rec"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   210
val set_inclN = "set_incl"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   211
val set_set_inclN = "set_set_incl"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   212
49121
9e0acaa470ab more work on FP sugar
blanchet
parents: 49119
diff changeset
   213
fun Inl_const LT RT = Const (@{const_name Inl}, LT --> mk_sumT (LT, RT));
9e0acaa470ab more work on FP sugar
blanchet
parents: 49119
diff changeset
   214
fun mk_Inl t RT = Inl_const (fastype_of t) RT $ t;
9e0acaa470ab more work on FP sugar
blanchet
parents: 49119
diff changeset
   215
9e0acaa470ab more work on FP sugar
blanchet
parents: 49119
diff changeset
   216
fun Inr_const LT RT = Const (@{const_name Inr}, RT --> mk_sumT (LT, RT));
9e0acaa470ab more work on FP sugar
blanchet
parents: 49119
diff changeset
   217
fun mk_Inr t LT = Inr_const LT (fastype_of t) $ t;
9e0acaa470ab more work on FP sugar
blanchet
parents: 49119
diff changeset
   218
9e0acaa470ab more work on FP sugar
blanchet
parents: 49119
diff changeset
   219
fun mk_InN [_] t 1 = t
9e0acaa470ab more work on FP sugar
blanchet
parents: 49119
diff changeset
   220
  | mk_InN (_ :: Ts) t 1 = mk_Inl t (mk_sumTN Ts)
9e0acaa470ab more work on FP sugar
blanchet
parents: 49119
diff changeset
   221
  | mk_InN (LT :: Ts) t m = mk_Inr (mk_InN Ts t (m - 1)) LT
9e0acaa470ab more work on FP sugar
blanchet
parents: 49119
diff changeset
   222
  | mk_InN Ts t _ = raise (TYPE ("mk_InN", Ts, [t]));
9e0acaa470ab more work on FP sugar
blanchet
parents: 49119
diff changeset
   223
49129
b5413cb7d860 define "case" constant
blanchet
parents: 49128
diff changeset
   224
fun mk_sum_case f g =
b5413cb7d860 define "case" constant
blanchet
parents: 49128
diff changeset
   225
  let
b5413cb7d860 define "case" constant
blanchet
parents: 49128
diff changeset
   226
    val fT = fastype_of f;
b5413cb7d860 define "case" constant
blanchet
parents: 49128
diff changeset
   227
    val gT = fastype_of g;
b5413cb7d860 define "case" constant
blanchet
parents: 49128
diff changeset
   228
  in
b5413cb7d860 define "case" constant
blanchet
parents: 49128
diff changeset
   229
    Const (@{const_name sum_case},
b5413cb7d860 define "case" constant
blanchet
parents: 49128
diff changeset
   230
      fT --> gT --> mk_sumT (domain_type fT, domain_type gT) --> range_type fT) $ f $ g
b5413cb7d860 define "case" constant
blanchet
parents: 49128
diff changeset
   231
  end;
b5413cb7d860 define "case" constant
blanchet
parents: 49128
diff changeset
   232
b5413cb7d860 define "case" constant
blanchet
parents: 49128
diff changeset
   233
fun mk_sum_caseN [f] = f
b5413cb7d860 define "case" constant
blanchet
parents: 49128
diff changeset
   234
  | mk_sum_caseN (f :: fs) = mk_sum_case f (mk_sum_caseN fs);
b5413cb7d860 define "case" constant
blanchet
parents: 49128
diff changeset
   235
49176
6d29d2db5f88 construct high-level iterator RHS
blanchet
parents: 49169
diff changeset
   236
fun dest_sumTN 1 T = [T]
6d29d2db5f88 construct high-level iterator RHS
blanchet
parents: 49169
diff changeset
   237
  | dest_sumTN n (Type (@{type_name sum}, [T, T'])) = T :: dest_sumTN (n - 1) T';
6d29d2db5f88 construct high-level iterator RHS
blanchet
parents: 49169
diff changeset
   238
6d29d2db5f88 construct high-level iterator RHS
blanchet
parents: 49169
diff changeset
   239
(* TODO: move something like this to "HOLogic"? *)
6d29d2db5f88 construct high-level iterator RHS
blanchet
parents: 49169
diff changeset
   240
fun dest_tupleT 0 @{typ unit} = []
6d29d2db5f88 construct high-level iterator RHS
blanchet
parents: 49169
diff changeset
   241
  | dest_tupleT 1 T = [T]
6d29d2db5f88 construct high-level iterator RHS
blanchet
parents: 49169
diff changeset
   242
  | dest_tupleT n (Type (@{type_name prod}, [T, T'])) = T :: dest_tupleT (n - 1) T';
6d29d2db5f88 construct high-level iterator RHS
blanchet
parents: 49169
diff changeset
   243
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   244
fun mk_Field r =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   245
  let val T = fst (dest_relT (fastype_of r));
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   246
  in Const (@{const_name Field}, mk_relT (T, T) --> HOLogic.mk_setT T) $ r end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   247
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   248
val mk_union = HOLogic.mk_binop @{const_name sup};
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   249
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   250
(*dangerous; use with monotonic, converging functions only!*)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   251
fun fixpoint eq f X = if subset eq (f X, X) then X else fixpoint eq f (f X);
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
(* stolen from "~~/src/HOL/Tools/Datatype/datatype_aux.ML" *)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   254
fun split_conj_thm th =
49119
1f605c36869c more work on FP sugar
blanchet
parents: 49075
diff changeset
   255
  ((th RS conjunct1) :: split_conj_thm (th RS conjunct2)) handle THM _ => [th];
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   256
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   257
fun split_conj_prems limit th =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   258
  let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   259
    fun split n i th =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   260
      if i = n then th else split n (i + 1) (conjI RSN (i, th)) handle THM _ => th;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   261
  in split limit 1 th end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   262
49125
5fc5211cf104 implemented "mk_exhaust_tac"
blanchet
parents: 49121
diff changeset
   263
local
5fc5211cf104 implemented "mk_exhaust_tac"
blanchet
parents: 49121
diff changeset
   264
  fun mk_sumEN' 1 = @{thm obj_sum_step}
5fc5211cf104 implemented "mk_exhaust_tac"
blanchet
parents: 49121
diff changeset
   265
    | mk_sumEN' n = mk_sumEN' (n - 1) RSN (2, @{thm obj_sum_step});
5fc5211cf104 implemented "mk_exhaust_tac"
blanchet
parents: 49121
diff changeset
   266
in
5fc5211cf104 implemented "mk_exhaust_tac"
blanchet
parents: 49121
diff changeset
   267
  fun mk_sumEN 1 = @{thm obj_sum_base}
5fc5211cf104 implemented "mk_exhaust_tac"
blanchet
parents: 49121
diff changeset
   268
    | mk_sumEN 2 = @{thm sumE}
5fc5211cf104 implemented "mk_exhaust_tac"
blanchet
parents: 49121
diff changeset
   269
    | mk_sumEN n = (mk_sumEN' (n - 2) RSN (2, @{thm obj_sumE})) OF replicate n (impI RS allI);
5fc5211cf104 implemented "mk_exhaust_tac"
blanchet
parents: 49121
diff changeset
   270
end;
5fc5211cf104 implemented "mk_exhaust_tac"
blanchet
parents: 49121
diff changeset
   271
49130
3c26e17b2849 implemented "mk_case_tac" -- and got rid of "cheat_tac"
blanchet
parents: 49129
diff changeset
   272
fun mk_sum_casesN 1 1 = @{thm refl}
3c26e17b2849 implemented "mk_case_tac" -- and got rid of "cheat_tac"
blanchet
parents: 49129
diff changeset
   273
  | mk_sum_casesN _ 1 = @{thm sum.cases(1)}
3c26e17b2849 implemented "mk_case_tac" -- and got rid of "cheat_tac"
blanchet
parents: 49129
diff changeset
   274
  | mk_sum_casesN 2 2 = @{thm sum.cases(2)}
3c26e17b2849 implemented "mk_case_tac" -- and got rid of "cheat_tac"
blanchet
parents: 49129
diff changeset
   275
  | mk_sum_casesN n m = trans OF [@{thm sum_case_step(2)}, mk_sum_casesN (n - 1) (m - 1)];
3c26e17b2849 implemented "mk_case_tac" -- and got rid of "cheat_tac"
blanchet
parents: 49129
diff changeset
   276
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   277
fun mk_tactics mid mcomp mcong snat bdco bdinf sbd inbd wpull =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   278
  [mid, mcomp, mcong] @ snat @ [bdco, bdinf] @ sbd @ [inbd, wpull];
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   279
49141
aca966dc18f6 added comment for Dmitriy
blanchet
parents: 49134
diff changeset
   280
(* FIXME: because of "@ lhss", the output could contain type variables that are not in the input;
aca966dc18f6 added comment for Dmitriy
blanchet
parents: 49134
diff changeset
   281
   also, "fp_sort" should put the "resBs" first and in the order in which they appear *)
49185
073d7d1b7488 respect order of/additional type variables supplied by the user in fixed point constructions;
traytel
parents: 49176
diff changeset
   282
fun fp_sort lhss NONE Ass = Library.sort (Term_Ord.typ_ord o pairself TFree)
073d7d1b7488 respect order of/additional type variables supplied by the user in fixed point constructions;
traytel
parents: 49176
diff changeset
   283
    (subtract (op =) lhss (fold (fold (insert (op =))) Ass [])) @ lhss
073d7d1b7488 respect order of/additional type variables supplied by the user in fixed point constructions;
traytel
parents: 49176
diff changeset
   284
  | fp_sort lhss (SOME resBs) Ass =
073d7d1b7488 respect order of/additional type variables supplied by the user in fixed point constructions;
traytel
parents: 49176
diff changeset
   285
    (subtract (op =) lhss (filter (fn T => exists (fn Ts => member (op =) Ts T) Ass) resBs)) @ lhss;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   286
49185
073d7d1b7488 respect order of/additional type variables supplied by the user in fixed point constructions;
traytel
parents: 49176
diff changeset
   287
fun mk_fp_bnf timer construct resBs bs sort lhss bnfs deadss livess unfold lthy =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   288
  let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   289
    val name = fold_rev (fn b => fn s => Binding.name_of b ^ s) bs "";
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   290
    fun qualify i bind =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   291
      let val namei = if i > 0 then name ^ string_of_int i else name;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   292
      in
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   293
        if member (op =) (#2 (Binding.dest bind)) (namei, true) then bind
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   294
        else Binding.prefix_name namei bind
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   295
      end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   296
49132
81487fc17185 added robustness
blanchet
parents: 49130
diff changeset
   297
    val Ass = map (map dest_TFree) livess;
49185
073d7d1b7488 respect order of/additional type variables supplied by the user in fixed point constructions;
traytel
parents: 49176
diff changeset
   298
    val resDs = (case resBs of NONE => [] | SOME Ts => fold (subtract (op =)) Ass Ts);
073d7d1b7488 respect order of/additional type variables supplied by the user in fixed point constructions;
traytel
parents: 49176
diff changeset
   299
    val Ds = fold (fold Term.add_tfreesT) deadss [];
49132
81487fc17185 added robustness
blanchet
parents: 49130
diff changeset
   300
49156
2a361e09101b error message only in case of an error
traytel
parents: 49141
diff changeset
   301
    val _ = (case Library.inter (op =) Ds lhss of [] => ()
49132
81487fc17185 added robustness
blanchet
parents: 49130
diff changeset
   302
      | A :: _ => error ("Nonadmissible type recursion (cannot take fixed point of dead type \
81487fc17185 added robustness
blanchet
parents: 49130
diff changeset
   303
        \variable " ^ quote (Syntax.string_of_typ lthy (TFree A)) ^ ")"));
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   304
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   305
    val timer = time (timer "Construction of BNFs");
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   306
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   307
    val ((kill_poss, _), (bnfs', (unfold', lthy'))) =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   308
      normalize_bnfs qualify Ass Ds sort bnfs unfold lthy;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   309
49132
81487fc17185 added robustness
blanchet
parents: 49130
diff changeset
   310
    val Dss = map3 (append oo map o nth) livess kill_poss deadss;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   311
49185
073d7d1b7488 respect order of/additional type variables supplied by the user in fixed point constructions;
traytel
parents: 49176
diff changeset
   312
    val ((bnfs'', deadss), lthy'') =
49223
blanchet
parents: 49222
diff changeset
   313
      fold_map3 (seal_bnf unfold') (map (Binding.prefix_name preN) bs) Dss bnfs' lthy'
49185
073d7d1b7488 respect order of/additional type variables supplied by the user in fixed point constructions;
traytel
parents: 49176
diff changeset
   314
      |>> split_list;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   315
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   316
    val timer = time (timer "Normalization & sealing of BNFs");
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   317
49185
073d7d1b7488 respect order of/additional type variables supplied by the user in fixed point constructions;
traytel
parents: 49176
diff changeset
   318
    val res = construct resBs bs (map TFree resDs, deadss) bnfs'' lthy'';
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   319
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   320
    val timer = time (timer "FP construction in total");
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   321
  in
49226
510c6d4a73ec fixed and enabled iterator/recursor theorems
blanchet
parents: 49225
diff changeset
   322
    (bnfs'', res)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   323
  end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   324
49169
937a0fadddfb honor mixfix specifications
blanchet
parents: 49156
diff changeset
   325
fun fp_bnf construct bs mixfixes resBs eqs lthy =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   326
  let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   327
    val timer = time (Timer.startRealTimer ());
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   328
    val (lhss, rhss) = split_list eqs;
49185
073d7d1b7488 respect order of/additional type variables supplied by the user in fixed point constructions;
traytel
parents: 49176
diff changeset
   329
    val sort = fp_sort lhss (SOME resBs);
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   330
    val ((bnfs, (Dss, Ass)), (unfold, lthy')) = apfst (apsnd split_list o split_list)
49223
blanchet
parents: 49222
diff changeset
   331
      (fold_map2 (fn b => bnf_of_typ Smart_Inline (Binding.prefix_name rawN b) I sort) bs rhss
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   332
        (empty_unfold, lthy));
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   333
  in
49185
073d7d1b7488 respect order of/additional type variables supplied by the user in fixed point constructions;
traytel
parents: 49176
diff changeset
   334
    mk_fp_bnf timer (construct mixfixes) (SOME resBs) bs sort lhss bnfs Dss Ass unfold lthy'
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   335
  end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   336
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   337
fun fp_bnf_cmd construct (bs, (raw_lhss, raw_bnfs)) lthy =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   338
  let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   339
    val timer = time (Timer.startRealTimer ());
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   340
    val lhss = map (dest_TFree o Syntax.read_typ lthy) raw_lhss;
49185
073d7d1b7488 respect order of/additional type variables supplied by the user in fixed point constructions;
traytel
parents: 49176
diff changeset
   341
    val sort = fp_sort lhss NONE;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   342
    val ((bnfs, (Dss, Ass)), (unfold, lthy')) = apfst (apsnd split_list o split_list)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   343
      (fold_map2 (fn b => fn rawT =>
49223
blanchet
parents: 49222
diff changeset
   344
        (bnf_of_typ Smart_Inline (Binding.prefix_name rawN b) I sort (Syntax.read_typ lthy rawT)))
49185
073d7d1b7488 respect order of/additional type variables supplied by the user in fixed point constructions;
traytel
parents: 49176
diff changeset
   345
      bs raw_bnfs (empty_unfold, lthy));
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   346
  in
49207
4634c217b77b completed iter/rec proofs
blanchet
parents: 49200
diff changeset
   347
    snd (mk_fp_bnf timer (construct (map (K NoSyn) bs)) NONE bs sort lhss bnfs Dss Ass unfold lthy')
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   348
  end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   349
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   350
end;