src/HOL/Codatatype/Tools/bnf_fp_sugar.ML
author blanchet
Thu, 06 Sep 2012 12:21:33 +0200
changeset 49184 83fdea0c4779
parent 49183 0cc46e2dee7e
child 49199 7c9a3c67c55d
permissions -rw-r--r--
gracefully handle shadowing case (fourth step of sugar localization)
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
49112
4de4635d8f93 started work on sugared "(co)data" commands
blanchet
parents:
diff changeset
     1
(*  Title:      HOL/Codatatype/Tools/bnf_fp_sugar.ML
4de4635d8f93 started work on sugared "(co)data" commands
blanchet
parents:
diff changeset
     2
    Author:     Jasmin Blanchette, TU Muenchen
4de4635d8f93 started work on sugared "(co)data" commands
blanchet
parents:
diff changeset
     3
    Copyright   2012
4de4635d8f93 started work on sugared "(co)data" commands
blanchet
parents:
diff changeset
     4
4de4635d8f93 started work on sugared "(co)data" commands
blanchet
parents:
diff changeset
     5
Sugar for constructing LFPs and GFPs.
4de4635d8f93 started work on sugared "(co)data" commands
blanchet
parents:
diff changeset
     6
*)
4de4635d8f93 started work on sugared "(co)data" commands
blanchet
parents:
diff changeset
     7
4de4635d8f93 started work on sugared "(co)data" commands
blanchet
parents:
diff changeset
     8
signature BNF_FP_SUGAR =
4de4635d8f93 started work on sugared "(co)data" commands
blanchet
parents:
diff changeset
     9
sig
4de4635d8f93 started work on sugared "(co)data" commands
blanchet
parents:
diff changeset
    10
end;
4de4635d8f93 started work on sugared "(co)data" commands
blanchet
parents:
diff changeset
    11
4de4635d8f93 started work on sugared "(co)data" commands
blanchet
parents:
diff changeset
    12
structure BNF_FP_Sugar : BNF_FP_SUGAR =
4de4635d8f93 started work on sugared "(co)data" commands
blanchet
parents:
diff changeset
    13
struct
4de4635d8f93 started work on sugared "(co)data" commands
blanchet
parents:
diff changeset
    14
49119
1f605c36869c more work on FP sugar
blanchet
parents: 49112
diff changeset
    15
open BNF_Util
1f605c36869c more work on FP sugar
blanchet
parents: 49112
diff changeset
    16
open BNF_Wrap
1f605c36869c more work on FP sugar
blanchet
parents: 49112
diff changeset
    17
open BNF_FP_Util
1f605c36869c more work on FP sugar
blanchet
parents: 49112
diff changeset
    18
open BNF_LFP
1f605c36869c more work on FP sugar
blanchet
parents: 49112
diff changeset
    19
open BNF_GFP
49123
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents: 49121
diff changeset
    20
open BNF_FP_Sugar_Tactics
49119
1f605c36869c more work on FP sugar
blanchet
parents: 49112
diff changeset
    21
49129
b5413cb7d860 define "case" constant
blanchet
parents: 49127
diff changeset
    22
val caseN = "case";
b5413cb7d860 define "case" constant
blanchet
parents: 49127
diff changeset
    23
49124
968e1b7de057 more work on FP sugar
blanchet
parents: 49123
diff changeset
    24
fun cannot_merge_types () = error "Mutually recursive types must have the same type parameters";
49119
1f605c36869c more work on FP sugar
blanchet
parents: 49112
diff changeset
    25
1f605c36869c more work on FP sugar
blanchet
parents: 49112
diff changeset
    26
fun merge_type_arg_constrained ctxt (T, c) (T', c') =
1f605c36869c more work on FP sugar
blanchet
parents: 49112
diff changeset
    27
  if T = T' then
1f605c36869c more work on FP sugar
blanchet
parents: 49112
diff changeset
    28
    (case (c, c') of
1f605c36869c more work on FP sugar
blanchet
parents: 49112
diff changeset
    29
      (_, NONE) => (T, c)
1f605c36869c more work on FP sugar
blanchet
parents: 49112
diff changeset
    30
    | (NONE, _) => (T, c')
1f605c36869c more work on FP sugar
blanchet
parents: 49112
diff changeset
    31
    | _ =>
1f605c36869c more work on FP sugar
blanchet
parents: 49112
diff changeset
    32
      if c = c' then
1f605c36869c more work on FP sugar
blanchet
parents: 49112
diff changeset
    33
        (T, c)
1f605c36869c more work on FP sugar
blanchet
parents: 49112
diff changeset
    34
      else
1f605c36869c more work on FP sugar
blanchet
parents: 49112
diff changeset
    35
        error ("Inconsistent sort constraints for type variable " ^
1f605c36869c more work on FP sugar
blanchet
parents: 49112
diff changeset
    36
          quote (Syntax.string_of_typ ctxt T)))
1f605c36869c more work on FP sugar
blanchet
parents: 49112
diff changeset
    37
  else
1f605c36869c more work on FP sugar
blanchet
parents: 49112
diff changeset
    38
    cannot_merge_types ();
1f605c36869c more work on FP sugar
blanchet
parents: 49112
diff changeset
    39
1f605c36869c more work on FP sugar
blanchet
parents: 49112
diff changeset
    40
fun merge_type_args_constrained ctxt (cAs, cAs') =
1f605c36869c more work on FP sugar
blanchet
parents: 49112
diff changeset
    41
  if length cAs = length cAs' then map2 (merge_type_arg_constrained ctxt) cAs cAs'
1f605c36869c more work on FP sugar
blanchet
parents: 49112
diff changeset
    42
  else cannot_merge_types ();
1f605c36869c more work on FP sugar
blanchet
parents: 49112
diff changeset
    43
49121
9e0acaa470ab more work on FP sugar
blanchet
parents: 49119
diff changeset
    44
fun type_args_constrained_of (((cAs, _), _), _) = cAs;
9e0acaa470ab more work on FP sugar
blanchet
parents: 49119
diff changeset
    45
val type_args_of = map fst o type_args_constrained_of;
49129
b5413cb7d860 define "case" constant
blanchet
parents: 49127
diff changeset
    46
fun type_binder_of (((_, b), _), _) = b;
49181
blanchet
parents: 49180
diff changeset
    47
fun mixfix_of ((_, mx), _) = mx;
49121
9e0acaa470ab more work on FP sugar
blanchet
parents: 49119
diff changeset
    48
fun ctr_specs_of (_, ctr_specs) = ctr_specs;
49119
1f605c36869c more work on FP sugar
blanchet
parents: 49112
diff changeset
    49
49121
9e0acaa470ab more work on FP sugar
blanchet
parents: 49119
diff changeset
    50
fun disc_of (((disc, _), _), _) = disc;
9e0acaa470ab more work on FP sugar
blanchet
parents: 49119
diff changeset
    51
fun ctr_of (((_, ctr), _), _) = ctr;
9e0acaa470ab more work on FP sugar
blanchet
parents: 49119
diff changeset
    52
fun args_of ((_, args), _) = args;
49181
blanchet
parents: 49180
diff changeset
    53
fun ctr_mixfix_of (_, mx) = mx;
49119
1f605c36869c more work on FP sugar
blanchet
parents: 49112
diff changeset
    54
49134
846264f80f16 optionally provide extra dead variables to the FP constructions
blanchet
parents: 49130
diff changeset
    55
fun prepare_data prepare_typ gfp specs fake_lthy lthy =
49112
4de4635d8f93 started work on sugared "(co)data" commands
blanchet
parents:
diff changeset
    56
  let
49121
9e0acaa470ab more work on FP sugar
blanchet
parents: 49119
diff changeset
    57
    val constrained_As =
9e0acaa470ab more work on FP sugar
blanchet
parents: 49119
diff changeset
    58
      map (map (apfst (prepare_typ fake_lthy)) o type_args_constrained_of) specs
49119
1f605c36869c more work on FP sugar
blanchet
parents: 49112
diff changeset
    59
      |> Library.foldr1 (merge_type_args_constrained lthy);
49121
9e0acaa470ab more work on FP sugar
blanchet
parents: 49119
diff changeset
    60
    val As = map fst constrained_As;
49134
846264f80f16 optionally provide extra dead variables to the FP constructions
blanchet
parents: 49130
diff changeset
    61
    val As' = map dest_TFree As;
49119
1f605c36869c more work on FP sugar
blanchet
parents: 49112
diff changeset
    62
49121
9e0acaa470ab more work on FP sugar
blanchet
parents: 49119
diff changeset
    63
    val _ = (case duplicates (op =) As of [] => ()
49165
c6ccaf6df93c check type variables on rhs
blanchet
parents: 49161
diff changeset
    64
      | A :: _ => error ("Duplicate type parameter " ^ quote (Syntax.string_of_typ lthy A)));
49119
1f605c36869c more work on FP sugar
blanchet
parents: 49112
diff changeset
    65
1f605c36869c more work on FP sugar
blanchet
parents: 49112
diff changeset
    66
    (* TODO: use sort constraints on type args *)
1f605c36869c more work on FP sugar
blanchet
parents: 49112
diff changeset
    67
1f605c36869c more work on FP sugar
blanchet
parents: 49112
diff changeset
    68
    val N = length specs;
1f605c36869c more work on FP sugar
blanchet
parents: 49112
diff changeset
    69
49182
b8517107ffc5 read the real types off the constant types, rather than using the fake parser types (second step of sugar localization)
blanchet
parents: 49181
diff changeset
    70
    fun mk_fake_T b =
49121
9e0acaa470ab more work on FP sugar
blanchet
parents: 49119
diff changeset
    71
      Type (fst (Term.dest_Type (Proof_Context.read_type_name fake_lthy true (Binding.name_of b))),
9e0acaa470ab more work on FP sugar
blanchet
parents: 49119
diff changeset
    72
        As);
9e0acaa470ab more work on FP sugar
blanchet
parents: 49119
diff changeset
    73
49129
b5413cb7d860 define "case" constant
blanchet
parents: 49127
diff changeset
    74
    val bs = map type_binder_of specs;
49182
b8517107ffc5 read the real types off the constant types, rather than using the fake parser types (second step of sugar localization)
blanchet
parents: 49181
diff changeset
    75
    val fake_Ts = map mk_fake_T bs;
49121
9e0acaa470ab more work on FP sugar
blanchet
parents: 49119
diff changeset
    76
49181
blanchet
parents: 49180
diff changeset
    77
    val mixfixes = map mixfix_of specs;
49119
1f605c36869c more work on FP sugar
blanchet
parents: 49112
diff changeset
    78
1f605c36869c more work on FP sugar
blanchet
parents: 49112
diff changeset
    79
    val _ = (case duplicates Binding.eq_name bs of [] => ()
1f605c36869c more work on FP sugar
blanchet
parents: 49112
diff changeset
    80
      | b :: _ => error ("Duplicate type name declaration " ^ quote (Binding.name_of b)));
1f605c36869c more work on FP sugar
blanchet
parents: 49112
diff changeset
    81
49121
9e0acaa470ab more work on FP sugar
blanchet
parents: 49119
diff changeset
    82
    val ctr_specss = map ctr_specs_of specs;
49119
1f605c36869c more work on FP sugar
blanchet
parents: 49112
diff changeset
    83
49129
b5413cb7d860 define "case" constant
blanchet
parents: 49127
diff changeset
    84
    val disc_binderss = map (map disc_of) ctr_specss;
b5413cb7d860 define "case" constant
blanchet
parents: 49127
diff changeset
    85
    val ctr_binderss = map (map ctr_of) ctr_specss;
49121
9e0acaa470ab more work on FP sugar
blanchet
parents: 49119
diff changeset
    86
    val ctr_argsss = map (map args_of) ctr_specss;
49181
blanchet
parents: 49180
diff changeset
    87
    val ctr_mixfixess = map (map ctr_mixfix_of) ctr_specss;
49119
1f605c36869c more work on FP sugar
blanchet
parents: 49112
diff changeset
    88
49129
b5413cb7d860 define "case" constant
blanchet
parents: 49127
diff changeset
    89
    val sel_bindersss = map (map (map fst)) ctr_argsss;
49183
0cc46e2dee7e careful about constructor types w.r.t. fake context (third step of localization)
blanchet
parents: 49182
diff changeset
    90
    val fake_ctr_Tsss = map (map (map (prepare_typ fake_lthy o snd))) ctr_argsss;
49121
9e0acaa470ab more work on FP sugar
blanchet
parents: 49119
diff changeset
    91
49183
0cc46e2dee7e careful about constructor types w.r.t. fake context (third step of localization)
blanchet
parents: 49182
diff changeset
    92
    val rhs_As' = fold (fold (fold Term.add_tfreesT)) fake_ctr_Tsss [];
49167
68623861e0f2 print timing information
blanchet
parents: 49165
diff changeset
    93
    val _ = (case subtract (op =) As' rhs_As' of
49165
c6ccaf6df93c check type variables on rhs
blanchet
parents: 49161
diff changeset
    94
        [] => ()
c6ccaf6df93c check type variables on rhs
blanchet
parents: 49161
diff changeset
    95
      | A' :: _ => error ("Extra type variables on rhs: " ^
c6ccaf6df93c check type variables on rhs
blanchet
parents: 49161
diff changeset
    96
          quote (Syntax.string_of_typ lthy (TFree A'))));
c6ccaf6df93c check type variables on rhs
blanchet
parents: 49161
diff changeset
    97
49176
6d29d2db5f88 construct high-level iterator RHS
blanchet
parents: 49169
diff changeset
    98
    val ((Cs, Xs), _) =
49121
9e0acaa470ab more work on FP sugar
blanchet
parents: 49119
diff changeset
    99
      lthy
9e0acaa470ab more work on FP sugar
blanchet
parents: 49119
diff changeset
   100
      |> fold (fold (fn s => Variable.declare_typ (TFree (s, dummyS))) o type_args_of) specs
9e0acaa470ab more work on FP sugar
blanchet
parents: 49119
diff changeset
   101
      |> mk_TFrees N
49176
6d29d2db5f88 construct high-level iterator RHS
blanchet
parents: 49169
diff changeset
   102
      ||>> mk_TFrees N;
49119
1f605c36869c more work on FP sugar
blanchet
parents: 49112
diff changeset
   103
49176
6d29d2db5f88 construct high-level iterator RHS
blanchet
parents: 49169
diff changeset
   104
    fun is_same_recT (T as Type (s, Us)) (Type (s', Us')) =
49146
e32b1f748854 added a check
blanchet
parents: 49135
diff changeset
   105
        s = s' andalso (Us = Us' orelse error ("Illegal occurrence of recursive type " ^
e32b1f748854 added a check
blanchet
parents: 49135
diff changeset
   106
          quote (Syntax.string_of_typ fake_lthy T)))
49176
6d29d2db5f88 construct high-level iterator RHS
blanchet
parents: 49169
diff changeset
   107
      | is_same_recT _ _ = false;
49146
e32b1f748854 added a check
blanchet
parents: 49135
diff changeset
   108
49176
6d29d2db5f88 construct high-level iterator RHS
blanchet
parents: 49169
diff changeset
   109
    fun freeze_recXs (T as Type (s, Us)) =
49182
b8517107ffc5 read the real types off the constant types, rather than using the fake parser types (second step of sugar localization)
blanchet
parents: 49181
diff changeset
   110
        (case find_index (is_same_recT T) fake_Ts of
49176
6d29d2db5f88 construct high-level iterator RHS
blanchet
parents: 49169
diff changeset
   111
          ~1 => Type (s, map freeze_recXs Us)
6d29d2db5f88 construct high-level iterator RHS
blanchet
parents: 49169
diff changeset
   112
        | i => nth Xs i)
6d29d2db5f88 construct high-level iterator RHS
blanchet
parents: 49169
diff changeset
   113
      | freeze_recXs T = T;
49121
9e0acaa470ab more work on FP sugar
blanchet
parents: 49119
diff changeset
   114
49183
0cc46e2dee7e careful about constructor types w.r.t. fake context (third step of localization)
blanchet
parents: 49182
diff changeset
   115
    val ctr_TsssXs = map (map (map freeze_recXs)) fake_ctr_Tsss;
49176
6d29d2db5f88 construct high-level iterator RHS
blanchet
parents: 49169
diff changeset
   116
    val sum_prod_TsXs = map (mk_sumTN o map HOLogic.mk_tupleT) ctr_TsssXs;
49119
1f605c36869c more work on FP sugar
blanchet
parents: 49112
diff changeset
   117
49176
6d29d2db5f88 construct high-level iterator RHS
blanchet
parents: 49169
diff changeset
   118
    val eqs = map dest_TFree Xs ~~ sum_prod_TsXs;
49121
9e0acaa470ab more work on FP sugar
blanchet
parents: 49119
diff changeset
   119
49176
6d29d2db5f88 construct high-level iterator RHS
blanchet
parents: 49169
diff changeset
   120
    val ((raw_unfs, raw_flds, raw_fp_iters, raw_fp_recs, unf_flds, fld_unfs, fld_injects), lthy') =
49169
937a0fadddfb honor mixfix specifications
blanchet
parents: 49167
diff changeset
   121
      fp_bnf (if gfp then bnf_gfp else bnf_lfp) bs mixfixes As' eqs lthy;
49121
9e0acaa470ab more work on FP sugar
blanchet
parents: 49119
diff changeset
   122
49167
68623861e0f2 print timing information
blanchet
parents: 49165
diff changeset
   123
    val timer = time (Timer.startRealTimer ());
68623861e0f2 print timing information
blanchet
parents: 49165
diff changeset
   124
49176
6d29d2db5f88 construct high-level iterator RHS
blanchet
parents: 49169
diff changeset
   125
    fun mk_unf_or_fld get_T Ts t =
6d29d2db5f88 construct high-level iterator RHS
blanchet
parents: 49169
diff changeset
   126
      let val Type (_, Ts0) = get_T (fastype_of t) in
49124
968e1b7de057 more work on FP sugar
blanchet
parents: 49123
diff changeset
   127
        Term.subst_atomic_types (Ts0 ~~ Ts) t
49121
9e0acaa470ab more work on FP sugar
blanchet
parents: 49119
diff changeset
   128
      end;
49119
1f605c36869c more work on FP sugar
blanchet
parents: 49112
diff changeset
   129
49126
1bbd7a37fc29 implemented "mk_inject_tac"
blanchet
parents: 49125
diff changeset
   130
    val mk_unf = mk_unf_or_fld domain_type;
1bbd7a37fc29 implemented "mk_inject_tac"
blanchet
parents: 49125
diff changeset
   131
    val mk_fld = mk_unf_or_fld range_type;
49121
9e0acaa470ab more work on FP sugar
blanchet
parents: 49119
diff changeset
   132
49126
1bbd7a37fc29 implemented "mk_inject_tac"
blanchet
parents: 49125
diff changeset
   133
    val unfs = map (mk_unf As) raw_unfs;
49124
968e1b7de057 more work on FP sugar
blanchet
parents: 49123
diff changeset
   134
    val flds = map (mk_fld As) raw_flds;
968e1b7de057 more work on FP sugar
blanchet
parents: 49123
diff changeset
   135
49182
b8517107ffc5 read the real types off the constant types, rather than using the fake parser types (second step of sugar localization)
blanchet
parents: 49181
diff changeset
   136
    val fp_Ts = map (domain_type o fastype_of) unfs;
49183
0cc46e2dee7e careful about constructor types w.r.t. fake context (third step of localization)
blanchet
parents: 49182
diff changeset
   137
    val ctr_Tsss = map (map (map (Term.typ_subst_atomic (Xs ~~ fp_Ts)))) ctr_TsssXs;
49182
b8517107ffc5 read the real types off the constant types, rather than using the fake parser types (second step of sugar localization)
blanchet
parents: 49181
diff changeset
   138
49176
6d29d2db5f88 construct high-level iterator RHS
blanchet
parents: 49169
diff changeset
   139
    fun mk_fp_iter_or_rec Ts Us t =
49121
9e0acaa470ab more work on FP sugar
blanchet
parents: 49119
diff changeset
   140
      let
49176
6d29d2db5f88 construct high-level iterator RHS
blanchet
parents: 49169
diff changeset
   141
        val (binders, body) = strip_type (fastype_of t);
6d29d2db5f88 construct high-level iterator RHS
blanchet
parents: 49169
diff changeset
   142
        val Type (_, Ts0) = if gfp then body else List.last binders;
6d29d2db5f88 construct high-level iterator RHS
blanchet
parents: 49169
diff changeset
   143
        val Us0 = map (if gfp then domain_type else body_type) (fst (split_last binders));
6d29d2db5f88 construct high-level iterator RHS
blanchet
parents: 49169
diff changeset
   144
      in
6d29d2db5f88 construct high-level iterator RHS
blanchet
parents: 49169
diff changeset
   145
        Term.subst_atomic_types (Ts0 @ Us0 ~~ Ts @ Us) t
6d29d2db5f88 construct high-level iterator RHS
blanchet
parents: 49169
diff changeset
   146
      end;
6d29d2db5f88 construct high-level iterator RHS
blanchet
parents: 49169
diff changeset
   147
6d29d2db5f88 construct high-level iterator RHS
blanchet
parents: 49169
diff changeset
   148
    val fp_iters = map (mk_fp_iter_or_rec As Cs) raw_fp_iters;
6d29d2db5f88 construct high-level iterator RHS
blanchet
parents: 49169
diff changeset
   149
    val fp_recs = map (mk_fp_iter_or_rec As Cs) raw_fp_recs;
6d29d2db5f88 construct high-level iterator RHS
blanchet
parents: 49169
diff changeset
   150
6d29d2db5f88 construct high-level iterator RHS
blanchet
parents: 49169
diff changeset
   151
    fun pour_sugar_on_type ((((((((((((((b, fp_T), C), fld), unf), fp_iter), fp_rec), fld_unf),
6d29d2db5f88 construct high-level iterator RHS
blanchet
parents: 49169
diff changeset
   152
          unf_fld), fld_inject), ctr_binders), ctr_mixfixes), ctr_Tss), disc_binders), sel_binderss)
6d29d2db5f88 construct high-level iterator RHS
blanchet
parents: 49169
diff changeset
   153
        no_defs_lthy =
6d29d2db5f88 construct high-level iterator RHS
blanchet
parents: 49169
diff changeset
   154
      let
6d29d2db5f88 construct high-level iterator RHS
blanchet
parents: 49169
diff changeset
   155
        val n = length ctr_Tss;
49121
9e0acaa470ab more work on FP sugar
blanchet
parents: 49119
diff changeset
   156
        val ks = 1 upto n;
9e0acaa470ab more work on FP sugar
blanchet
parents: 49119
diff changeset
   157
        val ms = map length ctr_Tss;
9e0acaa470ab more work on FP sugar
blanchet
parents: 49119
diff changeset
   158
49124
968e1b7de057 more work on FP sugar
blanchet
parents: 49123
diff changeset
   159
        val unf_T = domain_type (fastype_of fld);
49121
9e0acaa470ab more work on FP sugar
blanchet
parents: 49119
diff changeset
   160
        val prod_Ts = map HOLogic.mk_tupleT ctr_Tss;
49134
846264f80f16 optionally provide extra dead variables to the FP constructions
blanchet
parents: 49130
diff changeset
   161
        val case_Ts = map (fn Ts => Ts ---> C) ctr_Tss;
49119
1f605c36869c more work on FP sugar
blanchet
parents: 49112
diff changeset
   162
49176
6d29d2db5f88 construct high-level iterator RHS
blanchet
parents: 49169
diff changeset
   163
        val ((((u, v), fs), xss), _) =
49124
968e1b7de057 more work on FP sugar
blanchet
parents: 49123
diff changeset
   164
          lthy
49176
6d29d2db5f88 construct high-level iterator RHS
blanchet
parents: 49169
diff changeset
   165
          |> yield_singleton (mk_Frees "u") unf_T
6d29d2db5f88 construct high-level iterator RHS
blanchet
parents: 49169
diff changeset
   166
          ||>> yield_singleton (mk_Frees "v") fp_T
6d29d2db5f88 construct high-level iterator RHS
blanchet
parents: 49169
diff changeset
   167
          ||>> mk_Frees "f" case_Ts
49124
968e1b7de057 more work on FP sugar
blanchet
parents: 49123
diff changeset
   168
          ||>> mk_Freess "x" ctr_Tss;
49121
9e0acaa470ab more work on FP sugar
blanchet
parents: 49119
diff changeset
   169
49129
b5413cb7d860 define "case" constant
blanchet
parents: 49127
diff changeset
   170
        val ctr_rhss =
49121
9e0acaa470ab more work on FP sugar
blanchet
parents: 49119
diff changeset
   171
          map2 (fn k => fn xs =>
9e0acaa470ab more work on FP sugar
blanchet
parents: 49119
diff changeset
   172
            fold_rev Term.lambda xs (fld $ mk_InN prod_Ts (HOLogic.mk_tuple xs) k)) ks xss;
9e0acaa470ab more work on FP sugar
blanchet
parents: 49119
diff changeset
   173
49130
3c26e17b2849 implemented "mk_case_tac" -- and got rid of "cheat_tac"
blanchet
parents: 49129
diff changeset
   174
        val case_binder = Binding.suffix_name ("_" ^ caseN) b;
49129
b5413cb7d860 define "case" constant
blanchet
parents: 49127
diff changeset
   175
49134
846264f80f16 optionally provide extra dead variables to the FP constructions
blanchet
parents: 49130
diff changeset
   176
        val case_rhs =
49176
6d29d2db5f88 construct high-level iterator RHS
blanchet
parents: 49169
diff changeset
   177
          fold_rev Term.lambda (fs @ [v]) (mk_sum_caseN (map2 mk_uncurried_fun fs xss) $ (unf $ v));
49129
b5413cb7d860 define "case" constant
blanchet
parents: 49127
diff changeset
   178
49130
3c26e17b2849 implemented "mk_case_tac" -- and got rid of "cheat_tac"
blanchet
parents: 49129
diff changeset
   179
        val (((raw_ctrs, raw_ctr_defs), (raw_case, raw_case_def)), (lthy', lthy)) = no_defs_lthy
49169
937a0fadddfb honor mixfix specifications
blanchet
parents: 49167
diff changeset
   180
          |> apfst split_list o fold_map3 (fn b => fn mx => fn rhs =>
937a0fadddfb honor mixfix specifications
blanchet
parents: 49167
diff changeset
   181
               Local_Theory.define ((b, mx), ((Thm.def_binding b, []), rhs)) #>> apsnd snd)
937a0fadddfb honor mixfix specifications
blanchet
parents: 49167
diff changeset
   182
             ctr_binders ctr_mixfixes ctr_rhss
49130
3c26e17b2849 implemented "mk_case_tac" -- and got rid of "cheat_tac"
blanchet
parents: 49129
diff changeset
   183
          ||>> (Local_Theory.define ((case_binder, NoSyn), ((Thm.def_binding case_binder, []),
3c26e17b2849 implemented "mk_case_tac" -- and got rid of "cheat_tac"
blanchet
parents: 49129
diff changeset
   184
             case_rhs)) #>> apsnd snd)
49121
9e0acaa470ab more work on FP sugar
blanchet
parents: 49119
diff changeset
   185
          ||> `Local_Theory.restore;
9e0acaa470ab more work on FP sugar
blanchet
parents: 49119
diff changeset
   186
9e0acaa470ab more work on FP sugar
blanchet
parents: 49119
diff changeset
   187
        (*transforms defined frees into consts (and more)*)
9e0acaa470ab more work on FP sugar
blanchet
parents: 49119
diff changeset
   188
        val phi = Proof_Context.export_morphism lthy lthy';
9e0acaa470ab more work on FP sugar
blanchet
parents: 49119
diff changeset
   189
9e0acaa470ab more work on FP sugar
blanchet
parents: 49119
diff changeset
   190
        val ctr_defs = map (Morphism.thm phi) raw_ctr_defs;
49130
3c26e17b2849 implemented "mk_case_tac" -- and got rid of "cheat_tac"
blanchet
parents: 49129
diff changeset
   191
        val case_def = Morphism.thm phi raw_case_def;
3c26e17b2849 implemented "mk_case_tac" -- and got rid of "cheat_tac"
blanchet
parents: 49129
diff changeset
   192
49123
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents: 49121
diff changeset
   193
        val ctrs = map (Morphism.term phi) raw_ctrs;
49130
3c26e17b2849 implemented "mk_case_tac" -- and got rid of "cheat_tac"
blanchet
parents: 49129
diff changeset
   194
        val casex = Morphism.term phi raw_case;
49121
9e0acaa470ab more work on FP sugar
blanchet
parents: 49119
diff changeset
   195
49135
de13b454fa31 fixed some type issues in sugar "exhaust_tac"
blanchet
parents: 49134
diff changeset
   196
        fun exhaust_tac {context = ctxt, ...} =
49123
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents: 49121
diff changeset
   197
          let
49135
de13b454fa31 fixed some type issues in sugar "exhaust_tac"
blanchet
parents: 49134
diff changeset
   198
            val fld_iff_unf_thm =
de13b454fa31 fixed some type issues in sugar "exhaust_tac"
blanchet
parents: 49134
diff changeset
   199
              let
de13b454fa31 fixed some type issues in sugar "exhaust_tac"
blanchet
parents: 49134
diff changeset
   200
                val goal =
de13b454fa31 fixed some type issues in sugar "exhaust_tac"
blanchet
parents: 49134
diff changeset
   201
                  fold_rev Logic.all [u, v]
de13b454fa31 fixed some type issues in sugar "exhaust_tac"
blanchet
parents: 49134
diff changeset
   202
                    (mk_Trueprop_eq (HOLogic.mk_eq (v, fld $ u), HOLogic.mk_eq (unf $ v, u)));
de13b454fa31 fixed some type issues in sugar "exhaust_tac"
blanchet
parents: 49134
diff changeset
   203
              in
de13b454fa31 fixed some type issues in sugar "exhaust_tac"
blanchet
parents: 49134
diff changeset
   204
                Skip_Proof.prove lthy [] [] goal (fn {context = ctxt, ...} =>
49176
6d29d2db5f88 construct high-level iterator RHS
blanchet
parents: 49169
diff changeset
   205
                  mk_fld_iff_unf_tac ctxt (map (SOME o certifyT lthy) [unf_T, fp_T])
6d29d2db5f88 construct high-level iterator RHS
blanchet
parents: 49169
diff changeset
   206
                    (certify lthy fld) (certify lthy unf) fld_unf unf_fld)
49135
de13b454fa31 fixed some type issues in sugar "exhaust_tac"
blanchet
parents: 49134
diff changeset
   207
                |> Thm.close_derivation
de13b454fa31 fixed some type issues in sugar "exhaust_tac"
blanchet
parents: 49134
diff changeset
   208
                |> Morphism.thm phi
de13b454fa31 fixed some type issues in sugar "exhaust_tac"
blanchet
parents: 49134
diff changeset
   209
              end;
de13b454fa31 fixed some type issues in sugar "exhaust_tac"
blanchet
parents: 49134
diff changeset
   210
de13b454fa31 fixed some type issues in sugar "exhaust_tac"
blanchet
parents: 49134
diff changeset
   211
            val sumEN_thm' =
de13b454fa31 fixed some type issues in sugar "exhaust_tac"
blanchet
parents: 49134
diff changeset
   212
              Local_Defs.unfold lthy @{thms all_unit_eq}
de13b454fa31 fixed some type issues in sugar "exhaust_tac"
blanchet
parents: 49134
diff changeset
   213
                (Drule.instantiate' (map (SOME o certifyT lthy) prod_Ts) [] (mk_sumEN n))
de13b454fa31 fixed some type issues in sugar "exhaust_tac"
blanchet
parents: 49134
diff changeset
   214
              |> Morphism.thm phi;
49123
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents: 49121
diff changeset
   215
          in
49161
a8e74375d971 fixed (n + 1)st bug in "mk_exhaust_tac" -- arose with uncurried constructors
blanchet
parents: 49157
diff changeset
   216
            mk_exhaust_tac ctxt n ctr_defs fld_iff_unf_thm sumEN_thm'
49123
263b0e330d8b more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents: 49121
diff changeset
   217
          end;
49121
9e0acaa470ab more work on FP sugar
blanchet
parents: 49119
diff changeset
   218
49126
1bbd7a37fc29 implemented "mk_inject_tac"
blanchet
parents: 49125
diff changeset
   219
        val inject_tacss =
1bbd7a37fc29 implemented "mk_inject_tac"
blanchet
parents: 49125
diff changeset
   220
          map2 (fn 0 => K []
1bbd7a37fc29 implemented "mk_inject_tac"
blanchet
parents: 49125
diff changeset
   221
                 | _ => fn ctr_def => [fn {context = ctxt, ...} =>
1bbd7a37fc29 implemented "mk_inject_tac"
blanchet
parents: 49125
diff changeset
   222
                     mk_inject_tac ctxt ctr_def fld_inject])
1bbd7a37fc29 implemented "mk_inject_tac"
blanchet
parents: 49125
diff changeset
   223
            ms ctr_defs;
1bbd7a37fc29 implemented "mk_inject_tac"
blanchet
parents: 49125
diff changeset
   224
49127
f7326a0d7f19 implemented "mk_half_distinct_tac"
blanchet
parents: 49126
diff changeset
   225
        val half_distinct_tacss =
f7326a0d7f19 implemented "mk_half_distinct_tac"
blanchet
parents: 49126
diff changeset
   226
          map (map (fn (def, def') => fn {context = ctxt, ...} =>
f7326a0d7f19 implemented "mk_half_distinct_tac"
blanchet
parents: 49126
diff changeset
   227
            mk_half_distinct_tac ctxt fld_inject [def, def'])) (mk_half_pairss ctr_defs);
f7326a0d7f19 implemented "mk_half_distinct_tac"
blanchet
parents: 49126
diff changeset
   228
49130
3c26e17b2849 implemented "mk_case_tac" -- and got rid of "cheat_tac"
blanchet
parents: 49129
diff changeset
   229
        val case_tacs =
3c26e17b2849 implemented "mk_case_tac" -- and got rid of "cheat_tac"
blanchet
parents: 49129
diff changeset
   230
          map3 (fn k => fn m => fn ctr_def => fn {context = ctxt, ...} =>
3c26e17b2849 implemented "mk_case_tac" -- and got rid of "cheat_tac"
blanchet
parents: 49129
diff changeset
   231
            mk_case_tac ctxt n k m case_def ctr_def unf_fld) ks ms ctr_defs;
49121
9e0acaa470ab more work on FP sugar
blanchet
parents: 49119
diff changeset
   232
9e0acaa470ab more work on FP sugar
blanchet
parents: 49119
diff changeset
   233
        val tacss = [exhaust_tac] :: inject_tacss @ half_distinct_tacss @ [case_tacs];
49134
846264f80f16 optionally provide extra dead variables to the FP constructions
blanchet
parents: 49130
diff changeset
   234
49176
6d29d2db5f88 construct high-level iterator RHS
blanchet
parents: 49169
diff changeset
   235
        (* (co)iterators, (co)recursors, (co)induction *)
6d29d2db5f88 construct high-level iterator RHS
blanchet
parents: 49169
diff changeset
   236
6d29d2db5f88 construct high-level iterator RHS
blanchet
parents: 49169
diff changeset
   237
        val is_recT = member (op =) fp_Ts;
6d29d2db5f88 construct high-level iterator RHS
blanchet
parents: 49169
diff changeset
   238
6d29d2db5f88 construct high-level iterator RHS
blanchet
parents: 49169
diff changeset
   239
        val ns = map length ctr_Tsss;
6d29d2db5f88 construct high-level iterator RHS
blanchet
parents: 49169
diff changeset
   240
        val mss = map (map length) ctr_Tsss;
6d29d2db5f88 construct high-level iterator RHS
blanchet
parents: 49169
diff changeset
   241
        val Css = map2 replicate ns Cs;
6d29d2db5f88 construct high-level iterator RHS
blanchet
parents: 49169
diff changeset
   242
49134
846264f80f16 optionally provide extra dead variables to the FP constructions
blanchet
parents: 49130
diff changeset
   243
        fun sugar_lfp lthy =
846264f80f16 optionally provide extra dead variables to the FP constructions
blanchet
parents: 49130
diff changeset
   244
          let
49176
6d29d2db5f88 construct high-level iterator RHS
blanchet
parents: 49169
diff changeset
   245
            val fp_y_Ts = map domain_type (fst (split_last (binder_types (fastype_of fp_iter))));
6d29d2db5f88 construct high-level iterator RHS
blanchet
parents: 49169
diff changeset
   246
            val y_prod_Tss = map2 dest_sumTN ns fp_y_Ts;
6d29d2db5f88 construct high-level iterator RHS
blanchet
parents: 49169
diff changeset
   247
            val y_Tsss = map2 (map2 dest_tupleT) mss y_prod_Tss;
6d29d2db5f88 construct high-level iterator RHS
blanchet
parents: 49169
diff changeset
   248
            val g_Tss = map2 (map2 (curry (op --->))) y_Tsss Css;
6d29d2db5f88 construct high-level iterator RHS
blanchet
parents: 49169
diff changeset
   249
            val iter_T = flat g_Tss ---> fp_T --> C;
49146
e32b1f748854 added a check
blanchet
parents: 49135
diff changeset
   250
49176
6d29d2db5f88 construct high-level iterator RHS
blanchet
parents: 49169
diff changeset
   251
            val ((gss, ysss), _) =
6d29d2db5f88 construct high-level iterator RHS
blanchet
parents: 49169
diff changeset
   252
              lthy
6d29d2db5f88 construct high-level iterator RHS
blanchet
parents: 49169
diff changeset
   253
              |> mk_Freess "f" g_Tss
49177
db8ce685073f introduced and used "mk_Freesss", and simplified "mk_Freess(')"
blanchet
parents: 49176
diff changeset
   254
              ||>> mk_Freesss "x" y_Tsss;
49134
846264f80f16 optionally provide extra dead variables to the FP constructions
blanchet
parents: 49130
diff changeset
   255
49146
e32b1f748854 added a check
blanchet
parents: 49135
diff changeset
   256
            val iter_rhs =
49176
6d29d2db5f88 construct high-level iterator RHS
blanchet
parents: 49169
diff changeset
   257
              fold_rev (fold_rev Term.lambda) gss
6d29d2db5f88 construct high-level iterator RHS
blanchet
parents: 49169
diff changeset
   258
                (Term.list_comb (fp_iter, map2 (mk_sum_caseN oo map2 mk_uncurried_fun) gss ysss));
49134
846264f80f16 optionally provide extra dead variables to the FP constructions
blanchet
parents: 49130
diff changeset
   259
          in
846264f80f16 optionally provide extra dead variables to the FP constructions
blanchet
parents: 49130
diff changeset
   260
            lthy
846264f80f16 optionally provide extra dead variables to the FP constructions
blanchet
parents: 49130
diff changeset
   261
          end;
846264f80f16 optionally provide extra dead variables to the FP constructions
blanchet
parents: 49130
diff changeset
   262
846264f80f16 optionally provide extra dead variables to the FP constructions
blanchet
parents: 49130
diff changeset
   263
        fun sugar_gfp lthy = lthy;
49119
1f605c36869c more work on FP sugar
blanchet
parents: 49112
diff changeset
   264
      in
49130
3c26e17b2849 implemented "mk_case_tac" -- and got rid of "cheat_tac"
blanchet
parents: 49129
diff changeset
   265
        wrap_data tacss ((ctrs, casex), (disc_binders, sel_binderss)) lthy'
49134
846264f80f16 optionally provide extra dead variables to the FP constructions
blanchet
parents: 49130
diff changeset
   266
        |> (if gfp then sugar_gfp else sugar_lfp)
49119
1f605c36869c more work on FP sugar
blanchet
parents: 49112
diff changeset
   267
      end;
49167
68623861e0f2 print timing information
blanchet
parents: 49165
diff changeset
   268
68623861e0f2 print timing information
blanchet
parents: 49165
diff changeset
   269
    val lthy'' =
49176
6d29d2db5f88 construct high-level iterator RHS
blanchet
parents: 49169
diff changeset
   270
      fold pour_sugar_on_type (bs ~~ fp_Ts ~~ Cs ~~ flds ~~ unfs ~~ fp_iters ~~ fp_recs ~~
6d29d2db5f88 construct high-level iterator RHS
blanchet
parents: 49169
diff changeset
   271
        fld_unfs ~~ unf_flds ~~ fld_injects ~~ ctr_binderss ~~ ctr_mixfixess ~~ ctr_Tsss ~~
6d29d2db5f88 construct high-level iterator RHS
blanchet
parents: 49169
diff changeset
   272
        disc_binderss ~~ sel_bindersss) lthy';
49167
68623861e0f2 print timing information
blanchet
parents: 49165
diff changeset
   273
68623861e0f2 print timing information
blanchet
parents: 49165
diff changeset
   274
    val timer = time (timer ("Constructors, discriminators, selectors, etc., for the new " ^
68623861e0f2 print timing information
blanchet
parents: 49165
diff changeset
   275
      (if gfp then "co" else "") ^ "datatype"));
49112
4de4635d8f93 started work on sugared "(co)data" commands
blanchet
parents:
diff changeset
   276
  in
49167
68623861e0f2 print timing information
blanchet
parents: 49165
diff changeset
   277
    (timer; lthy'')
49112
4de4635d8f93 started work on sugared "(co)data" commands
blanchet
parents:
diff changeset
   278
  end;
4de4635d8f93 started work on sugared "(co)data" commands
blanchet
parents:
diff changeset
   279
49121
9e0acaa470ab more work on FP sugar
blanchet
parents: 49119
diff changeset
   280
fun data_cmd info specs lthy =
9e0acaa470ab more work on FP sugar
blanchet
parents: 49119
diff changeset
   281
  let
49184
83fdea0c4779 gracefully handle shadowing case (fourth step of sugar localization)
blanchet
parents: 49183
diff changeset
   282
    (*the "perhaps o try" below helps gracefully handles the case where the new type is defined in a
83fdea0c4779 gracefully handle shadowing case (fourth step of sugar localization)
blanchet
parents: 49183
diff changeset
   283
      locale and shadows an existing global type*)
49179
f9d48d479c84 don't throw away the context when hacking the theory (first step to localize the sugar code)
blanchet
parents: 49177
diff changeset
   284
    val fake_thy = Theory.copy
49184
83fdea0c4779 gracefully handle shadowing case (fourth step of sugar localization)
blanchet
parents: 49183
diff changeset
   285
      #> fold (fn spec => perhaps (try (Sign.add_type lthy
83fdea0c4779 gracefully handle shadowing case (fourth step of sugar localization)
blanchet
parents: 49183
diff changeset
   286
        (type_binder_of spec, length (type_args_constrained_of spec), mixfix_of spec)))) specs;
49179
f9d48d479c84 don't throw away the context when hacking the theory (first step to localize the sugar code)
blanchet
parents: 49177
diff changeset
   287
    val fake_lthy = Proof_Context.background_theory fake_thy lthy;
49121
9e0acaa470ab more work on FP sugar
blanchet
parents: 49119
diff changeset
   288
  in
9e0acaa470ab more work on FP sugar
blanchet
parents: 49119
diff changeset
   289
    prepare_data Syntax.read_typ info specs fake_lthy lthy
9e0acaa470ab more work on FP sugar
blanchet
parents: 49119
diff changeset
   290
  end;
49119
1f605c36869c more work on FP sugar
blanchet
parents: 49112
diff changeset
   291
49129
b5413cb7d860 define "case" constant
blanchet
parents: 49127
diff changeset
   292
val parse_opt_binding_colon = Scan.optional (Parse.binding --| Parse.$$$ ":") no_binder
49119
1f605c36869c more work on FP sugar
blanchet
parents: 49112
diff changeset
   293
49112
4de4635d8f93 started work on sugared "(co)data" commands
blanchet
parents:
diff changeset
   294
val parse_ctr_arg =
49119
1f605c36869c more work on FP sugar
blanchet
parents: 49112
diff changeset
   295
  Parse.$$$ "(" |-- parse_opt_binding_colon -- Parse.typ --| Parse.$$$ ")" ||
49129
b5413cb7d860 define "case" constant
blanchet
parents: 49127
diff changeset
   296
  (Parse.typ >> pair no_binder);
49112
4de4635d8f93 started work on sugared "(co)data" commands
blanchet
parents:
diff changeset
   297
4de4635d8f93 started work on sugared "(co)data" commands
blanchet
parents:
diff changeset
   298
val parse_single_spec =
4de4635d8f93 started work on sugared "(co)data" commands
blanchet
parents:
diff changeset
   299
  Parse.type_args_constrained -- Parse.binding -- Parse.opt_mixfix --
49119
1f605c36869c more work on FP sugar
blanchet
parents: 49112
diff changeset
   300
  (@{keyword "="} |-- Parse.enum1 "|" (parse_opt_binding_colon -- Parse.binding --
1f605c36869c more work on FP sugar
blanchet
parents: 49112
diff changeset
   301
    Scan.repeat parse_ctr_arg -- Parse.opt_mixfix));
49112
4de4635d8f93 started work on sugared "(co)data" commands
blanchet
parents:
diff changeset
   302
4de4635d8f93 started work on sugared "(co)data" commands
blanchet
parents:
diff changeset
   303
val _ =
4de4635d8f93 started work on sugared "(co)data" commands
blanchet
parents:
diff changeset
   304
  Outer_Syntax.local_theory @{command_spec "data"} "define BNF-based inductive datatypes"
49134
846264f80f16 optionally provide extra dead variables to the FP constructions
blanchet
parents: 49130
diff changeset
   305
    (Parse.and_list1 parse_single_spec >> data_cmd false);
49112
4de4635d8f93 started work on sugared "(co)data" commands
blanchet
parents:
diff changeset
   306
4de4635d8f93 started work on sugared "(co)data" commands
blanchet
parents:
diff changeset
   307
val _ =
4de4635d8f93 started work on sugared "(co)data" commands
blanchet
parents:
diff changeset
   308
  Outer_Syntax.local_theory @{command_spec "codata"} "define BNF-based coinductive datatypes"
49134
846264f80f16 optionally provide extra dead variables to the FP constructions
blanchet
parents: 49130
diff changeset
   309
    (Parse.and_list1 parse_single_spec >> data_cmd true);
49112
4de4635d8f93 started work on sugared "(co)data" commands
blanchet
parents:
diff changeset
   310
4de4635d8f93 started work on sugared "(co)data" commands
blanchet
parents:
diff changeset
   311
end;