src/HOL/Codatatype/Tools/bnf_fp_sugar.ML
author blanchet
Tue, 04 Sep 2012 16:17:22 +0200
changeset 50141 1bbd7a37fc29
parent 50140 5fc5211cf104
child 50142 f7326a0d7f19
permissions -rw-r--r--
implemented "mk_inject_tac"
blanchet@50127
     1
(*  Title:      HOL/Codatatype/Tools/bnf_fp_sugar.ML
blanchet@50127
     2
    Author:     Jasmin Blanchette, TU Muenchen
blanchet@50127
     3
    Copyright   2012
blanchet@50127
     4
blanchet@50127
     5
Sugar for constructing LFPs and GFPs.
blanchet@50127
     6
*)
blanchet@50127
     7
blanchet@50127
     8
signature BNF_FP_SUGAR =
blanchet@50127
     9
sig
blanchet@50127
    10
end;
blanchet@50127
    11
blanchet@50127
    12
structure BNF_FP_Sugar : BNF_FP_SUGAR =
blanchet@50127
    13
struct
blanchet@50127
    14
blanchet@50134
    15
open BNF_Util
blanchet@50134
    16
open BNF_Wrap
blanchet@50134
    17
open BNF_FP_Util
blanchet@50134
    18
open BNF_LFP
blanchet@50134
    19
open BNF_GFP
blanchet@50138
    20
open BNF_FP_Sugar_Tactics
blanchet@50134
    21
blanchet@50139
    22
fun cannot_merge_types () = error "Mutually recursive types must have the same type parameters";
blanchet@50134
    23
blanchet@50134
    24
fun merge_type_arg_constrained ctxt (T, c) (T', c') =
blanchet@50134
    25
  if T = T' then
blanchet@50134
    26
    (case (c, c') of
blanchet@50134
    27
      (_, NONE) => (T, c)
blanchet@50134
    28
    | (NONE, _) => (T, c')
blanchet@50134
    29
    | _ =>
blanchet@50134
    30
      if c = c' then
blanchet@50134
    31
        (T, c)
blanchet@50134
    32
      else
blanchet@50134
    33
        error ("Inconsistent sort constraints for type variable " ^
blanchet@50134
    34
          quote (Syntax.string_of_typ ctxt T)))
blanchet@50134
    35
  else
blanchet@50134
    36
    cannot_merge_types ();
blanchet@50134
    37
blanchet@50134
    38
fun merge_type_args_constrained ctxt (cAs, cAs') =
blanchet@50134
    39
  if length cAs = length cAs' then map2 (merge_type_arg_constrained ctxt) cAs cAs'
blanchet@50134
    40
  else cannot_merge_types ();
blanchet@50134
    41
blanchet@50136
    42
fun type_args_constrained_of (((cAs, _), _), _) = cAs;
blanchet@50136
    43
val type_args_of = map fst o type_args_constrained_of;
blanchet@50136
    44
fun type_name_of (((_, b), _), _) = b;
blanchet@50136
    45
fun mixfix_of_typ ((_, mx), _) = mx;
blanchet@50136
    46
fun ctr_specs_of (_, ctr_specs) = ctr_specs;
blanchet@50134
    47
blanchet@50136
    48
fun disc_of (((disc, _), _), _) = disc;
blanchet@50136
    49
fun ctr_of (((_, ctr), _), _) = ctr;
blanchet@50136
    50
fun args_of ((_, args), _) = args;
blanchet@50136
    51
fun mixfix_of_ctr (_, mx) = mx;
blanchet@50134
    52
blanchet@50134
    53
val lfp_info = bnf_lfp;
blanchet@50134
    54
val gfp_info = bnf_gfp;
blanchet@50134
    55
blanchet@50136
    56
fun prepare_data prepare_typ construct specs fake_lthy lthy =
blanchet@50127
    57
  let
blanchet@50136
    58
    val constrained_As =
blanchet@50136
    59
      map (map (apfst (prepare_typ fake_lthy)) o type_args_constrained_of) specs
blanchet@50134
    60
      |> Library.foldr1 (merge_type_args_constrained lthy);
blanchet@50136
    61
    val As = map fst constrained_As;
blanchet@50134
    62
blanchet@50136
    63
    val _ = (case duplicates (op =) As of [] => ()
blanchet@50134
    64
      | T :: _ => error ("Duplicate type parameter " ^ quote (Syntax.string_of_typ lthy T)));
blanchet@50134
    65
blanchet@50134
    66
    (* TODO: check that no type variables occur in the rhss that's not in the lhss *)
blanchet@50134
    67
    (* TODO: use sort constraints on type args *)
blanchet@50134
    68
blanchet@50134
    69
    val N = length specs;
blanchet@50134
    70
blanchet@50136
    71
    fun mk_T b =
blanchet@50136
    72
      Type (fst (Term.dest_Type (Proof_Context.read_type_name fake_lthy true (Binding.name_of b))),
blanchet@50136
    73
        As);
blanchet@50136
    74
blanchet@50136
    75
    val bs = map type_name_of specs;
blanchet@50136
    76
    val Ts = map mk_T bs;
blanchet@50136
    77
blanchet@50136
    78
    val mixfixes = map mixfix_of_typ specs;
blanchet@50134
    79
blanchet@50134
    80
    val _ = (case duplicates Binding.eq_name bs of [] => ()
blanchet@50134
    81
      | b :: _ => error ("Duplicate type name declaration " ^ quote (Binding.name_of b)));
blanchet@50134
    82
blanchet@50136
    83
    val ctr_specss = map ctr_specs_of specs;
blanchet@50134
    84
blanchet@50136
    85
    val disc_namess = map (map disc_of) ctr_specss;
blanchet@50136
    86
    val ctr_namess = map (map ctr_of) ctr_specss;
blanchet@50136
    87
    val ctr_argsss = map (map args_of) ctr_specss;
blanchet@50136
    88
    val ctr_mixfixess = map (map mixfix_of_ctr) ctr_specss;
blanchet@50134
    89
blanchet@50134
    90
    val sel_namesss = map (map (map fst)) ctr_argsss;
blanchet@50136
    91
    val ctr_Tsss = map (map (map (prepare_typ fake_lthy o snd))) ctr_argsss;
blanchet@50134
    92
blanchet@50136
    93
    val (Bs, C) =
blanchet@50136
    94
      lthy
blanchet@50136
    95
      |> fold (fold (fn s => Variable.declare_typ (TFree (s, dummyS))) o type_args_of) specs
blanchet@50136
    96
      |> mk_TFrees N
blanchet@50136
    97
      ||> the_single o fst o mk_TFrees 1;
blanchet@50134
    98
blanchet@50136
    99
    fun freeze_rec (T as Type (s, Ts')) =
blanchet@50136
   100
        (case find_index (curry (op =) T) Ts of
blanchet@50136
   101
          ~1 => Type (s, map freeze_rec Ts')
blanchet@50136
   102
        | i => nth Bs i)
blanchet@50136
   103
      | freeze_rec T = T;
blanchet@50134
   104
blanchet@50136
   105
    val ctr_TsssBs = map (map (map freeze_rec)) ctr_Tsss;
blanchet@50136
   106
    val sum_prod_TsBs = map (mk_sumTN o map HOLogic.mk_tupleT) ctr_TsssBs;
blanchet@50134
   107
blanchet@50136
   108
    val eqs = map dest_TFree Bs ~~ sum_prod_TsBs;
blanchet@50136
   109
blanchet@50141
   110
    val ((raw_unfs, raw_flds, unf_flds, fld_unfs, fld_injects), lthy') =
blanchet@50141
   111
      fp_bnf construct bs eqs lthy;
blanchet@50136
   112
blanchet@50141
   113
    fun mk_unf_or_fld get_foldedT Ts t =
blanchet@50139
   114
      let val Type (_, Ts0) = get_foldedT (fastype_of t) in
blanchet@50139
   115
        Term.subst_atomic_types (Ts0 ~~ Ts) t
blanchet@50136
   116
      end;
blanchet@50136
   117
blanchet@50141
   118
    val mk_unf = mk_unf_or_fld domain_type;
blanchet@50141
   119
    val mk_fld = mk_unf_or_fld range_type;
blanchet@50139
   120
blanchet@50141
   121
    val unfs = map (mk_unf As) raw_unfs;
blanchet@50136
   122
    val flds = map (mk_fld As) raw_flds;
blanchet@50136
   123
blanchet@50141
   124
    fun wrap_type (((((((((T, fld), unf), fld_unf), unf_fld), fld_inject), ctr_names), ctr_Tss),
blanchet@50141
   125
        disc_names), sel_namess) no_defs_lthy =
blanchet@50134
   126
      let
blanchet@50136
   127
        val n = length ctr_names;
blanchet@50136
   128
        val ks = 1 upto n;
blanchet@50136
   129
        val ms = map length ctr_Tss;
blanchet@50136
   130
blanchet@50139
   131
        val unf_T = domain_type (fastype_of fld);
blanchet@50139
   132
blanchet@50136
   133
        val prod_Ts = map HOLogic.mk_tupleT ctr_Tss;
blanchet@50136
   134
blanchet@50139
   135
        val (((u, v), xss), _) =
blanchet@50139
   136
          lthy
blanchet@50139
   137
          |> yield_singleton (mk_Frees "u") unf_T
blanchet@50139
   138
          ||>> yield_singleton (mk_Frees "v") T
blanchet@50139
   139
          ||>> mk_Freess "x" ctr_Tss;
blanchet@50136
   140
blanchet@50136
   141
        val rhss =
blanchet@50136
   142
          map2 (fn k => fn xs =>
blanchet@50136
   143
            fold_rev Term.lambda xs (fld $ mk_InN prod_Ts (HOLogic.mk_tuple xs) k)) ks xss;
blanchet@50136
   144
blanchet@50136
   145
        val ((raw_ctrs, raw_ctr_defs), (lthy', lthy)) = no_defs_lthy
blanchet@50136
   146
          |> apfst split_list o fold_map2 (fn b => fn rhs =>
blanchet@50136
   147
               Local_Theory.define ((b, NoSyn), ((Thm.def_binding b, []), rhs)) #>> apsnd snd)
blanchet@50136
   148
             ctr_names rhss
blanchet@50136
   149
          ||> `Local_Theory.restore;
blanchet@50136
   150
blanchet@50136
   151
        val raw_caseof =
blanchet@50136
   152
          Const (@{const_name undefined}, map (fn Ts => Ts ---> C) ctr_Tss ---> T --> C);
blanchet@50136
   153
blanchet@50136
   154
        (*transforms defined frees into consts (and more)*)
blanchet@50136
   155
        val phi = Proof_Context.export_morphism lthy lthy';
blanchet@50136
   156
blanchet@50136
   157
        val ctr_defs = map (Morphism.thm phi) raw_ctr_defs;
blanchet@50138
   158
        val ctrs = map (Morphism.term phi) raw_ctrs;
blanchet@50138
   159
        val caseof = Morphism.term phi raw_caseof;
blanchet@50136
   160
blanchet@50138
   161
        val fld_iff_unf_thm =
blanchet@50138
   162
          let
blanchet@50139
   163
            val goal =
blanchet@50139
   164
              fold_rev Logic.all [u, v]
blanchet@50139
   165
                (mk_Trueprop_eq (HOLogic.mk_eq (v, fld $ u), HOLogic.mk_eq (unf $ v, u)));
blanchet@50138
   166
          in
blanchet@50138
   167
            Skip_Proof.prove lthy [] [] goal (fn {context = ctxt, ...} =>
blanchet@50139
   168
              mk_fld_iff_unf_tac ctxt (map (SOME o certifyT lthy) [unf_T, T]) (certify lthy fld)
blanchet@50138
   169
                (certify lthy unf) fld_unf unf_fld)
blanchet@50140
   170
            |> Thm.close_derivation
blanchet@50138
   171
          end;
blanchet@50136
   172
blanchet@50140
   173
        val sumEN_thm = mk_sumEN n;
blanchet@50140
   174
        val sumEN_thm' =
blanchet@50140
   175
          let val cTs = map (SOME o certifyT lthy) prod_Ts in
blanchet@50140
   176
            Local_Defs.unfold lthy @{thms all_unit_eq} (Drule.instantiate' cTs [] sumEN_thm)
blanchet@50140
   177
          end;
blanchet@50140
   178
blanchet@50140
   179
        fun exhaust_tac {context = ctxt, ...} =
blanchet@50140
   180
          mk_exhaust_tac ctxt n ms ctr_defs fld_iff_unf_thm sumEN_thm';
blanchet@50140
   181
blanchet@50141
   182
        val inject_tacss =
blanchet@50141
   183
          map2 (fn 0 => K []
blanchet@50141
   184
                 | _ => fn ctr_def => [fn {context = ctxt, ...} =>
blanchet@50141
   185
                     mk_inject_tac ctxt ctr_def fld_inject])
blanchet@50141
   186
            ms ctr_defs;
blanchet@50141
   187
blanchet@50141
   188
        (*###*)
blanchet@50136
   189
        fun cheat_tac {context = ctxt, ...} = Skip_Proof.cheat_tac (Proof_Context.theory_of ctxt);
blanchet@50136
   190
blanchet@50136
   191
        val half_distinct_tacss = map (map (K cheat_tac)) (mk_half_pairss ks);
blanchet@50136
   192
blanchet@50136
   193
        val case_tacs = map (K cheat_tac) ks;
blanchet@50136
   194
blanchet@50136
   195
        val tacss = [exhaust_tac] :: inject_tacss @ half_distinct_tacss @ [case_tacs];
blanchet@50134
   196
      in
blanchet@50136
   197
        wrap_data tacss ((ctrs, caseof), (disc_names, sel_namess)) lthy'
blanchet@50134
   198
      end;
blanchet@50127
   199
  in
blanchet@50139
   200
    lthy'
blanchet@50141
   201
    |> fold wrap_type (Ts ~~ flds ~~ unfs ~~ fld_unfs ~~ unf_flds ~~ fld_injects ~~ ctr_namess ~~
blanchet@50141
   202
      ctr_Tsss ~~ disc_namess ~~ sel_namesss)
blanchet@50127
   203
  end;
blanchet@50127
   204
blanchet@50136
   205
fun data_cmd info specs lthy =
blanchet@50136
   206
  let
blanchet@50136
   207
    val fake_lthy =
blanchet@50136
   208
      Proof_Context.theory_of lthy
blanchet@50136
   209
      |> Theory.copy
blanchet@50136
   210
      |> Sign.add_types_global (map (fn spec =>
blanchet@50136
   211
        (type_name_of spec, length (type_args_constrained_of spec), mixfix_of_typ spec)) specs)
blanchet@50136
   212
      |> Proof_Context.init_global
blanchet@50136
   213
  in
blanchet@50136
   214
    prepare_data Syntax.read_typ info specs fake_lthy lthy
blanchet@50136
   215
  end;
blanchet@50134
   216
blanchet@50134
   217
val parse_opt_binding_colon = Scan.optional (Parse.binding --| Parse.$$$ ":") no_name
blanchet@50134
   218
blanchet@50127
   219
val parse_ctr_arg =
blanchet@50134
   220
  Parse.$$$ "(" |-- parse_opt_binding_colon -- Parse.typ --| Parse.$$$ ")" ||
blanchet@50134
   221
  (Parse.typ >> pair no_name);
blanchet@50127
   222
blanchet@50127
   223
val parse_single_spec =
blanchet@50127
   224
  Parse.type_args_constrained -- Parse.binding -- Parse.opt_mixfix --
blanchet@50134
   225
  (@{keyword "="} |-- Parse.enum1 "|" (parse_opt_binding_colon -- Parse.binding --
blanchet@50134
   226
    Scan.repeat parse_ctr_arg -- Parse.opt_mixfix));
blanchet@50127
   227
blanchet@50127
   228
val _ =
blanchet@50127
   229
  Outer_Syntax.local_theory @{command_spec "data"} "define BNF-based inductive datatypes"
blanchet@50134
   230
    (Parse.and_list1 parse_single_spec >> data_cmd lfp_info);
blanchet@50127
   231
blanchet@50127
   232
val _ =
blanchet@50127
   233
  Outer_Syntax.local_theory @{command_spec "codata"} "define BNF-based coinductive datatypes"
blanchet@50134
   234
    (Parse.and_list1 parse_single_spec >> data_cmd gfp_info);
blanchet@50127
   235
blanchet@50127
   236
end;