src/HOL/Tools/BNF/bnf_lfp_compat.ML
author blanchet
Mon, 25 Apr 2016 19:23:20 +0200
changeset 63047 2146553e96c6
parent 62497 5b5b704f4811
child 63064 2f18172214c8
permissions -rw-r--r--
generalize code to avoid making assumptions about type variable names
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
55061
a0adf838e2d1 adjusted comments
blanchet
parents: 55060
diff changeset
     1
(*  Title:      HOL/Tools/BNF/bnf_lfp_compat.ML
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
     2
    Author:     Jasmin Blanchette, TU Muenchen
58117
9608028d8f43 more compatibility between old- and new-style datatypes
blanchet
parents: 58113
diff changeset
     3
    Copyright   2013, 2014
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
     4
58147
967444d352b8 more compatibility functions
blanchet
parents: 58146
diff changeset
     5
Compatibility layer with the old datatype package. Partly based on
58119
blanchet
parents: 58117
diff changeset
     6
blanchet
parents: 58117
diff changeset
     7
    Title:      HOL/Tools/Old_Datatype/old_datatype_data.ML
blanchet
parents: 58117
diff changeset
     8
    Author:     Stefan Berghofer, TU Muenchen
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
     9
*)
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    10
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    11
signature BNF_LFP_COMPAT =
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    12
sig
58354
04ac60da613e support (finite values of) codatatypes in Quickcheck
blanchet
parents: 58340
diff changeset
    13
  datatype preference = Keep_Nesting | Include_GFPs | Kill_Type_Args
58123
62765d39539f implemented compatibility definition of datatype
blanchet
parents: 58122
diff changeset
    14
58354
04ac60da613e support (finite values of) codatatypes in Quickcheck
blanchet
parents: 58340
diff changeset
    15
  val get_all: theory -> preference list -> Old_Datatype_Aux.info Symtab.table
04ac60da613e support (finite values of) codatatypes in Quickcheck
blanchet
parents: 58340
diff changeset
    16
  val get_info: theory -> preference list -> string -> Old_Datatype_Aux.info option
04ac60da613e support (finite values of) codatatypes in Quickcheck
blanchet
parents: 58340
diff changeset
    17
  val the_info: theory -> preference list -> string -> Old_Datatype_Aux.info
58129
3ec65a7f2b50 ported Refute to use new datatypes when possible
blanchet
parents: 58126
diff changeset
    18
  val the_spec: theory -> string -> (string * sort) list * (string * typ list) list
58354
04ac60da613e support (finite values of) codatatypes in Quickcheck
blanchet
parents: 58340
diff changeset
    19
  val the_descr: theory -> preference list -> string list ->
58117
9608028d8f43 more compatibility between old- and new-style datatypes
blanchet
parents: 58113
diff changeset
    20
    Old_Datatype_Aux.descr * (string * sort) list * string list * string
9608028d8f43 more compatibility between old- and new-style datatypes
blanchet
parents: 58113
diff changeset
    21
    * (string list * string list) * (typ list * typ list)
58129
3ec65a7f2b50 ported Refute to use new datatypes when possible
blanchet
parents: 58126
diff changeset
    22
  val get_constrs: theory -> string -> (string * typ) list option
58354
04ac60da613e support (finite values of) codatatypes in Quickcheck
blanchet
parents: 58340
diff changeset
    23
  val interpretation: string -> preference list ->
58123
62765d39539f implemented compatibility definition of datatype
blanchet
parents: 58122
diff changeset
    24
    (Old_Datatype_Aux.config -> string list -> theory -> theory) -> theory -> theory
58125
a2ba381607fb more work on compatibility interfaces
blanchet
parents: 58123
diff changeset
    25
  val datatype_compat: string list -> local_theory -> local_theory
a2ba381607fb more work on compatibility interfaces
blanchet
parents: 58123
diff changeset
    26
  val datatype_compat_global: string list -> theory -> theory
58117
9608028d8f43 more compatibility between old- and new-style datatypes
blanchet
parents: 58113
diff changeset
    27
  val datatype_compat_cmd: string list -> local_theory -> local_theory
58354
04ac60da613e support (finite values of) codatatypes in Quickcheck
blanchet
parents: 58340
diff changeset
    28
  val add_datatype: preference list -> Old_Datatype_Aux.spec list -> theory -> string list * theory
60003
ba8fa0c38d66 renamed ML funs
blanchet
parents: 59936
diff changeset
    29
  val primrec: (binding * typ option * mixfix) list -> (Attrib.binding * term) list ->
58684
56b9eab818ca tuned whitespace
blanchet
parents: 58665
diff changeset
    30
    local_theory -> (term list * thm list) * local_theory
60003
ba8fa0c38d66 renamed ML funs
blanchet
parents: 59936
diff changeset
    31
  val primrec_global: (binding * typ option * mixfix) list ->
58147
967444d352b8 more compatibility functions
blanchet
parents: 58146
diff changeset
    32
    (Attrib.binding * term) list -> theory -> (term list * thm list) * theory
60003
ba8fa0c38d66 renamed ML funs
blanchet
parents: 59936
diff changeset
    33
  val primrec_overloaded: (string * (string * typ) * bool) list ->
58147
967444d352b8 more compatibility functions
blanchet
parents: 58146
diff changeset
    34
    (binding * typ option * mixfix) list -> (Attrib.binding * term) list -> theory ->
967444d352b8 more compatibility functions
blanchet
parents: 58146
diff changeset
    35
    (term list * thm list) * theory
60003
ba8fa0c38d66 renamed ML funs
blanchet
parents: 59936
diff changeset
    36
  val primrec_simple: ((binding * typ) * mixfix) list -> term list ->
58220
a2afad27a0b1 wildcards in plugins
blanchet
parents: 58219
diff changeset
    37
    local_theory -> (string list * (term list * thm list)) * local_theory
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    38
end;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    39
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    40
structure BNF_LFP_Compat : BNF_LFP_COMPAT =
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    41
struct
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    42
54006
9fe1bd54d437 renamed theory file
blanchet
parents: 54003
diff changeset
    43
open Ctr_Sugar
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    44
open BNF_Util
58211
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
    45
open BNF_Tactics
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    46
open BNF_FP_Util
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    47
open BNF_FP_Def_Sugar
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    48
open BNF_FP_N2M_Sugar
58123
62765d39539f implemented compatibility definition of datatype
blanchet
parents: 58122
diff changeset
    49
open BNF_LFP
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    50
58211
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
    51
val compat_N = "compat_";
58213
blanchet
parents: 58211
diff changeset
    52
val rec_split_N = "rec_split_";
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    53
58354
04ac60da613e support (finite values of) codatatypes in Quickcheck
blanchet
parents: 58340
diff changeset
    54
datatype preference = Keep_Nesting | Include_GFPs | Kill_Type_Args;
58123
62765d39539f implemented compatibility definition of datatype
blanchet
parents: 58122
diff changeset
    55
58213
blanchet
parents: 58211
diff changeset
    56
fun mk_split_rec_rhs ctxt fpTs Cs (recs as rec1 :: _) =
58211
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
    57
  let
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
    58
    fun repair_rec_arg_args [] [] = []
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
    59
      | repair_rec_arg_args ((g_T as Type (@{type_name fun}, _)) :: g_Ts) (g :: gs) =
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
    60
        let
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
    61
          val (x_Ts, body_T) = strip_type g_T;
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
    62
        in
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
    63
          (case try HOLogic.dest_prodT body_T of
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
    64
            NONE => [g]
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
    65
          | SOME (fst_T, _) =>
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
    66
            if member (op =) fpTs fst_T then
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
    67
              let val (xs, _) = mk_Frees "x" x_Ts ctxt in
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
    68
                map (fn mk_proj => fold_rev Term.lambda xs (mk_proj (Term.list_comb (g, xs))))
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
    69
                  [HOLogic.mk_fst, HOLogic.mk_snd]
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
    70
              end
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
    71
            else
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
    72
              [g])
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
    73
          :: repair_rec_arg_args g_Ts gs
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
    74
        end
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
    75
      | repair_rec_arg_args (g_T :: g_Ts) (g :: gs) =
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
    76
        if member (op =) fpTs g_T then
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
    77
          let
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
    78
            val j = find_index (member (op =) Cs) g_Ts;
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
    79
            val h = nth gs j;
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
    80
            val g_Ts' = nth_drop j g_Ts;
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
    81
            val gs' = nth_drop j gs;
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
    82
          in
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
    83
            [g, h] :: repair_rec_arg_args g_Ts' gs'
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
    84
          end
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
    85
        else
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
    86
          [g] :: repair_rec_arg_args g_Ts gs;
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
    87
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
    88
    fun repair_back_rec_arg f_T f' =
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
    89
      let
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
    90
        val g_Ts = Term.binder_types f_T;
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
    91
        val (gs, _) = mk_Frees "g" g_Ts ctxt;
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
    92
      in
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
    93
        fold_rev Term.lambda gs (Term.list_comb (f',
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
    94
          flat_rec_arg_args (repair_rec_arg_args g_Ts gs)))
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
    95
      end;
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
    96
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
    97
    val f_Ts = binder_fun_types (fastype_of rec1);
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
    98
    val (fs', _) = mk_Frees "f" (replicate (length f_Ts) Term.dummyT) ctxt;
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
    99
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
   100
    fun mk_rec' recx =
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
   101
      fold_rev Term.lambda fs' (Term.list_comb (recx, map2 repair_back_rec_arg f_Ts fs'))
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
   102
      |> Syntax.check_term ctxt;
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
   103
  in
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
   104
    map mk_rec' recs
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
   105
  end;
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
   106
58213
blanchet
parents: 58211
diff changeset
   107
fun define_split_recs fpTs Cs recs lthy =
58211
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
   108
  let
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
   109
    val b_names = Name.variant_list [] (map base_name_of_typ fpTs);
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
   110
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
   111
    fun mk_binding b_name =
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
   112
      Binding.qualify true (compat_N ^ b_name)
58213
blanchet
parents: 58211
diff changeset
   113
        (Binding.prefix_name rec_split_N (Binding.name b_name));
58211
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
   114
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
   115
    val bs = map mk_binding b_names;
58213
blanchet
parents: 58211
diff changeset
   116
    val rhss = mk_split_rec_rhs lthy fpTs Cs recs;
58211
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
   117
  in
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58461
diff changeset
   118
    @{fold_map 3} (define_co_rec_as Least_FP Cs) fpTs bs rhss lthy
58211
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
   119
  end;
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
   120
58218
a92acec845a7 no type-based lookup -- these fail in the general, ambiguous case
blanchet
parents: 58217
diff changeset
   121
fun mk_split_rec_thmss ctxt Xs ctrXs_Tsss ctrss rec0_thmss (recs as rec1 :: _) rec_defs =
58211
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
   122
  let
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
   123
    val f_Ts = binder_fun_types (fastype_of rec1);
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
   124
    val (fs, _) = mk_Frees "f" f_Ts ctxt;
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
   125
    val frecs = map (fn recx => Term.list_comb (recx, fs)) recs;
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
   126
58218
a92acec845a7 no type-based lookup -- these fail in the general, ambiguous case
blanchet
parents: 58217
diff changeset
   127
    val Xs_frecs = Xs ~~ frecs;
58211
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
   128
    val fss = unflat ctrss fs;
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
   129
63047
2146553e96c6 generalize code to avoid making assumptions about type variable names
blanchet
parents: 62497
diff changeset
   130
    fun mk_rec_call g n (Type (@{type_name fun}, [_, ran_T])) =
2146553e96c6 generalize code to avoid making assumptions about type variable names
blanchet
parents: 62497
diff changeset
   131
        Abs (Name.uu, Term.dummyT, mk_rec_call g (n + 1) ran_T)
58218
a92acec845a7 no type-based lookup -- these fail in the general, ambiguous case
blanchet
parents: 58217
diff changeset
   132
      | mk_rec_call g n X =
58211
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
   133
        let
58218
a92acec845a7 no type-based lookup -- these fail in the general, ambiguous case
blanchet
parents: 58217
diff changeset
   134
          val frec = the (AList.lookup (op =) Xs_frecs X);
58211
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
   135
          val xg = Term.list_comb (g, map Bound (n - 1 downto 0));
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
   136
        in frec $ xg end;
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
   137
58218
a92acec845a7 no type-based lookup -- these fail in the general, ambiguous case
blanchet
parents: 58217
diff changeset
   138
    fun mk_rec_arg_arg ctrXs_T g =
a92acec845a7 no type-based lookup -- these fail in the general, ambiguous case
blanchet
parents: 58217
diff changeset
   139
      g :: (if member (op =) Xs (body_type ctrXs_T) then [mk_rec_call g 0 ctrXs_T] else []);
58211
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
   140
58218
a92acec845a7 no type-based lookup -- these fail in the general, ambiguous case
blanchet
parents: 58217
diff changeset
   141
    fun mk_goal frec ctrXs_Ts ctr f =
58211
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
   142
      let
58218
a92acec845a7 no type-based lookup -- these fail in the general, ambiguous case
blanchet
parents: 58217
diff changeset
   143
        val ctr_Ts = binder_types (fastype_of ctr);
58217
blanchet
parents: 58214
diff changeset
   144
        val (gs, _) = mk_Frees "g" ctr_Ts ctxt;
58211
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
   145
        val gctr = Term.list_comb (ctr, gs);
58218
a92acec845a7 no type-based lookup -- these fail in the general, ambiguous case
blanchet
parents: 58217
diff changeset
   146
        val fgs = flat_rec_arg_args (map2 mk_rec_arg_arg ctrXs_Ts gs);
58211
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
   147
      in
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
   148
        fold_rev (fold_rev Logic.all) [fs, gs]
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
   149
          (mk_Trueprop_eq (frec $ gctr, Term.list_comb (f, fgs)))
63047
2146553e96c6 generalize code to avoid making assumptions about type variable names
blanchet
parents: 62497
diff changeset
   150
        |> Syntax.check_term ctxt
58211
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
   151
      end;
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
   152
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58461
diff changeset
   153
    val goalss = @{map 4} (@{map 3} o mk_goal) frecs ctrXs_Tsss ctrss fss;
58211
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
   154
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
   155
    fun tac ctxt =
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
   156
      unfold_thms_tac ctxt (@{thms o_apply fst_conv snd_conv} @ rec_defs @ flat rec0_thmss) THEN
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 60004
diff changeset
   157
      HEADGOAL (rtac ctxt refl);
58211
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
   158
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
   159
    fun prove goal =
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
   160
      Goal.prove_sorry ctxt [] [] goal (tac o #context)
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
   161
      |> Thm.close_derivation;
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
   162
  in
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
   163
    map (map prove) goalss
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
   164
  end;
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
   165
58218
a92acec845a7 no type-based lookup -- these fail in the general, ambiguous case
blanchet
parents: 58217
diff changeset
   166
fun define_split_rec_derive_induct_rec_thms Xs fpTs ctrXs_Tsss ctrss inducts induct recs0 rec_thmss
58217
blanchet
parents: 58214
diff changeset
   167
    lthy =
58211
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
   168
  let
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
   169
    val thy = Proof_Context.theory_of lthy;
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
   170
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
   171
    (* imperfect: will not yield the expected theorem for functions taking a large number of
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
   172
       arguments *)
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
   173
    val repair_induct = unfold_thms lthy @{thms all_mem_range};
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
   174
58214
bd1754377965 properly note theorems for split recursors
blanchet
parents: 58213
diff changeset
   175
    val inducts' = map repair_induct inducts;
58211
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
   176
    val induct' = repair_induct induct;
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
   177
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
   178
    val Cs = map ((fn TVar ((s, _), S) => TFree (s, S)) o body_type o fastype_of) recs0;
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
   179
    val recs = map2 (mk_co_rec thy Least_FP Cs) fpTs recs0;
58213
blanchet
parents: 58211
diff changeset
   180
    val ((recs', rec'_defs), lthy') = define_split_recs fpTs Cs recs lthy |>> split_list;
58218
a92acec845a7 no type-based lookup -- these fail in the general, ambiguous case
blanchet
parents: 58217
diff changeset
   181
    val rec'_thmss = mk_split_rec_thmss lthy' Xs ctrXs_Tsss ctrss rec_thmss recs' rec'_defs;
58211
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
   182
  in
58214
bd1754377965 properly note theorems for split recursors
blanchet
parents: 58213
diff changeset
   183
    ((inducts', induct', recs', rec'_thmss), lthy')
58211
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
   184
  end;
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
   185
58219
61b852f90161 improved 'datatype_compat' further for recursion through functions
blanchet
parents: 58218
diff changeset
   186
fun body_rec_indices (Old_Datatype_Aux.DtRec kk) = [kk]
61b852f90161 improved 'datatype_compat' further for recursion through functions
blanchet
parents: 58218
diff changeset
   187
  | body_rec_indices (Old_Datatype_Aux.DtType (@{type_name fun}, [_, D])) = body_rec_indices D
61b852f90161 improved 'datatype_compat' further for recursion through functions
blanchet
parents: 58218
diff changeset
   188
  | body_rec_indices _ = [];
61b852f90161 improved 'datatype_compat' further for recursion through functions
blanchet
parents: 58218
diff changeset
   189
56453
00548d372f02 support deeply nested datatypes in 'datatype_compat'
blanchet
parents: 56452
diff changeset
   190
fun reindex_desc desc =
00548d372f02 support deeply nested datatypes in 'datatype_compat'
blanchet
parents: 56452
diff changeset
   191
  let
00548d372f02 support deeply nested datatypes in 'datatype_compat'
blanchet
parents: 56452
diff changeset
   192
    val kks = map fst desc;
00548d372f02 support deeply nested datatypes in 'datatype_compat'
blanchet
parents: 56452
diff changeset
   193
    val perm_kks = sort int_ord kks;
00548d372f02 support deeply nested datatypes in 'datatype_compat'
blanchet
parents: 56452
diff changeset
   194
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 57983
diff changeset
   195
    fun perm_dtyp (Old_Datatype_Aux.DtType (s, Ds)) = Old_Datatype_Aux.DtType (s, map perm_dtyp Ds)
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 57983
diff changeset
   196
      | perm_dtyp (Old_Datatype_Aux.DtRec kk) =
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 57983
diff changeset
   197
        Old_Datatype_Aux.DtRec (find_index (curry (op =) kk) kks)
58113
blanchet
parents: 58112
diff changeset
   198
      | perm_dtyp D = D;
56453
00548d372f02 support deeply nested datatypes in 'datatype_compat'
blanchet
parents: 56452
diff changeset
   199
  in
00548d372f02 support deeply nested datatypes in 'datatype_compat'
blanchet
parents: 56452
diff changeset
   200
    if perm_kks = kks then
00548d372f02 support deeply nested datatypes in 'datatype_compat'
blanchet
parents: 56452
diff changeset
   201
      desc
00548d372f02 support deeply nested datatypes in 'datatype_compat'
blanchet
parents: 56452
diff changeset
   202
    else
00548d372f02 support deeply nested datatypes in 'datatype_compat'
blanchet
parents: 56452
diff changeset
   203
      perm_kks ~~
00548d372f02 support deeply nested datatypes in 'datatype_compat'
blanchet
parents: 56452
diff changeset
   204
      map (fn (_, (s, Ds, sDss)) => (s, map perm_dtyp Ds, map (apsnd (map perm_dtyp)) sDss)) desc
58113
blanchet
parents: 58112
diff changeset
   205
  end;
56453
00548d372f02 support deeply nested datatypes in 'datatype_compat'
blanchet
parents: 56452
diff changeset
   206
58354
04ac60da613e support (finite values of) codatatypes in Quickcheck
blanchet
parents: 58340
diff changeset
   207
fun mk_infos_of_mutually_recursive_new_datatypes prefs check_names fpT_names0 lthy =
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   208
  let
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   209
    val thy = Proof_Context.theory_of lthy;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   210
58315
6d8458bc6e27 tuning terminology
blanchet
parents: 58300
diff changeset
   211
    fun not_datatype s = error (quote s ^ " is not a datatype");
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   212
    fun not_mutually_recursive ss =
58315
6d8458bc6e27 tuning terminology
blanchet
parents: 58300
diff changeset
   213
      error ("{" ^ commas ss ^ "} is not a complete set of mutually recursive datatypes");
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   214
58354
04ac60da613e support (finite values of) codatatypes in Quickcheck
blanchet
parents: 58340
diff changeset
   215
    fun checked_fp_sugar_of s =
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   216
      (case fp_sugar_of lthy s of
58354
04ac60da613e support (finite values of) codatatypes in Quickcheck
blanchet
parents: 58340
diff changeset
   217
        SOME (fp_sugar as {fp, ...}) =>
04ac60da613e support (finite values of) codatatypes in Quickcheck
blanchet
parents: 58340
diff changeset
   218
        if member (op =) prefs Include_GFPs orelse fp = Least_FP then fp_sugar else not_datatype s
04ac60da613e support (finite values of) codatatypes in Quickcheck
blanchet
parents: 58340
diff changeset
   219
      | NONE => not_datatype s);
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   220
58229
cece11f6262a export right sorts
blanchet
parents: 58220
diff changeset
   221
    val fpTs0 as Type (_, var_As) :: _ =
58354
04ac60da613e support (finite values of) codatatypes in Quickcheck
blanchet
parents: 58340
diff changeset
   222
      map (#T o checked_fp_sugar_of o fst o dest_Type)
04ac60da613e support (finite values of) codatatypes in Quickcheck
blanchet
parents: 58340
diff changeset
   223
        (#Ts (#fp_res (checked_fp_sugar_of (hd fpT_names0))));
56455
1ff66e72628b allow arguments to 'datatype_compat' in disorder
blanchet
parents: 56453
diff changeset
   224
    val fpT_names = map (fst o dest_Type) fpTs0;
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   225
58117
9608028d8f43 more compatibility between old- and new-style datatypes
blanchet
parents: 58113
diff changeset
   226
    val _ = check_names (op =) (fpT_names0, fpT_names) orelse not_mutually_recursive fpT_names0;
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   227
58117
9608028d8f43 more compatibility between old- and new-style datatypes
blanchet
parents: 58113
diff changeset
   228
    val (As_names, _) = Variable.variant_fixes (map (fn TVar ((s, _), _) => s) var_As) lthy;
56452
0c98c9118407 preserve user type variable names to avoid mismatches/confusion
blanchet
parents: 56002
diff changeset
   229
    val As = map2 (fn s => fn TVar (_, S) => TFree (s, S)) As_names var_As;
56455
1ff66e72628b allow arguments to 'datatype_compat' in disorder
blanchet
parents: 56453
diff changeset
   230
    val fpTs = map (fn s => Type (s, As)) fpT_names;
56453
00548d372f02 support deeply nested datatypes in 'datatype_compat'
blanchet
parents: 56452
diff changeset
   231
00548d372f02 support deeply nested datatypes in 'datatype_compat'
blanchet
parents: 56452
diff changeset
   232
    val nn_fp = length fpTs;
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   233
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 57983
diff changeset
   234
    val mk_dtyp = Old_Datatype_Aux.dtyp_of_typ (map (apsnd (map Term.dest_TFree) o dest_Type) fpTs);
56453
00548d372f02 support deeply nested datatypes in 'datatype_compat'
blanchet
parents: 56452
diff changeset
   235
00548d372f02 support deeply nested datatypes in 'datatype_compat'
blanchet
parents: 56452
diff changeset
   236
    fun mk_ctr_descr Ts = mk_ctr Ts #> dest_Const ##> (binder_types #> map mk_dtyp);
00548d372f02 support deeply nested datatypes in 'datatype_compat'
blanchet
parents: 56452
diff changeset
   237
    fun mk_typ_descr index (Type (T_name, Ts)) ({ctrs, ...} : ctr_sugar) =
00548d372f02 support deeply nested datatypes in 'datatype_compat'
blanchet
parents: 56452
diff changeset
   238
      (index, (T_name, map mk_dtyp Ts, map (mk_ctr_descr Ts) ctrs));
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   239
58460
a88eb33058f7 refactor fp_sugar move theorems
desharna
parents: 58459
diff changeset
   240
    val fp_ctr_sugars = map (#ctr_sugar o #fp_ctr_sugar o checked_fp_sugar_of) fpT_names;
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 58461
diff changeset
   241
    val orig_descr = @{map 3} mk_typ_descr (0 upto nn_fp - 1) fpTs fp_ctr_sugars;
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 57983
diff changeset
   242
    val all_infos = Old_Datatype_Data.get_all thy;
58117
9608028d8f43 more compatibility between old- and new-style datatypes
blanchet
parents: 58113
diff changeset
   243
    val (orig_descr' :: nested_descrs) =
58354
04ac60da613e support (finite values of) codatatypes in Quickcheck
blanchet
parents: 58340
diff changeset
   244
      if member (op =) prefs Keep_Nesting then [orig_descr]
58125
a2ba381607fb more work on compatibility interfaces
blanchet
parents: 58123
diff changeset
   245
      else fst (Old_Datatype_Aux.unfold_datatypes lthy orig_descr all_infos orig_descr nn_fp);
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   246
57798
018dc778cbcc more correct clique computation for N2M
blanchet
parents: 57794
diff changeset
   247
    fun cliquify_descr [] = []
018dc778cbcc more correct clique computation for N2M
blanchet
parents: 57794
diff changeset
   248
      | cliquify_descr [entry] = [[entry]]
018dc778cbcc more correct clique computation for N2M
blanchet
parents: 57794
diff changeset
   249
      | cliquify_descr (full_descr as (_, (T_name1, _, _)) :: _) =
018dc778cbcc more correct clique computation for N2M
blanchet
parents: 57794
diff changeset
   250
        let
018dc778cbcc more correct clique computation for N2M
blanchet
parents: 57794
diff changeset
   251
          val nn =
018dc778cbcc more correct clique computation for N2M
blanchet
parents: 57794
diff changeset
   252
            if member (op =) fpT_names T_name1 then
018dc778cbcc more correct clique computation for N2M
blanchet
parents: 57794
diff changeset
   253
              nn_fp
018dc778cbcc more correct clique computation for N2M
blanchet
parents: 57794
diff changeset
   254
            else
018dc778cbcc more correct clique computation for N2M
blanchet
parents: 57794
diff changeset
   255
              (case Symtab.lookup all_infos T_name1 of
018dc778cbcc more correct clique computation for N2M
blanchet
parents: 57794
diff changeset
   256
                SOME {descr, ...} =>
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 57983
diff changeset
   257
                length (filter_out (exists Old_Datatype_Aux.is_rec_type o #2 o snd) descr)
57798
018dc778cbcc more correct clique computation for N2M
blanchet
parents: 57794
diff changeset
   258
              | NONE => raise Fail "unknown old-style datatype");
018dc778cbcc more correct clique computation for N2M
blanchet
parents: 57794
diff changeset
   259
        in
018dc778cbcc more correct clique computation for N2M
blanchet
parents: 57794
diff changeset
   260
          chop nn full_descr ||> cliquify_descr |> op ::
018dc778cbcc more correct clique computation for N2M
blanchet
parents: 57794
diff changeset
   261
        end;
018dc778cbcc more correct clique computation for N2M
blanchet
parents: 57794
diff changeset
   262
56453
00548d372f02 support deeply nested datatypes in 'datatype_compat'
blanchet
parents: 56452
diff changeset
   263
    (* put nested types before the types that nest them, as needed for N2M *)
56484
c451cf8b29c8 thread mutual cliques through
blanchet
parents: 56455
diff changeset
   264
    val descrs = burrow reindex_desc (orig_descr' :: rev nested_descrs);
58340
blanchet
parents: 58335
diff changeset
   265
    val (mutual_cliques, descr) =
57798
018dc778cbcc more correct clique computation for N2M
blanchet
parents: 57794
diff changeset
   266
      split_list (flat (map_index (fn (i, descr) => map (pair i) descr)
018dc778cbcc more correct clique computation for N2M
blanchet
parents: 57794
diff changeset
   267
        (maps cliquify_descr descrs)));
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   268
58211
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
   269
    val fpTs' = Old_Datatype_Aux.get_rec_types descr;
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
   270
    val nn = length fpTs';
55485
bdfb607543f4 better handling of recursion through functions
blanchet
parents: 55481
diff changeset
   271
58354
04ac60da613e support (finite values of) codatatypes in Quickcheck
blanchet
parents: 58340
diff changeset
   272
    val fp_sugars = map (checked_fp_sugar_of o fst o dest_Type) fpTs';
58219
61b852f90161 improved 'datatype_compat' further for recursion through functions
blanchet
parents: 58218
diff changeset
   273
    val ctr_Tsss = map (map (map (Old_Datatype_Aux.typ_of_dtyp descr) o snd) o #3 o snd) descr;
61b852f90161 improved 'datatype_compat' further for recursion through functions
blanchet
parents: 58218
diff changeset
   274
    val kkssss = map (map (map body_rec_indices o snd) o #3 o snd) descr;
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   275
55772
367ec44763fd correct most general type for mutual recursion when several identical types are involved
blanchet
parents: 55571
diff changeset
   276
    val callers = map (fn kk => Var ((Name.uu, kk), @{typ "unit => unit"})) (0 upto nn - 1);
367ec44763fd correct most general type for mutual recursion when several identical types are involved
blanchet
parents: 55571
diff changeset
   277
367ec44763fd correct most general type for mutual recursion when several identical types are involved
blanchet
parents: 55571
diff changeset
   278
    fun apply_comps n kk =
55966
972f0aa7091b balance tuples that represent curried functions
blanchet
parents: 55951
diff changeset
   279
      mk_partial_compN n (replicate n HOLogic.unitT ---> HOLogic.unitT) (nth callers kk);
55772
367ec44763fd correct most general type for mutual recursion when several identical types are involved
blanchet
parents: 55571
diff changeset
   280
58219
61b852f90161 improved 'datatype_compat' further for recursion through functions
blanchet
parents: 58218
diff changeset
   281
    val callssss = map2 (map2 (map2 (map o apply_comps o num_binder_types))) ctr_Tsss kkssss;
56453
00548d372f02 support deeply nested datatypes in 'datatype_compat'
blanchet
parents: 56452
diff changeset
   282
58211
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
   283
    val b_names = Name.variant_list [] (map base_name_of_typ fpTs');
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
   284
    val compat_b_names = map (prefix compat_N) b_names;
56453
00548d372f02 support deeply nested datatypes in 'datatype_compat'
blanchet
parents: 56452
diff changeset
   285
    val compat_bs = map Binding.name compat_b_names;
55772
367ec44763fd correct most general type for mutual recursion when several identical types are involved
blanchet
parents: 55571
diff changeset
   286
58217
blanchet
parents: 58214
diff changeset
   287
    val ((fp_sugars', (lfp_sugar_thms', _)), lthy') =
54286
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54267
diff changeset
   288
      if nn > nn_fp then
58340
blanchet
parents: 58335
diff changeset
   289
        mutualize_fp_sugars (K true) Least_FP mutual_cliques compat_bs fpTs' callers callssss
blanchet
parents: 58335
diff changeset
   290
          fp_sugars lthy
54286
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54267
diff changeset
   291
      else
58217
blanchet
parents: 58214
diff changeset
   292
        ((fp_sugars, (NONE, NONE)), lthy);
blanchet
parents: 58214
diff changeset
   293
58460
a88eb33058f7 refactor fp_sugar move theorems
desharna
parents: 58459
diff changeset
   294
    fun mk_ctr_of ({fp_ctr_sugar = {ctr_sugar = {ctrs, ...}, ...}, ...} : fp_sugar) (Type (_, Ts)) =
a88eb33058f7 refactor fp_sugar move theorems
desharna
parents: 58459
diff changeset
   295
      map (mk_ctr Ts) ctrs;
58218
a92acec845a7 no type-based lookup -- these fail in the general, ambiguous case
blanchet
parents: 58217
diff changeset
   296
    val substAT = Term.typ_subst_atomic (var_As ~~ As);
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   297
58217
blanchet
parents: 58214
diff changeset
   298
    val Xs' = map #X fp_sugars';
58460
a88eb33058f7 refactor fp_sugar move theorems
desharna
parents: 58459
diff changeset
   299
    val ctrXs_Tsss' = map (map (map substAT) o #ctrXs_Tss o #fp_ctr_sugar) fp_sugars';
58217
blanchet
parents: 58214
diff changeset
   300
    val ctrss' = map2 mk_ctr_of fp_sugars' fpTs';
58461
75ee8d49c724 refactor fp_sugar move theorems
desharna
parents: 58460
diff changeset
   301
    val {fp_co_induct_sugar = {common_co_inducts = induct :: _, ...}, ...} :: _ = fp_sugars';
58459
f70bffabd7cf refactor fp_sugar move theorems
desharna
parents: 58358
diff changeset
   302
    val inducts = map (hd o #co_inducts o #fp_co_induct_sugar) fp_sugars';
58461
75ee8d49c724 refactor fp_sugar move theorems
desharna
parents: 58460
diff changeset
   303
    val recs = map (#co_rec o #fp_co_induct_sugar) fp_sugars';
58459
f70bffabd7cf refactor fp_sugar move theorems
desharna
parents: 58358
diff changeset
   304
    val rec_thmss = map (#co_rec_thms o #fp_co_induct_sugar) fp_sugars';
58211
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
   305
58218
a92acec845a7 no type-based lookup -- these fail in the general, ambiguous case
blanchet
parents: 58217
diff changeset
   306
    fun is_nested_rec_type (Type (@{type_name fun}, [_, T])) = member (op =) Xs' (body_type T)
58211
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
   307
      | is_nested_rec_type _ = false;
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
   308
58217
blanchet
parents: 58214
diff changeset
   309
    val ((lfp_sugar_thms'', (inducts', induct', recs', rec'_thmss)), lthy'') =
58354
04ac60da613e support (finite values of) codatatypes in Quickcheck
blanchet
parents: 58340
diff changeset
   310
      if member (op =) prefs Keep_Nesting orelse
04ac60da613e support (finite values of) codatatypes in Quickcheck
blanchet
parents: 58340
diff changeset
   311
         not (exists (exists (exists is_nested_rec_type)) ctrXs_Tsss') then
04ac60da613e support (finite values of) codatatypes in Quickcheck
blanchet
parents: 58340
diff changeset
   312
        ((lfp_sugar_thms', (inducts, induct, recs, rec_thmss)), lthy')
04ac60da613e support (finite values of) codatatypes in Quickcheck
blanchet
parents: 58340
diff changeset
   313
      else
58218
a92acec845a7 no type-based lookup -- these fail in the general, ambiguous case
blanchet
parents: 58217
diff changeset
   314
        define_split_rec_derive_induct_rec_thms Xs' fpTs' ctrXs_Tsss' ctrss' inducts induct recs
58217
blanchet
parents: 58214
diff changeset
   315
          rec_thmss lthy'
58214
bd1754377965 properly note theorems for split recursors
blanchet
parents: 58213
diff changeset
   316
        |>> `(fn (inducts', induct', _, rec'_thmss) =>
58354
04ac60da613e support (finite values of) codatatypes in Quickcheck
blanchet
parents: 58340
diff changeset
   317
          SOME ((inducts', induct', mk_induct_attrs ctrss'), (rec'_thmss, [])));
58211
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
   318
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
   319
    val rec'_names = map (fst o dest_Const) recs';
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
   320
    val rec'_thms = flat rec'_thmss;
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
   321
58460
a88eb33058f7 refactor fp_sugar move theorems
desharna
parents: 58459
diff changeset
   322
    fun mk_info (kk, {T = Type (T_name0, _), fp_ctr_sugar = {ctr_sugar = {casex, exhaust, nchotomy,
a88eb33058f7 refactor fp_sugar move theorems
desharna
parents: 58459
diff changeset
   323
        injects, distincts, case_thms, case_cong, case_cong_weak, split,
a88eb33058f7 refactor fp_sugar move theorems
desharna
parents: 58459
diff changeset
   324
        split_asm, ...}, ...}, ...} : fp_sugar) =
55539
0819931d652d simplified data structure by reducing the incidence of clumsy indices
blanchet
parents: 55531
diff changeset
   325
      (T_name0,
58211
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
   326
       {index = kk, descr = descr, inject = injects, distinct = distincts, induct = induct',
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
   327
        inducts = inducts', exhaust = exhaust, nchotomy = nchotomy, rec_names = rec'_names,
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
   328
        rec_rewrites = rec'_thms, case_name = fst (dest_Const casex), case_rewrites = case_thms,
57983
6edc3529bb4e reordered some (co)datatype property names for more consistency
blanchet
parents: 57807
diff changeset
   329
        case_cong = case_cong, case_cong_weak = case_cong_weak, split = split,
57794
73052169b213 tuning whitespace
blanchet
parents: 57633
diff changeset
   330
        split_asm = split_asm});
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   331
58217
blanchet
parents: 58214
diff changeset
   332
    val infos = map_index mk_info (take nn_fp fp_sugars');
58117
9608028d8f43 more compatibility between old- and new-style datatypes
blanchet
parents: 58113
diff changeset
   333
  in
58217
blanchet
parents: 58214
diff changeset
   334
    (nn, b_names, compat_b_names, lfp_sugar_thms'', infos, lthy'')
58117
9608028d8f43 more compatibility between old- and new-style datatypes
blanchet
parents: 58113
diff changeset
   335
  end;
9608028d8f43 more compatibility between old- and new-style datatypes
blanchet
parents: 58113
diff changeset
   336
58354
04ac60da613e support (finite values of) codatatypes in Quickcheck
blanchet
parents: 58340
diff changeset
   337
fun infos_of_new_datatype_mutual_cluster lthy prefs fpT_name =
58358
cdce4471d590 tweaked compatibility layer
blanchet
parents: 58354
diff changeset
   338
  let
cdce4471d590 tweaked compatibility layer
blanchet
parents: 58354
diff changeset
   339
    fun get prefs =
cdce4471d590 tweaked compatibility layer
blanchet
parents: 58354
diff changeset
   340
      #5 (mk_infos_of_mutually_recursive_new_datatypes prefs subset [fpT_name] lthy)
cdce4471d590 tweaked compatibility layer
blanchet
parents: 58354
diff changeset
   341
      handle ERROR _ => [];
cdce4471d590 tweaked compatibility layer
blanchet
parents: 58354
diff changeset
   342
  in
cdce4471d590 tweaked compatibility layer
blanchet
parents: 58354
diff changeset
   343
    (case get prefs of
cdce4471d590 tweaked compatibility layer
blanchet
parents: 58354
diff changeset
   344
      [] => if member (op =) prefs Keep_Nesting then [] else get (Keep_Nesting :: prefs)
cdce4471d590 tweaked compatibility layer
blanchet
parents: 58354
diff changeset
   345
    | infos => infos)
cdce4471d590 tweaked compatibility layer
blanchet
parents: 58354
diff changeset
   346
  end;
58117
9608028d8f43 more compatibility between old- and new-style datatypes
blanchet
parents: 58113
diff changeset
   347
58354
04ac60da613e support (finite values of) codatatypes in Quickcheck
blanchet
parents: 58340
diff changeset
   348
fun get_all thy prefs =
58117
9608028d8f43 more compatibility between old- and new-style datatypes
blanchet
parents: 58113
diff changeset
   349
  let
59644
wenzelm
parents: 59058
diff changeset
   350
    val ctxt = Proof_Context.init_global thy;
58117
9608028d8f43 more compatibility between old- and new-style datatypes
blanchet
parents: 58113
diff changeset
   351
    val old_info_tab = Old_Datatype_Data.get_all thy;
9608028d8f43 more compatibility between old- and new-style datatypes
blanchet
parents: 58113
diff changeset
   352
    val new_T_names = BNF_FP_Def_Sugar.fp_sugars_of_global thy
9608028d8f43 more compatibility between old- and new-style datatypes
blanchet
parents: 58113
diff changeset
   353
      |> map_filter (try (fn {T = Type (s, _), fp_res_index = 0, ...} => s));
58354
04ac60da613e support (finite values of) codatatypes in Quickcheck
blanchet
parents: 58340
diff changeset
   354
    val new_infos =
59644
wenzelm
parents: 59058
diff changeset
   355
      maps (infos_of_new_datatype_mutual_cluster ctxt (insert (op =) Keep_Nesting prefs))
58354
04ac60da613e support (finite values of) codatatypes in Quickcheck
blanchet
parents: 58340
diff changeset
   356
        new_T_names;
58117
9608028d8f43 more compatibility between old- and new-style datatypes
blanchet
parents: 58113
diff changeset
   357
  in
58354
04ac60da613e support (finite values of) codatatypes in Quickcheck
blanchet
parents: 58340
diff changeset
   358
    fold (if member (op =) prefs Keep_Nesting then Symtab.update else Symtab.default) new_infos
58123
62765d39539f implemented compatibility definition of datatype
blanchet
parents: 58122
diff changeset
   359
      old_info_tab
58117
9608028d8f43 more compatibility between old- and new-style datatypes
blanchet
parents: 58113
diff changeset
   360
  end;
9608028d8f43 more compatibility between old- and new-style datatypes
blanchet
parents: 58113
diff changeset
   361
58354
04ac60da613e support (finite values of) codatatypes in Quickcheck
blanchet
parents: 58340
diff changeset
   362
fun get_one get_old get_new thy prefs x =
04ac60da613e support (finite values of) codatatypes in Quickcheck
blanchet
parents: 58340
diff changeset
   363
  let
04ac60da613e support (finite values of) codatatypes in Quickcheck
blanchet
parents: 58340
diff changeset
   364
    val (get_fst, get_snd) = (get_old thy, get_new thy) |> member (op =) prefs Keep_Nesting ? swap;
04ac60da613e support (finite values of) codatatypes in Quickcheck
blanchet
parents: 58340
diff changeset
   365
  in
58117
9608028d8f43 more compatibility between old- and new-style datatypes
blanchet
parents: 58113
diff changeset
   366
    (case get_fst x of NONE => get_snd x | res => res)
9608028d8f43 more compatibility between old- and new-style datatypes
blanchet
parents: 58113
diff changeset
   367
  end;
9608028d8f43 more compatibility between old- and new-style datatypes
blanchet
parents: 58113
diff changeset
   368
58354
04ac60da613e support (finite values of) codatatypes in Quickcheck
blanchet
parents: 58340
diff changeset
   369
fun get_info_of_new_datatype prefs thy T_name =
59644
wenzelm
parents: 59058
diff changeset
   370
  let val ctxt = Proof_Context.init_global thy in
wenzelm
parents: 59058
diff changeset
   371
    AList.lookup (op =) (infos_of_new_datatype_mutual_cluster ctxt prefs T_name) T_name
58117
9608028d8f43 more compatibility between old- and new-style datatypes
blanchet
parents: 58113
diff changeset
   372
  end;
9608028d8f43 more compatibility between old- and new-style datatypes
blanchet
parents: 58113
diff changeset
   373
58354
04ac60da613e support (finite values of) codatatypes in Quickcheck
blanchet
parents: 58340
diff changeset
   374
fun get_info thy prefs =
04ac60da613e support (finite values of) codatatypes in Quickcheck
blanchet
parents: 58340
diff changeset
   375
  get_one Old_Datatype_Data.get_info (get_info_of_new_datatype prefs) thy prefs;
58117
9608028d8f43 more compatibility between old- and new-style datatypes
blanchet
parents: 58113
diff changeset
   376
58354
04ac60da613e support (finite values of) codatatypes in Quickcheck
blanchet
parents: 58340
diff changeset
   377
fun the_info thy prefs T_name =
04ac60da613e support (finite values of) codatatypes in Quickcheck
blanchet
parents: 58340
diff changeset
   378
  (case get_info thy prefs T_name of
58117
9608028d8f43 more compatibility between old- and new-style datatypes
blanchet
parents: 58113
diff changeset
   379
    SOME info => info
9608028d8f43 more compatibility between old- and new-style datatypes
blanchet
parents: 58113
diff changeset
   380
  | NONE => error ("Unknown datatype " ^ quote T_name));
9608028d8f43 more compatibility between old- and new-style datatypes
blanchet
parents: 58113
diff changeset
   381
58129
3ec65a7f2b50 ported Refute to use new datatypes when possible
blanchet
parents: 58126
diff changeset
   382
fun the_spec thy T_name =
58117
9608028d8f43 more compatibility between old- and new-style datatypes
blanchet
parents: 58113
diff changeset
   383
  let
58354
04ac60da613e support (finite values of) codatatypes in Quickcheck
blanchet
parents: 58340
diff changeset
   384
    val {descr, index, ...} = the_info thy [Keep_Nesting, Include_GFPs] T_name;
58117
9608028d8f43 more compatibility between old- and new-style datatypes
blanchet
parents: 58113
diff changeset
   385
    val (_, Ds, ctrs0) = the (AList.lookup (op =) descr index);
58129
3ec65a7f2b50 ported Refute to use new datatypes when possible
blanchet
parents: 58126
diff changeset
   386
    val tfrees = map Old_Datatype_Aux.dest_DtTFree Ds;
58117
9608028d8f43 more compatibility between old- and new-style datatypes
blanchet
parents: 58113
diff changeset
   387
    val ctrs = map (apsnd (map (Old_Datatype_Aux.typ_of_dtyp descr))) ctrs0;
58129
3ec65a7f2b50 ported Refute to use new datatypes when possible
blanchet
parents: 58126
diff changeset
   388
  in (tfrees, ctrs) end;
58117
9608028d8f43 more compatibility between old- and new-style datatypes
blanchet
parents: 58113
diff changeset
   389
58354
04ac60da613e support (finite values of) codatatypes in Quickcheck
blanchet
parents: 58340
diff changeset
   390
fun the_descr thy prefs (T_names0 as T_name01 :: _) =
58117
9608028d8f43 more compatibility between old- and new-style datatypes
blanchet
parents: 58113
diff changeset
   391
  let
9608028d8f43 more compatibility between old- and new-style datatypes
blanchet
parents: 58113
diff changeset
   392
    fun not_mutually_recursive ss =
9608028d8f43 more compatibility between old- and new-style datatypes
blanchet
parents: 58113
diff changeset
   393
      error ("{" ^ commas ss ^ "} is not a complete set of mutually recursive datatypes");
9608028d8f43 more compatibility between old- and new-style datatypes
blanchet
parents: 58113
diff changeset
   394
58354
04ac60da613e support (finite values of) codatatypes in Quickcheck
blanchet
parents: 58340
diff changeset
   395
    val info = the_info thy prefs T_name01;
58117
9608028d8f43 more compatibility between old- and new-style datatypes
blanchet
parents: 58113
diff changeset
   396
    val descr = #descr info;
9608028d8f43 more compatibility between old- and new-style datatypes
blanchet
parents: 58113
diff changeset
   397
9608028d8f43 more compatibility between old- and new-style datatypes
blanchet
parents: 58113
diff changeset
   398
    val (_, Ds, _) = the (AList.lookup (op =) descr (#index info));
9608028d8f43 more compatibility between old- and new-style datatypes
blanchet
parents: 58113
diff changeset
   399
    val vs = map Old_Datatype_Aux.dest_DtTFree Ds;
9608028d8f43 more compatibility between old- and new-style datatypes
blanchet
parents: 58113
diff changeset
   400
9608028d8f43 more compatibility between old- and new-style datatypes
blanchet
parents: 58113
diff changeset
   401
    fun is_DtTFree (Old_Datatype_Aux.DtTFree _) = true
9608028d8f43 more compatibility between old- and new-style datatypes
blanchet
parents: 58113
diff changeset
   402
      | is_DtTFree _ = false;
9608028d8f43 more compatibility between old- and new-style datatypes
blanchet
parents: 58113
diff changeset
   403
9608028d8f43 more compatibility between old- and new-style datatypes
blanchet
parents: 58113
diff changeset
   404
    val k = find_index (fn (_, (_, dTs, _)) => not (forall is_DtTFree dTs)) descr;
9608028d8f43 more compatibility between old- and new-style datatypes
blanchet
parents: 58113
diff changeset
   405
    val protoTs as (dataTs, _) = chop k descr
59058
a78612c67ec0 renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents: 58684
diff changeset
   406
      |> (apply2 o map)
58117
9608028d8f43 more compatibility between old- and new-style datatypes
blanchet
parents: 58113
diff changeset
   407
        (fn (_, (T_name, Ds, _)) => (T_name, map (Old_Datatype_Aux.typ_of_dtyp descr) Ds));
9608028d8f43 more compatibility between old- and new-style datatypes
blanchet
parents: 58113
diff changeset
   408
9608028d8f43 more compatibility between old- and new-style datatypes
blanchet
parents: 58113
diff changeset
   409
    val T_names = map fst dataTs;
9608028d8f43 more compatibility between old- and new-style datatypes
blanchet
parents: 58113
diff changeset
   410
    val _ = eq_set (op =) (T_names, T_names0) orelse not_mutually_recursive T_names0
9608028d8f43 more compatibility between old- and new-style datatypes
blanchet
parents: 58113
diff changeset
   411
59058
a78612c67ec0 renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents: 58684
diff changeset
   412
    val (Ts, Us) = apply2 (map Type) protoTs;
58117
9608028d8f43 more compatibility between old- and new-style datatypes
blanchet
parents: 58113
diff changeset
   413
9608028d8f43 more compatibility between old- and new-style datatypes
blanchet
parents: 58113
diff changeset
   414
    val names = map Long_Name.base_name T_names;
9608028d8f43 more compatibility between old- and new-style datatypes
blanchet
parents: 58113
diff changeset
   415
    val (auxnames, _) = Name.make_context names
9608028d8f43 more compatibility between old- and new-style datatypes
blanchet
parents: 58113
diff changeset
   416
      |> fold_map (Name.variant o Old_Datatype_Aux.name_of_typ) Us;
9608028d8f43 more compatibility between old- and new-style datatypes
blanchet
parents: 58113
diff changeset
   417
    val prefix = space_implode "_" names;
9608028d8f43 more compatibility between old- and new-style datatypes
blanchet
parents: 58113
diff changeset
   418
  in
9608028d8f43 more compatibility between old- and new-style datatypes
blanchet
parents: 58113
diff changeset
   419
    (descr, vs, T_names, prefix, (names, auxnames), (Ts, Us))
9608028d8f43 more compatibility between old- and new-style datatypes
blanchet
parents: 58113
diff changeset
   420
  end;
9608028d8f43 more compatibility between old- and new-style datatypes
blanchet
parents: 58113
diff changeset
   421
58129
3ec65a7f2b50 ported Refute to use new datatypes when possible
blanchet
parents: 58126
diff changeset
   422
fun get_constrs thy T_name =
3ec65a7f2b50 ported Refute to use new datatypes when possible
blanchet
parents: 58126
diff changeset
   423
  try (the_spec thy) T_name
58119
blanchet
parents: 58117
diff changeset
   424
  |> Option.map (fn (tfrees, ctrs) =>
blanchet
parents: 58117
diff changeset
   425
    let
blanchet
parents: 58117
diff changeset
   426
      fun varify_tfree (s, S) = TVar ((s, 0), S);
blanchet
parents: 58117
diff changeset
   427
      fun varify_typ (TFree x) = varify_tfree x
blanchet
parents: 58117
diff changeset
   428
        | varify_typ T = T;
blanchet
parents: 58117
diff changeset
   429
blanchet
parents: 58117
diff changeset
   430
      val dataT = Type (T_name, map varify_tfree tfrees);
blanchet
parents: 58117
diff changeset
   431
blanchet
parents: 58117
diff changeset
   432
      fun mk_ctr_typ Ts = map (Term.map_atyps varify_typ) Ts ---> dataT;
blanchet
parents: 58117
diff changeset
   433
    in
blanchet
parents: 58117
diff changeset
   434
      map (apsnd mk_ctr_typ) ctrs
blanchet
parents: 58117
diff changeset
   435
    end);
58117
9608028d8f43 more compatibility between old- and new-style datatypes
blanchet
parents: 58113
diff changeset
   436
58354
04ac60da613e support (finite values of) codatatypes in Quickcheck
blanchet
parents: 58340
diff changeset
   437
fun old_interpretation_of prefs f config T_names thy =
04ac60da613e support (finite values of) codatatypes in Quickcheck
blanchet
parents: 58340
diff changeset
   438
  if not (member (op =) prefs Keep_Nesting) orelse
04ac60da613e support (finite values of) codatatypes in Quickcheck
blanchet
parents: 58340
diff changeset
   439
     exists (is_none o fp_sugar_of_global thy) T_names then
58122
eaac3e01158a implemented compatibility interpretation
blanchet
parents: 58119
diff changeset
   440
    f config T_names thy
eaac3e01158a implemented compatibility interpretation
blanchet
parents: 58119
diff changeset
   441
  else
eaac3e01158a implemented compatibility interpretation
blanchet
parents: 58119
diff changeset
   442
    thy;
eaac3e01158a implemented compatibility interpretation
blanchet
parents: 58119
diff changeset
   443
58354
04ac60da613e support (finite values of) codatatypes in Quickcheck
blanchet
parents: 58340
diff changeset
   444
fun new_interpretation_of prefs f (fp_sugars : fp_sugar list) thy =
58122
eaac3e01158a implemented compatibility interpretation
blanchet
parents: 58119
diff changeset
   445
  let val T_names = map (fst o dest_Type o #T) fp_sugars in
58354
04ac60da613e support (finite values of) codatatypes in Quickcheck
blanchet
parents: 58340
diff changeset
   446
    if (member (op =) prefs Include_GFPs orelse forall (curry (op =) Least_FP o #fp) fp_sugars)
04ac60da613e support (finite values of) codatatypes in Quickcheck
blanchet
parents: 58340
diff changeset
   447
       andalso (member (op =) prefs Keep_Nesting orelse
58146
d91c1e50b36e codatatypes are not datatypes
blanchet
parents: 58137
diff changeset
   448
         exists (is_none o Old_Datatype_Data.get_info thy) T_names) then
58123
62765d39539f implemented compatibility definition of datatype
blanchet
parents: 58122
diff changeset
   449
      f Old_Datatype_Aux.default_config T_names thy
58122
eaac3e01158a implemented compatibility interpretation
blanchet
parents: 58119
diff changeset
   450
    else
eaac3e01158a implemented compatibility interpretation
blanchet
parents: 58119
diff changeset
   451
      thy
eaac3e01158a implemented compatibility interpretation
blanchet
parents: 58119
diff changeset
   452
  end;
eaac3e01158a implemented compatibility interpretation
blanchet
parents: 58119
diff changeset
   453
58354
04ac60da613e support (finite values of) codatatypes in Quickcheck
blanchet
parents: 58340
diff changeset
   454
fun interpretation name prefs f =
58659
6c9821c32dd5 Local_Interpretation is superseded by Plugin with formal Plugin_Name management, avoiding undeclared strings;
wenzelm
parents: 58634
diff changeset
   455
  Old_Datatype_Data.interpretation (old_interpretation_of prefs f)
6c9821c32dd5 Local_Interpretation is superseded by Plugin with formal Plugin_Name management, avoiding undeclared strings;
wenzelm
parents: 58634
diff changeset
   456
  #> fp_sugars_interpretation name (Local_Theory.background_theory o new_interpretation_of prefs f);
58117
9608028d8f43 more compatibility between old- and new-style datatypes
blanchet
parents: 58113
diff changeset
   457
9608028d8f43 more compatibility between old- and new-style datatypes
blanchet
parents: 58113
diff changeset
   458
val code_nitpicksimp_simp_attrs = Code.add_default_eqn_attrib :: @{attributes [nitpick_simp, simp]};
9608028d8f43 more compatibility between old- and new-style datatypes
blanchet
parents: 58113
diff changeset
   459
58125
a2ba381607fb more work on compatibility interfaces
blanchet
parents: 58123
diff changeset
   460
fun datatype_compat fpT_names lthy =
58117
9608028d8f43 more compatibility between old- and new-style datatypes
blanchet
parents: 58113
diff changeset
   461
  let
9608028d8f43 more compatibility between old- and new-style datatypes
blanchet
parents: 58113
diff changeset
   462
    val (nn, b_names, compat_b_names, lfp_sugar_thms, infos, lthy') =
58354
04ac60da613e support (finite values of) codatatypes in Quickcheck
blanchet
parents: 58340
diff changeset
   463
      mk_infos_of_mutually_recursive_new_datatypes [] eq_set fpT_names lthy;
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   464
53746
bd038e48526d have "datatype_new_compat" register induction and recursion theorems in nested case
blanchet
parents: 53309
diff changeset
   465
    val all_notes =
bd038e48526d have "datatype_new_compat" register induction and recursion theorems in nested case
blanchet
parents: 53309
diff changeset
   466
      (case lfp_sugar_thms of
bd038e48526d have "datatype_new_compat" register induction and recursion theorems in nested case
blanchet
parents: 53309
diff changeset
   467
        NONE => []
58214
bd1754377965 properly note theorems for split recursors
blanchet
parents: 58213
diff changeset
   468
      | SOME ((inducts, induct, induct_attrs), (rec_thmss, _)) =>
53746
bd038e48526d have "datatype_new_compat" register induction and recursion theorems in nested case
blanchet
parents: 53309
diff changeset
   469
        let
58211
c1f3fa32d322 extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents: 58189
diff changeset
   470
          val common_name = compat_N ^ mk_common_name b_names;
58117
9608028d8f43 more compatibility between old- and new-style datatypes
blanchet
parents: 58113
diff changeset
   471
53746
bd038e48526d have "datatype_new_compat" register induction and recursion theorems in nested case
blanchet
parents: 53309
diff changeset
   472
          val common_notes =
58214
bd1754377965 properly note theorems for split recursors
blanchet
parents: 58213
diff changeset
   473
            (if nn > 1 then [(inductN, [induct], induct_attrs)] else [])
53746
bd038e48526d have "datatype_new_compat" register induction and recursion theorems in nested case
blanchet
parents: 53309
diff changeset
   474
            |> filter_out (null o #2)
bd038e48526d have "datatype_new_compat" register induction and recursion theorems in nested case
blanchet
parents: 53309
diff changeset
   475
            |> map (fn (thmN, thms, attrs) =>
bd038e48526d have "datatype_new_compat" register induction and recursion theorems in nested case
blanchet
parents: 53309
diff changeset
   476
              ((Binding.qualify true common_name (Binding.name thmN), attrs), [(thms, [])]));
bd038e48526d have "datatype_new_compat" register induction and recursion theorems in nested case
blanchet
parents: 53309
diff changeset
   477
bd038e48526d have "datatype_new_compat" register induction and recursion theorems in nested case
blanchet
parents: 53309
diff changeset
   478
          val notes =
58214
bd1754377965 properly note theorems for split recursors
blanchet
parents: 58213
diff changeset
   479
            [(inductN, map single inducts, induct_attrs),
55479
ece4910c3ea0 improved 'datatype_new_compat': generate more fixpoint equations for types like 'datatype_new x = C (x list) (x list)' (here, one equation for each x list instead of a single for both), for higher compatibility + code generation attributes on the recursor
blanchet
parents: 55409
diff changeset
   480
             (recN, rec_thmss, code_nitpicksimp_simp_attrs)]
53746
bd038e48526d have "datatype_new_compat" register induction and recursion theorems in nested case
blanchet
parents: 53309
diff changeset
   481
            |> filter_out (null o #2)
bd038e48526d have "datatype_new_compat" register induction and recursion theorems in nested case
blanchet
parents: 53309
diff changeset
   482
            |> maps (fn (thmN, thmss, attrs) =>
bd038e48526d have "datatype_new_compat" register induction and recursion theorems in nested case
blanchet
parents: 53309
diff changeset
   483
              if forall null thmss then
bd038e48526d have "datatype_new_compat" register induction and recursion theorems in nested case
blanchet
parents: 53309
diff changeset
   484
                []
bd038e48526d have "datatype_new_compat" register induction and recursion theorems in nested case
blanchet
parents: 53309
diff changeset
   485
              else
bd038e48526d have "datatype_new_compat" register induction and recursion theorems in nested case
blanchet
parents: 53309
diff changeset
   486
                map2 (fn b_name => fn thms =>
bd038e48526d have "datatype_new_compat" register induction and recursion theorems in nested case
blanchet
parents: 53309
diff changeset
   487
                    ((Binding.qualify true b_name (Binding.name thmN), attrs), [(thms, [])]))
bd038e48526d have "datatype_new_compat" register induction and recursion theorems in nested case
blanchet
parents: 53309
diff changeset
   488
                  compat_b_names thmss);
bd038e48526d have "datatype_new_compat" register induction and recursion theorems in nested case
blanchet
parents: 53309
diff changeset
   489
        in
bd038e48526d have "datatype_new_compat" register induction and recursion theorems in nested case
blanchet
parents: 53309
diff changeset
   490
          common_notes @ notes
bd038e48526d have "datatype_new_compat" register induction and recursion theorems in nested case
blanchet
parents: 53309
diff changeset
   491
        end);
bd038e48526d have "datatype_new_compat" register induction and recursion theorems in nested case
blanchet
parents: 53309
diff changeset
   492
bd038e48526d have "datatype_new_compat" register induction and recursion theorems in nested case
blanchet
parents: 53309
diff changeset
   493
    val register_interpret =
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 57983
diff changeset
   494
      Old_Datatype_Data.register infos
58113
blanchet
parents: 58112
diff changeset
   495
      #> Old_Datatype_Data.interpretation_data (Old_Datatype_Aux.default_config, map fst infos);
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   496
  in
58117
9608028d8f43 more compatibility between old- and new-style datatypes
blanchet
parents: 58113
diff changeset
   497
    lthy'
53746
bd038e48526d have "datatype_new_compat" register induction and recursion theorems in nested case
blanchet
parents: 53309
diff changeset
   498
    |> Local_Theory.raw_theory register_interpret
57633
4ff8c090d580 repaired named derivations
blanchet
parents: 56484
diff changeset
   499
    |> Local_Theory.notes all_notes
4ff8c090d580 repaired named derivations
blanchet
parents: 56484
diff changeset
   500
    |> snd
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   501
  end;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   502
58665
50b229a5a097 tuned signature;
wenzelm
parents: 58659
diff changeset
   503
val datatype_compat_global = Named_Target.theory_map o datatype_compat;
58125
a2ba381607fb more work on compatibility interfaces
blanchet
parents: 58123
diff changeset
   504
a2ba381607fb more work on compatibility interfaces
blanchet
parents: 58123
diff changeset
   505
fun datatype_compat_cmd raw_fpT_names lthy =
a2ba381607fb more work on compatibility interfaces
blanchet
parents: 58123
diff changeset
   506
  let
a2ba381607fb more work on compatibility interfaces
blanchet
parents: 58123
diff changeset
   507
    val fpT_names =
a2ba381607fb more work on compatibility interfaces
blanchet
parents: 58123
diff changeset
   508
      map (fst o dest_Type o Proof_Context.read_type_name {proper = true, strict = false} lthy)
a2ba381607fb more work on compatibility interfaces
blanchet
parents: 58123
diff changeset
   509
        raw_fpT_names;
a2ba381607fb more work on compatibility interfaces
blanchet
parents: 58123
diff changeset
   510
  in
a2ba381607fb more work on compatibility interfaces
blanchet
parents: 58123
diff changeset
   511
    datatype_compat fpT_names lthy
a2ba381607fb more work on compatibility interfaces
blanchet
parents: 58123
diff changeset
   512
  end;
a2ba381607fb more work on compatibility interfaces
blanchet
parents: 58123
diff changeset
   513
58354
04ac60da613e support (finite values of) codatatypes in Quickcheck
blanchet
parents: 58340
diff changeset
   514
fun add_datatype prefs old_specs thy =
58123
62765d39539f implemented compatibility definition of datatype
blanchet
parents: 58122
diff changeset
   515
  let
62765d39539f implemented compatibility definition of datatype
blanchet
parents: 58122
diff changeset
   516
    val fpT_names = map (Sign.full_name thy o #1 o fst) old_specs;
62765d39539f implemented compatibility definition of datatype
blanchet
parents: 58122
diff changeset
   517
58300
055afb5f7df8 speed up old Nominal by killing type variables
blanchet
parents: 58271
diff changeset
   518
    fun new_type_args_of (s, S) =
58354
04ac60da613e support (finite values of) codatatypes in Quickcheck
blanchet
parents: 58340
diff changeset
   519
      (if member (op =) prefs Kill_Type_Args then NONE else SOME Binding.empty,
04ac60da613e support (finite values of) codatatypes in Quickcheck
blanchet
parents: 58340
diff changeset
   520
       (TFree (s, @{sort type}), S));
58123
62765d39539f implemented compatibility definition of datatype
blanchet
parents: 58122
diff changeset
   521
    fun new_ctr_spec_of (b, Ts, mx) = (((Binding.empty, b), map (pair Binding.empty) Ts), mx);
62765d39539f implemented compatibility definition of datatype
blanchet
parents: 58122
diff changeset
   522
62765d39539f implemented compatibility definition of datatype
blanchet
parents: 58122
diff changeset
   523
    fun new_spec_of ((b, old_tyargs, mx), old_ctr_specs) =
62765d39539f implemented compatibility definition of datatype
blanchet
parents: 58122
diff changeset
   524
      (((((map new_type_args_of old_tyargs, b), mx), map new_ctr_spec_of old_ctr_specs),
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 60728
diff changeset
   525
        (Binding.empty, Binding.empty, Binding.empty)), []);
58123
62765d39539f implemented compatibility definition of datatype
blanchet
parents: 58122
diff changeset
   526
62765d39539f implemented compatibility definition of datatype
blanchet
parents: 58122
diff changeset
   527
    val new_specs = map new_spec_of old_specs;
62765d39539f implemented compatibility definition of datatype
blanchet
parents: 58122
diff changeset
   528
  in
62765d39539f implemented compatibility definition of datatype
blanchet
parents: 58122
diff changeset
   529
    (fpT_names,
62765d39539f implemented compatibility definition of datatype
blanchet
parents: 58122
diff changeset
   530
     thy
58665
50b229a5a097 tuned signature;
wenzelm
parents: 58659
diff changeset
   531
     |> Named_Target.theory_map
50b229a5a097 tuned signature;
wenzelm
parents: 58659
diff changeset
   532
       (co_datatypes Least_FP construct_lfp (default_ctr_options, new_specs))
58354
04ac60da613e support (finite values of) codatatypes in Quickcheck
blanchet
parents: 58340
diff changeset
   533
     |> not (member (op =) prefs Keep_Nesting) ? perhaps (try (datatype_compat_global fpT_names)))
58123
62765d39539f implemented compatibility definition of datatype
blanchet
parents: 58122
diff changeset
   534
  end;
62765d39539f implemented compatibility definition of datatype
blanchet
parents: 58122
diff changeset
   535
60004
e27e7be1f2f6 have 'primrec' return definitions
blanchet
parents: 60003
diff changeset
   536
fun old_of_new f (ts, _, simpss) = (ts, f simpss);
e27e7be1f2f6 have 'primrec' return definitions
blanchet
parents: 60003
diff changeset
   537
e27e7be1f2f6 have 'primrec' return definitions
blanchet
parents: 60003
diff changeset
   538
val primrec = apfst (old_of_new flat) ooo BNF_LFP_Rec_Sugar.primrec;
e27e7be1f2f6 have 'primrec' return definitions
blanchet
parents: 60003
diff changeset
   539
val primrec_global = apfst (old_of_new flat) ooo BNF_LFP_Rec_Sugar.primrec_global;
e27e7be1f2f6 have 'primrec' return definitions
blanchet
parents: 60003
diff changeset
   540
val primrec_overloaded = apfst (old_of_new flat) oooo BNF_LFP_Rec_Sugar.primrec_overloaded;
62497
5b5b704f4811 respect qualification when noting theorems in prim(co)rec
traytel
parents: 62324
diff changeset
   541
val primrec_simple = apfst (apfst fst o apsnd (old_of_new (flat o snd))) ooo
5b5b704f4811 respect qualification when noting theorems in prim(co)rec
traytel
parents: 62324
diff changeset
   542
  BNF_LFP_Rec_Sugar.primrec_simple;
58126
3831312eb476 added primrec compatibility function
blanchet
parents: 58125
diff changeset
   543
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   544
val _ =
59936
b8ffc3dc9e24 @{command_spec} is superseded by @{command_keyword};
wenzelm
parents: 59644
diff changeset
   545
  Outer_Syntax.local_theory @{command_keyword datatype_compat}
58315
6d8458bc6e27 tuning terminology
blanchet
parents: 58300
diff changeset
   546
    "register datatypes as old-style datatypes and derive old-style properties"
55531
601ca8efa000 renamed 'datatype_new_compat' to 'datatype_compat'
blanchet
parents: 55486
diff changeset
   547
    (Scan.repeat1 Parse.type_const >> datatype_compat_cmd);
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   548
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   549
end;