src/HOL/Tools/Datatype/datatype_realizer.ML
author wenzelm
Sun, 07 Mar 2010 12:19:47 +0100
changeset 35625 9c818cab0dd0
parent 35364 b8c62d60195c
child 35845 e5980f0ad025
permissions -rw-r--r--
modernized structure Object_Logic;
haftmann@33967
     1
(*  Title:      HOL/Tools/Datatype/datatype_realizer.ML
berghofe@13467
     2
    Author:     Stefan Berghofer, TU Muenchen
berghofe@13467
     3
haftmann@33967
     4
Program extraction from proofs involving datatypes:
haftmann@33967
     5
realizers for induction and case analysis.
berghofe@13467
     6
*)
berghofe@13467
     7
berghofe@13467
     8
signature DATATYPE_REALIZER =
berghofe@13467
     9
sig
haftmann@31737
    10
  val add_dt_realizers: Datatype.config -> string list -> theory -> theory
haftmann@24699
    11
  val setup: theory -> theory
berghofe@13467
    12
end;
berghofe@13467
    13
haftmann@33967
    14
structure Datatype_Realizer : DATATYPE_REALIZER =
berghofe@13467
    15
struct
berghofe@13467
    16
haftmann@33967
    17
open Datatype_Aux;
berghofe@13467
    18
wenzelm@35364
    19
fun subsets i j =
wenzelm@35364
    20
  if i <= j then
wenzelm@35364
    21
    let val is = subsets (i+1) j
wenzelm@35364
    22
    in map (fn ks => i::ks) is @ is end
wenzelm@35364
    23
  else [[]];
berghofe@13467
    24
berghofe@13467
    25
fun forall_intr_prf (t, prf) =
berghofe@13467
    26
  let val (a, T) = (case t of Var ((a, _), T) => (a, T) | Free p => p)
skalberg@15531
    27
  in Abst (a, SOME T, Proofterm.prf_abstract_over t prf) end;
berghofe@13467
    28
berghofe@13467
    29
fun prf_of thm =
wenzelm@28814
    30
  Reconstruct.reconstruct_proof (Thm.theory_of_thm thm) (Thm.prop_of thm) (Thm.proof_of thm);
berghofe@13467
    31
berghofe@13467
    32
fun is_unit t = snd (strip_type (fastype_of t)) = HOLogic.unitT;
berghofe@13467
    33
berghofe@13725
    34
fun tname_of (Type (s, _)) = s
berghofe@13725
    35
  | tname_of _ = "";
berghofe@13725
    36
haftmann@32724
    37
fun make_ind sorts ({descr, rec_names, rec_rewrites, induct, ...} : info) is thy =
berghofe@13467
    38
  let
berghofe@13467
    39
    val recTs = get_rec_types descr sorts;
berghofe@13467
    40
    val pnames = if length descr = 1 then ["P"]
berghofe@13467
    41
      else map (fn i => "P" ^ string_of_int i) (1 upto length descr);
berghofe@13467
    42
berghofe@13467
    43
    val rec_result_Ts = map (fn ((i, _), P) =>
berghofe@13467
    44
      if i mem is then TFree ("'" ^ P, HOLogic.typeS) else HOLogic.unitT)
berghofe@13467
    45
        (descr ~~ pnames);
berghofe@13467
    46
berghofe@13467
    47
    fun make_pred i T U r x =
berghofe@13467
    48
      if i mem is then
skalberg@15570
    49
        Free (List.nth (pnames, i), T --> U --> HOLogic.boolT) $ r $ x
skalberg@15570
    50
      else Free (List.nth (pnames, i), U --> HOLogic.boolT) $ x;
berghofe@13467
    51
berghofe@13467
    52
    fun mk_all i s T t =
berghofe@13467
    53
      if i mem is then list_all_free ([(s, T)], t) else t;
berghofe@13467
    54
haftmann@31458
    55
    val (prems, rec_fns) = split_list (flat (fst (fold_map
haftmann@31458
    56
      (fn ((i, (_, _, constrs)), T) => fold_map (fn (cname, cargs) => fn j =>
berghofe@13467
    57
        let
berghofe@13467
    58
          val Ts = map (typ_of_dtyp descr sorts) cargs;
haftmann@33967
    59
          val tnames = Name.variant_list pnames (Datatype_Prop.make_tnames Ts);
haftmann@31458
    60
          val recs = filter (is_rec_type o fst o fst) (cargs ~~ tnames ~~ Ts);
berghofe@13467
    61
          val frees = tnames ~~ Ts;
berghofe@13467
    62
berghofe@13467
    63
          fun mk_prems vs [] = 
berghofe@13467
    64
                let
haftmann@31458
    65
                  val rT = nth (rec_result_Ts) i;
berghofe@13467
    66
                  val vs' = filter_out is_unit vs;
berghofe@13467
    67
                  val f = mk_Free "f" (map fastype_of vs' ---> rT) j;
wenzelm@18929
    68
                  val f' = Envir.eta_contract (list_abs_free
berghofe@13467
    69
                    (map dest_Free vs, if i mem is then list_comb (f, vs')
berghofe@13467
    70
                      else HOLogic.unit));
berghofe@13467
    71
                in (HOLogic.mk_Trueprop (make_pred i rT T (list_comb (f, vs'))
berghofe@13467
    72
                  (list_comb (Const (cname, Ts ---> T), map Free frees))), f')
berghofe@13467
    73
                end
berghofe@13641
    74
            | mk_prems vs (((dt, s), T) :: ds) = 
berghofe@13467
    75
                let
berghofe@13641
    76
                  val k = body_index dt;
berghofe@13641
    77
                  val (Us, U) = strip_type T;
berghofe@13641
    78
                  val i = length Us;
haftmann@31458
    79
                  val rT = nth (rec_result_Ts) k;
berghofe@13641
    80
                  val r = Free ("r" ^ s, Us ---> rT);
berghofe@13467
    81
                  val (p, f) = mk_prems (vs @ [r]) ds
berghofe@13641
    82
                in (mk_all k ("r" ^ s) (Us ---> rT) (Logic.mk_implies
berghofe@13641
    83
                  (list_all (map (pair "x") Us, HOLogic.mk_Trueprop
berghofe@13641
    84
                    (make_pred k rT U (app_bnds r i)
berghofe@13641
    85
                      (app_bnds (Free (s, T)) i))), p)), f)
berghofe@13467
    86
                end
berghofe@13467
    87
haftmann@31458
    88
        in (apfst (curry list_all_free frees) (mk_prems (map Free frees) recs), j + 1) end)
haftmann@31458
    89
          constrs) (descr ~~ recTs) 1)));
berghofe@13467
    90
 
berghofe@13467
    91
    fun mk_proj j [] t = t
berghofe@13467
    92
      | mk_proj j (i :: is) t = if null is then t else
wenzelm@23577
    93
          if (j: int) = i then HOLogic.mk_fst t
berghofe@13467
    94
          else mk_proj j is (HOLogic.mk_snd t);
berghofe@13467
    95
haftmann@33967
    96
    val tnames = Datatype_Prop.make_tnames recTs;
berghofe@13467
    97
    val fTs = map fastype_of rec_fns;
berghofe@13467
    98
    val ps = map (fn ((((i, _), T), U), s) => Abs ("x", T, make_pred i U T
berghofe@13467
    99
      (list_comb (Const (s, fTs ---> T --> U), rec_fns) $ Bound 0) (Bound 0)))
berghofe@13467
   100
        (descr ~~ recTs ~~ rec_result_Ts ~~ rec_names);
berghofe@13467
   101
    val r = if null is then Extraction.nullt else
wenzelm@32952
   102
      foldr1 HOLogic.mk_prod (map_filter (fn (((((i, _), T), U), s), tname) =>
skalberg@15531
   103
        if i mem is then SOME
berghofe@13467
   104
          (list_comb (Const (s, fTs ---> T --> U), rec_fns) $ Free (tname, T))
skalberg@15531
   105
        else NONE) (descr ~~ recTs ~~ rec_result_Ts ~~ rec_names ~~ tnames));
wenzelm@35364
   106
    val concl = HOLogic.mk_Trueprop (foldr1 (HOLogic.mk_binop @{const_name "op &"})
berghofe@13467
   107
      (map (fn ((((i, _), T), U), tname) =>
berghofe@13467
   108
        make_pred i U T (mk_proj i is r) (Free (tname, T)))
berghofe@13467
   109
          (descr ~~ recTs ~~ rec_result_Ts ~~ tnames)));
wenzelm@22578
   110
    val cert = cterm_of thy;
berghofe@13467
   111
    val inst = map (pairself cert) (map head_of (HOLogic.dest_conj
haftmann@32709
   112
      (HOLogic.dest_Trueprop (concl_of induct))) ~~ ps);
berghofe@13467
   113
wenzelm@17959
   114
    val thm = OldGoals.simple_prove_goal_cterm (cert (Logic.list_implies (prems, concl)))
berghofe@13467
   115
      (fn prems =>
berghofe@13467
   116
         [rewrite_goals_tac (map mk_meta_eq [fst_conv, snd_conv]),
haftmann@32709
   117
          rtac (cterm_instantiate inst induct) 1,
wenzelm@35625
   118
          ALLGOALS Object_Logic.atomize_prems_tac,
haftmann@26359
   119
          rewrite_goals_tac (@{thm o_def} :: map mk_meta_eq rec_rewrites),
berghofe@13467
   120
          REPEAT ((resolve_tac prems THEN_ALL_NEW (fn i =>
berghofe@13467
   121
            REPEAT (etac allE i) THEN atac i)) 1)]);
berghofe@13467
   122
haftmann@32709
   123
    val ind_name = Thm.get_name induct;
skalberg@15570
   124
    val vs = map (fn i => List.nth (pnames, i)) is;
haftmann@18358
   125
    val (thm', thy') = thy
wenzelm@30449
   126
      |> Sign.root_path
wenzelm@30449
   127
      |> PureThy.store_thm
wenzelm@30449
   128
        (Binding.qualified_name (space_implode "_" (ind_name :: vs @ ["correctness"])), thm)
wenzelm@24712
   129
      ||> Sign.restore_naming thy;
berghofe@13467
   130
haftmann@33967
   131
    val ivs = rev (Term.add_vars (Logic.varify (Datatype_Prop.make_ind [descr] sorts)) []);
wenzelm@22691
   132
    val rvs = rev (Thm.fold_terms Term.add_vars thm' []);
wenzelm@35364
   133
    val ivs1 = map Var (filter_out (fn (_, T) =>  (* FIXME set (!??) *)
wenzelm@35364
   134
      tname_of (body_type T) mem [@{type_abbrev set}, @{type_name bool}]) ivs);
wenzelm@33035
   135
    val ivs2 = map (fn (ixn, _) => Var (ixn, the (AList.lookup (op =) rvs ixn))) ivs;
berghofe@13467
   136
wenzelm@33345
   137
    val prf =
wenzelm@33345
   138
      List.foldr forall_intr_prf
wenzelm@33345
   139
        (fold_rev (fn (f, p) => fn prf =>
wenzelm@33345
   140
          (case head_of (strip_abs_body f) of
wenzelm@33345
   141
             Free (s, T) =>
wenzelm@33345
   142
               let val T' = Logic.varifyT T
wenzelm@33345
   143
               in Abst (s, SOME T', Proofterm.prf_abstract_over
wenzelm@33345
   144
                 (Var ((s, 0), T')) (AbsP ("H", SOME p, prf)))
wenzelm@33345
   145
               end
wenzelm@33345
   146
          | _ => AbsP ("H", SOME p, prf)))
wenzelm@33345
   147
            (rec_fns ~~ prems_of thm)
wenzelm@33345
   148
            (Proofterm.proof_combP (prf_of thm', map PBound (length prems - 1 downto 0)))) ivs2;
berghofe@13467
   149
wenzelm@33345
   150
    val r' =
wenzelm@33345
   151
      if null is then r
wenzelm@33345
   152
      else Logic.varify (fold_rev lambda
wenzelm@33345
   153
        (map Logic.unvarify ivs1 @ filter_out is_unit
wenzelm@33345
   154
            (map (head_of o strip_abs_body) rec_fns)) r);
berghofe@13467
   155
berghofe@13467
   156
  in Extraction.add_realizers_i [(ind_name, (vs, r', prf))] thy' end;
berghofe@13467
   157
berghofe@13467
   158
haftmann@32709
   159
fun make_casedists sorts ({index, descr, case_name, case_rewrites, exhaust, ...} : info) thy =
berghofe@13467
   160
  let
wenzelm@19806
   161
    val cert = cterm_of thy;
berghofe@13467
   162
    val rT = TFree ("'P", HOLogic.typeS);
berghofe@13467
   163
    val rT' = TVar (("'P", 0), HOLogic.typeS);
berghofe@13467
   164
berghofe@13467
   165
    fun make_casedist_prem T (cname, cargs) =
berghofe@13467
   166
      let
berghofe@13467
   167
        val Ts = map (typ_of_dtyp descr sorts) cargs;
haftmann@33967
   168
        val frees = Name.variant_list ["P", "y"] (Datatype_Prop.make_tnames Ts) ~~ Ts;
berghofe@13467
   169
        val free_ts = map Free frees;
wenzelm@30364
   170
        val r = Free ("r" ^ Long_Name.base_name cname, Ts ---> rT)
berghofe@13467
   171
      in (r, list_all_free (frees, Logic.mk_implies (HOLogic.mk_Trueprop
berghofe@13467
   172
        (HOLogic.mk_eq (Free ("y", T), list_comb (Const (cname, Ts ---> T), free_ts))),
berghofe@13467
   173
          HOLogic.mk_Trueprop (Free ("P", rT --> HOLogic.boolT) $
berghofe@13467
   174
            list_comb (r, free_ts)))))
berghofe@13467
   175
      end;
berghofe@13467
   176
haftmann@17521
   177
    val SOME (_, _, constrs) = AList.lookup (op =) descr index;
skalberg@15570
   178
    val T = List.nth (get_rec_types descr sorts, index);
berghofe@13467
   179
    val (rs, prems) = split_list (map (make_casedist_prem T) constrs);
berghofe@13467
   180
    val r = Const (case_name, map fastype_of rs ---> T --> rT);
berghofe@13467
   181
wenzelm@32013
   182
    val y = Var (("y", 0), Logic.varifyT T);
berghofe@13467
   183
    val y' = Free ("y", T);
berghofe@13467
   184
wenzelm@17959
   185
    val thm = OldGoals.prove_goalw_cterm [] (cert (Logic.list_implies (prems,
berghofe@13467
   186
      HOLogic.mk_Trueprop (Free ("P", rT --> HOLogic.boolT) $
berghofe@13467
   187
        list_comb (r, rs @ [y'])))))
berghofe@13467
   188
      (fn prems =>
haftmann@32709
   189
         [rtac (cterm_instantiate [(cert y, cert y')] exhaust) 1,
berghofe@13467
   190
          ALLGOALS (EVERY'
berghofe@13467
   191
            [asm_simp_tac (HOL_basic_ss addsimps case_rewrites),
berghofe@13467
   192
             resolve_tac prems, asm_simp_tac HOL_basic_ss])]);
berghofe@13467
   193
haftmann@32709
   194
    val exh_name = Thm.get_name exhaust;
haftmann@18358
   195
    val (thm', thy') = thy
wenzelm@30449
   196
      |> Sign.root_path
wenzelm@30449
   197
      |> PureThy.store_thm (Binding.qualified_name (exh_name ^ "_P_correctness"), thm)
wenzelm@24712
   198
      ||> Sign.restore_naming thy;
berghofe@13467
   199
berghofe@13725
   200
    val P = Var (("P", 0), rT' --> HOLogic.boolT);
berghofe@13467
   201
    val prf = forall_intr_prf (y, forall_intr_prf (P,
wenzelm@33345
   202
      fold_rev (fn (p, r) => fn prf =>
wenzelm@33345
   203
          forall_intr_prf (Logic.varify r, AbsP ("H", SOME (Logic.varify p), prf)))
wenzelm@33345
   204
        (prems ~~ rs)
wenzelm@33345
   205
        (Proofterm.proof_combP (prf_of thm', map PBound (length prems - 1 downto 0)))));
wenzelm@32013
   206
    val r' = Logic.varify (Abs ("y", T,
berghofe@13725
   207
      list_abs (map dest_Free rs, list_comb (r,
berghofe@13725
   208
        map Bound ((length rs - 1 downto 0) @ [length rs])))));
berghofe@13467
   209
berghofe@13467
   210
  in Extraction.add_realizers_i
berghofe@13467
   211
    [(exh_name, (["P"], r', prf)),
haftmann@32709
   212
     (exh_name, ([], Extraction.nullt, prf_of exhaust))] thy'
berghofe@13467
   213
  end;
berghofe@13467
   214
haftmann@31668
   215
fun add_dt_realizers config names thy =
wenzelm@25223
   216
  if ! Proofterm.proofs < 2 then thy
haftmann@24699
   217
  else let
haftmann@31668
   218
    val _ = message config "Adding realizers for induction and case analysis ..."
haftmann@31784
   219
    val infos = map (Datatype.the_info thy) names;
haftmann@24699
   220
    val info :: _ = infos;
haftmann@24699
   221
  in
haftmann@24699
   222
    thy
haftmann@24699
   223
    |> fold_rev (make_ind (#sorts info) info) (subsets 0 (length (#descr info) - 1))
haftmann@24699
   224
    |> fold_rev (make_casedists (#sorts info)) infos
haftmann@24699
   225
  end;
berghofe@13467
   226
haftmann@31723
   227
val setup = Datatype.interpretation add_dt_realizers;
berghofe@13467
   228
berghofe@13467
   229
end;