src/HOL/Tools/Datatype/datatype_realizer.ML
author wenzelm
Thu, 03 Feb 2011 19:21:12 +0100
changeset 41698 90597e044e5f
parent 41423 25df154b8ffc
child 42364 8c674b3b8e44
permissions -rw-r--r--
clarified Proofterm.proofs_enabled;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33338
diff changeset
     1
(*  Title:      HOL/Tools/Datatype/datatype_realizer.ML
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
     2
    Author:     Stefan Berghofer, TU Muenchen
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
     3
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33338
diff changeset
     4
Program extraction from proofs involving datatypes:
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33338
diff changeset
     5
realizers for induction and case analysis.
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
     6
*)
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
     7
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
     8
signature DATATYPE_REALIZER =
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
     9
sig
31737
b3f63611784e simplified names of common datatype types
haftmann
parents: 31723
diff changeset
    10
  val add_dt_realizers: Datatype.config -> string list -> theory -> theory
24699
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 23590
diff changeset
    11
  val setup: theory -> theory
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    12
end;
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    13
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33338
diff changeset
    14
structure Datatype_Realizer : DATATYPE_REALIZER =
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    15
struct
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    16
35364
b8c62d60195c more antiquotations;
wenzelm
parents: 33969
diff changeset
    17
fun subsets i j =
b8c62d60195c more antiquotations;
wenzelm
parents: 33969
diff changeset
    18
  if i <= j then
b8c62d60195c more antiquotations;
wenzelm
parents: 33969
diff changeset
    19
    let val is = subsets (i+1) j
b8c62d60195c more antiquotations;
wenzelm
parents: 33969
diff changeset
    20
    in map (fn ks => i::ks) is @ is end
b8c62d60195c more antiquotations;
wenzelm
parents: 33969
diff changeset
    21
  else [[]];
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    22
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    23
fun prf_of thm =
28814
463c9e9111ae clarified Thm.proof_body_of vs. Thm.proof_of;
wenzelm
parents: 28799
diff changeset
    24
  Reconstruct.reconstruct_proof (Thm.theory_of_thm thm) (Thm.prop_of thm) (Thm.proof_of thm);
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    25
40844
5895c525739d more direct use of binder_types/body_type;
wenzelm
parents: 39557
diff changeset
    26
fun is_unit t = body_type (fastype_of t) = HOLogic.unitT;
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    27
13725
12404b452034 Changed format of realizers / correctness proofs.
berghofe
parents: 13708
diff changeset
    28
fun tname_of (Type (s, _)) = s
12404b452034 Changed format of realizers / correctness proofs.
berghofe
parents: 13708
diff changeset
    29
  | tname_of _ = "";
12404b452034 Changed format of realizers / correctness proofs.
berghofe
parents: 13708
diff changeset
    30
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40844
diff changeset
    31
fun make_ind sorts ({descr, rec_names, rec_rewrites, induct, ...} : Datatype_Aux.info) is thy =
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    32
  let
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40844
diff changeset
    33
    val recTs = Datatype_Aux.get_rec_types descr sorts;
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    34
    val pnames = if length descr = 1 then ["P"]
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    35
      else map (fn i => "P" ^ string_of_int i) (1 upto length descr);
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    36
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    37
    val rec_result_Ts = map (fn ((i, _), P) =>
36692
54b64d4ad524 farewell to old-style mem infixes -- type inference in situations with mem_int and mem_string should provide enough information to resolve the type of (op =)
haftmann
parents: 35845
diff changeset
    38
      if member (op =) is i then TFree ("'" ^ P, HOLogic.typeS) else HOLogic.unitT)
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    39
        (descr ~~ pnames);
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    40
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    41
    fun make_pred i T U r x =
36692
54b64d4ad524 farewell to old-style mem infixes -- type inference in situations with mem_int and mem_string should provide enough information to resolve the type of (op =)
haftmann
parents: 35845
diff changeset
    42
      if member (op =) is i then
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
    43
        Free (List.nth (pnames, i), T --> U --> HOLogic.boolT) $ r $ x
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
    44
      else Free (List.nth (pnames, i), U --> HOLogic.boolT) $ x;
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    45
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    46
    fun mk_all i s T t =
36692
54b64d4ad524 farewell to old-style mem infixes -- type inference in situations with mem_int and mem_string should provide enough information to resolve the type of (op =)
haftmann
parents: 35845
diff changeset
    47
      if member (op =) is i then list_all_free ([(s, T)], t) else t;
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    48
31458
b1cf26f2919b avoid Library.foldl_map
haftmann
parents: 30435
diff changeset
    49
    val (prems, rec_fns) = split_list (flat (fst (fold_map
b1cf26f2919b avoid Library.foldl_map
haftmann
parents: 30435
diff changeset
    50
      (fn ((i, (_, _, constrs)), T) => fold_map (fn (cname, cargs) => fn j =>
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    51
        let
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40844
diff changeset
    52
          val Ts = map (Datatype_Aux.typ_of_dtyp descr sorts) cargs;
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33338
diff changeset
    53
          val tnames = Name.variant_list pnames (Datatype_Prop.make_tnames Ts);
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40844
diff changeset
    54
          val recs = filter (Datatype_Aux.is_rec_type o fst o fst) (cargs ~~ tnames ~~ Ts);
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    55
          val frees = tnames ~~ Ts;
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    56
38977
e71e2c56479c eliminated really ancient OldGoals operations in favour of Goal.prove_internal (with its minimal checks and without normalization of result);
wenzelm
parents: 38795
diff changeset
    57
          fun mk_prems vs [] =
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    58
                let
31458
b1cf26f2919b avoid Library.foldl_map
haftmann
parents: 30435
diff changeset
    59
                  val rT = nth (rec_result_Ts) i;
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    60
                  val vs' = filter_out is_unit vs;
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40844
diff changeset
    61
                  val f = Datatype_Aux.mk_Free "f" (map fastype_of vs' ---> rT) j;
18929
d81435108688 Envir.(beta_)eta_contract;
wenzelm
parents: 18358
diff changeset
    62
                  val f' = Envir.eta_contract (list_abs_free
36692
54b64d4ad524 farewell to old-style mem infixes -- type inference in situations with mem_int and mem_string should provide enough information to resolve the type of (op =)
haftmann
parents: 35845
diff changeset
    63
                    (map dest_Free vs, if member (op =) is i then list_comb (f, vs')
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    64
                      else HOLogic.unit));
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    65
                in (HOLogic.mk_Trueprop (make_pred i rT T (list_comb (f, vs'))
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    66
                  (list_comb (Const (cname, Ts ---> T), map Free frees))), f')
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    67
                end
38977
e71e2c56479c eliminated really ancient OldGoals operations in favour of Goal.prove_internal (with its minimal checks and without normalization of result);
wenzelm
parents: 38795
diff changeset
    68
            | mk_prems vs (((dt, s), T) :: ds) =
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    69
                let
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40844
diff changeset
    70
                  val k = Datatype_Aux.body_index dt;
13641
63d1790a43ed Reimplemented parts of datatype package dealing with datatypes involving
berghofe
parents: 13467
diff changeset
    71
                  val (Us, U) = strip_type T;
63d1790a43ed Reimplemented parts of datatype package dealing with datatypes involving
berghofe
parents: 13467
diff changeset
    72
                  val i = length Us;
31458
b1cf26f2919b avoid Library.foldl_map
haftmann
parents: 30435
diff changeset
    73
                  val rT = nth (rec_result_Ts) k;
13641
63d1790a43ed Reimplemented parts of datatype package dealing with datatypes involving
berghofe
parents: 13467
diff changeset
    74
                  val r = Free ("r" ^ s, Us ---> rT);
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    75
                  val (p, f) = mk_prems (vs @ [r]) ds
13641
63d1790a43ed Reimplemented parts of datatype package dealing with datatypes involving
berghofe
parents: 13467
diff changeset
    76
                in (mk_all k ("r" ^ s) (Us ---> rT) (Logic.mk_implies
63d1790a43ed Reimplemented parts of datatype package dealing with datatypes involving
berghofe
parents: 13467
diff changeset
    77
                  (list_all (map (pair "x") Us, HOLogic.mk_Trueprop
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40844
diff changeset
    78
                    (make_pred k rT U (Datatype_Aux.app_bnds r i)
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40844
diff changeset
    79
                      (Datatype_Aux.app_bnds (Free (s, T)) i))), p)), f)
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    80
                end
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    81
31458
b1cf26f2919b avoid Library.foldl_map
haftmann
parents: 30435
diff changeset
    82
        in (apfst (curry list_all_free frees) (mk_prems (map Free frees) recs), j + 1) end)
b1cf26f2919b avoid Library.foldl_map
haftmann
parents: 30435
diff changeset
    83
          constrs) (descr ~~ recTs) 1)));
38977
e71e2c56479c eliminated really ancient OldGoals operations in favour of Goal.prove_internal (with its minimal checks and without normalization of result);
wenzelm
parents: 38795
diff changeset
    84
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    85
    fun mk_proj j [] t = t
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    86
      | mk_proj j (i :: is) t = if null is then t else
23577
c5b93c69afd3 avoid polymorphic equality;
wenzelm
parents: 22691
diff changeset
    87
          if (j: int) = i then HOLogic.mk_fst t
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    88
          else mk_proj j is (HOLogic.mk_snd t);
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    89
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33338
diff changeset
    90
    val tnames = Datatype_Prop.make_tnames recTs;
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    91
    val fTs = map fastype_of rec_fns;
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    92
    val ps = map (fn ((((i, _), T), U), s) => Abs ("x", T, make_pred i U T
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    93
      (list_comb (Const (s, fTs ---> T --> U), rec_fns) $ Bound 0) (Bound 0)))
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    94
        (descr ~~ recTs ~~ rec_result_Ts ~~ rec_names);
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    95
    val r = if null is then Extraction.nullt else
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32727
diff changeset
    96
      foldr1 HOLogic.mk_prod (map_filter (fn (((((i, _), T), U), s), tname) =>
36692
54b64d4ad524 farewell to old-style mem infixes -- type inference in situations with mem_int and mem_string should provide enough information to resolve the type of (op =)
haftmann
parents: 35845
diff changeset
    97
        if member (op =) is i then SOME
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    98
          (list_comb (Const (s, fTs ---> T --> U), rec_fns) $ Free (tname, T))
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15256
diff changeset
    99
        else NONE) (descr ~~ recTs ~~ rec_result_Ts ~~ rec_names ~~ tnames));
38795
848be46708dc formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
haftmann
parents: 37236
diff changeset
   100
    val concl = HOLogic.mk_Trueprop (foldr1 (HOLogic.mk_binop @{const_name HOL.conj})
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   101
      (map (fn ((((i, _), T), U), tname) =>
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   102
        make_pred i U T (mk_proj i is r) (Free (tname, T)))
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   103
          (descr ~~ recTs ~~ rec_result_Ts ~~ tnames)));
22578
b0eb5652f210 removed obsolete sign_of/sign_of_thm;
wenzelm
parents: 21646
diff changeset
   104
    val cert = cterm_of thy;
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   105
    val inst = map (pairself cert) (map head_of (HOLogic.dest_conj
32712
ec5976f4d3d8 registering split rules and projected induction rules; ML identifiers more close to Isar theorem names
haftmann
parents: 32013
diff changeset
   106
      (HOLogic.dest_Trueprop (concl_of induct))) ~~ ps);
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   107
38977
e71e2c56479c eliminated really ancient OldGoals operations in favour of Goal.prove_internal (with its minimal checks and without normalization of result);
wenzelm
parents: 38795
diff changeset
   108
    val thm = Goal.prove_internal (map cert prems) (cert concl)
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   109
      (fn prems =>
38977
e71e2c56479c eliminated really ancient OldGoals operations in favour of Goal.prove_internal (with its minimal checks and without normalization of result);
wenzelm
parents: 38795
diff changeset
   110
         EVERY [
e71e2c56479c eliminated really ancient OldGoals operations in favour of Goal.prove_internal (with its minimal checks and without normalization of result);
wenzelm
parents: 38795
diff changeset
   111
          rewrite_goals_tac (map mk_meta_eq [@{thm fst_conv}, @{thm snd_conv}]),
32712
ec5976f4d3d8 registering split rules and projected induction rules; ML identifiers more close to Isar theorem names
haftmann
parents: 32013
diff changeset
   112
          rtac (cterm_instantiate inst induct) 1,
35625
9c818cab0dd0 modernized structure Object_Logic;
wenzelm
parents: 35364
diff changeset
   113
          ALLGOALS Object_Logic.atomize_prems_tac,
26359
6d437bde2f1d more antiquotations
haftmann
parents: 25223
diff changeset
   114
          rewrite_goals_tac (@{thm o_def} :: map mk_meta_eq rec_rewrites),
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   115
          REPEAT ((resolve_tac prems THEN_ALL_NEW (fn i =>
38977
e71e2c56479c eliminated really ancient OldGoals operations in favour of Goal.prove_internal (with its minimal checks and without normalization of result);
wenzelm
parents: 38795
diff changeset
   116
            REPEAT (etac allE i) THEN atac i)) 1)])
e71e2c56479c eliminated really ancient OldGoals operations in favour of Goal.prove_internal (with its minimal checks and without normalization of result);
wenzelm
parents: 38795
diff changeset
   117
      |> Drule.export_without_context;
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   118
36744
6e1f3d609a68 renamed Thm.get_name -> Thm.derivation_name and Thm.put_name -> Thm.name_derivation, to emphasize the true nature of these operations;
wenzelm
parents: 36692
diff changeset
   119
    val ind_name = Thm.derivation_name induct;
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   120
    val vs = map (fn i => List.nth (pnames, i)) is;
18358
0a733e11021a re-oriented some result tuples in PureThy
haftmann
parents: 17959
diff changeset
   121
    val (thm', thy') = thy
30435
e62d6ecab6ad explicit Binding.qualified_name -- prevents implicitly qualified bstring;
wenzelm
parents: 30364
diff changeset
   122
      |> Sign.root_path
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 38977
diff changeset
   123
      |> Global_Theory.store_thm
30435
e62d6ecab6ad explicit Binding.qualified_name -- prevents implicitly qualified bstring;
wenzelm
parents: 30364
diff changeset
   124
        (Binding.qualified_name (space_implode "_" (ind_name :: vs @ ["correctness"])), thm)
24712
64ed05609568 proper Sign operations instead of Theory aliases;
wenzelm
parents: 24711
diff changeset
   125
      ||> Sign.restore_naming thy;
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   126
35845
e5980f0ad025 renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
wenzelm
parents: 35625
diff changeset
   127
    val ivs = rev (Term.add_vars (Logic.varify_global (Datatype_Prop.make_ind [descr] sorts)) []);
22691
290454649b8c Thm.fold_terms;
wenzelm
parents: 22596
diff changeset
   128
    val rvs = rev (Thm.fold_terms Term.add_vars thm' []);
37233
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36744
diff changeset
   129
    val ivs1 = map Var (filter_out (fn (_, T) =>
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36744
diff changeset
   130
      @{type_name bool} = tname_of (body_type T)) ivs);
33035
15eab423e573 standardized basic operations on type option;
wenzelm
parents: 32952
diff changeset
   131
    val ivs2 = map (fn (ixn, _) => Var (ixn, the (AList.lookup (op =) rvs ixn))) ivs;
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   132
33338
de76079f973a eliminated some old folds;
wenzelm
parents: 33035
diff changeset
   133
    val prf =
37233
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36744
diff changeset
   134
      Extraction.abs_corr_shyps thy' induct vs ivs2
33338
de76079f973a eliminated some old folds;
wenzelm
parents: 33035
diff changeset
   135
        (fold_rev (fn (f, p) => fn prf =>
de76079f973a eliminated some old folds;
wenzelm
parents: 33035
diff changeset
   136
          (case head_of (strip_abs_body f) of
de76079f973a eliminated some old folds;
wenzelm
parents: 33035
diff changeset
   137
             Free (s, T) =>
35845
e5980f0ad025 renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
wenzelm
parents: 35625
diff changeset
   138
               let val T' = Logic.varifyT_global T
33338
de76079f973a eliminated some old folds;
wenzelm
parents: 33035
diff changeset
   139
               in Abst (s, SOME T', Proofterm.prf_abstract_over
de76079f973a eliminated some old folds;
wenzelm
parents: 33035
diff changeset
   140
                 (Var ((s, 0), T')) (AbsP ("H", SOME p, prf)))
de76079f973a eliminated some old folds;
wenzelm
parents: 33035
diff changeset
   141
               end
de76079f973a eliminated some old folds;
wenzelm
parents: 33035
diff changeset
   142
          | _ => AbsP ("H", SOME p, prf)))
de76079f973a eliminated some old folds;
wenzelm
parents: 33035
diff changeset
   143
            (rec_fns ~~ prems_of thm)
37233
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36744
diff changeset
   144
            (Proofterm.proof_combP (prf_of thm', map PBound (length prems - 1 downto 0))));
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   145
33338
de76079f973a eliminated some old folds;
wenzelm
parents: 33035
diff changeset
   146
    val r' =
de76079f973a eliminated some old folds;
wenzelm
parents: 33035
diff changeset
   147
      if null is then r
35845
e5980f0ad025 renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
wenzelm
parents: 35625
diff changeset
   148
      else Logic.varify_global (fold_rev lambda
e5980f0ad025 renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
wenzelm
parents: 35625
diff changeset
   149
        (map Logic.unvarify_global ivs1 @ filter_out is_unit
33338
de76079f973a eliminated some old folds;
wenzelm
parents: 33035
diff changeset
   150
            (map (head_of o strip_abs_body) rec_fns)) r);
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   151
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   152
  in Extraction.add_realizers_i [(ind_name, (vs, r', prf))] thy' end;
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   153
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   154
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40844
diff changeset
   155
fun make_casedists sorts
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40844
diff changeset
   156
    ({index, descr, case_name, case_rewrites, exhaust, ...} : Datatype_Aux.info) thy =
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   157
  let
19806
f860b7a98445 renamed Type.(un)varifyT to Logic.(un)varifyT;
wenzelm
parents: 18929
diff changeset
   158
    val cert = cterm_of thy;
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   159
    val rT = TFree ("'P", HOLogic.typeS);
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   160
    val rT' = TVar (("'P", 0), HOLogic.typeS);
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   161
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   162
    fun make_casedist_prem T (cname, cargs) =
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   163
      let
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40844
diff changeset
   164
        val Ts = map (Datatype_Aux.typ_of_dtyp descr sorts) cargs;
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33338
diff changeset
   165
        val frees = Name.variant_list ["P", "y"] (Datatype_Prop.make_tnames Ts) ~~ Ts;
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   166
        val free_ts = map Free frees;
30364
577edc39b501 moved basic algebra of long names from structure NameSpace to Long_Name;
wenzelm
parents: 30280
diff changeset
   167
        val r = Free ("r" ^ Long_Name.base_name cname, Ts ---> rT)
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   168
      in (r, list_all_free (frees, Logic.mk_implies (HOLogic.mk_Trueprop
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   169
        (HOLogic.mk_eq (Free ("y", T), list_comb (Const (cname, Ts ---> T), free_ts))),
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   170
          HOLogic.mk_Trueprop (Free ("P", rT --> HOLogic.boolT) $
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   171
            list_comb (r, free_ts)))))
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   172
      end;
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   173
17521
0f1c48de39f5 introduced AList module in favor of assoc etc.
haftmann
parents: 16123
diff changeset
   174
    val SOME (_, _, constrs) = AList.lookup (op =) descr index;
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40844
diff changeset
   175
    val T = List.nth (Datatype_Aux.get_rec_types descr sorts, index);
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   176
    val (rs, prems) = split_list (map (make_casedist_prem T) constrs);
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   177
    val r = Const (case_name, map fastype_of rs ---> T --> rT);
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   178
35845
e5980f0ad025 renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
wenzelm
parents: 35625
diff changeset
   179
    val y = Var (("y", 0), Logic.varifyT_global T);
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   180
    val y' = Free ("y", T);
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   181
38977
e71e2c56479c eliminated really ancient OldGoals operations in favour of Goal.prove_internal (with its minimal checks and without normalization of result);
wenzelm
parents: 38795
diff changeset
   182
    val thm = Goal.prove_internal (map cert prems)
e71e2c56479c eliminated really ancient OldGoals operations in favour of Goal.prove_internal (with its minimal checks and without normalization of result);
wenzelm
parents: 38795
diff changeset
   183
      (cert (HOLogic.mk_Trueprop (Free ("P", rT --> HOLogic.boolT) $ list_comb (r, rs @ [y']))))
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   184
      (fn prems =>
38977
e71e2c56479c eliminated really ancient OldGoals operations in favour of Goal.prove_internal (with its minimal checks and without normalization of result);
wenzelm
parents: 38795
diff changeset
   185
         EVERY [
e71e2c56479c eliminated really ancient OldGoals operations in favour of Goal.prove_internal (with its minimal checks and without normalization of result);
wenzelm
parents: 38795
diff changeset
   186
          rtac (cterm_instantiate [(cert y, cert y')] exhaust) 1,
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   187
          ALLGOALS (EVERY'
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   188
            [asm_simp_tac (HOL_basic_ss addsimps case_rewrites),
38977
e71e2c56479c eliminated really ancient OldGoals operations in favour of Goal.prove_internal (with its minimal checks and without normalization of result);
wenzelm
parents: 38795
diff changeset
   189
             resolve_tac prems, asm_simp_tac HOL_basic_ss])])
e71e2c56479c eliminated really ancient OldGoals operations in favour of Goal.prove_internal (with its minimal checks and without normalization of result);
wenzelm
parents: 38795
diff changeset
   190
      |> Drule.export_without_context;
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   191
36744
6e1f3d609a68 renamed Thm.get_name -> Thm.derivation_name and Thm.put_name -> Thm.name_derivation, to emphasize the true nature of these operations;
wenzelm
parents: 36692
diff changeset
   192
    val exh_name = Thm.derivation_name exhaust;
18358
0a733e11021a re-oriented some result tuples in PureThy
haftmann
parents: 17959
diff changeset
   193
    val (thm', thy') = thy
30435
e62d6ecab6ad explicit Binding.qualified_name -- prevents implicitly qualified bstring;
wenzelm
parents: 30364
diff changeset
   194
      |> Sign.root_path
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 38977
diff changeset
   195
      |> Global_Theory.store_thm (Binding.qualified_name (exh_name ^ "_P_correctness"), thm)
24712
64ed05609568 proper Sign operations instead of Theory aliases;
wenzelm
parents: 24711
diff changeset
   196
      ||> Sign.restore_naming thy;
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   197
13725
12404b452034 Changed format of realizers / correctness proofs.
berghofe
parents: 13708
diff changeset
   198
    val P = Var (("P", 0), rT' --> HOLogic.boolT);
37233
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36744
diff changeset
   199
    val prf = Extraction.abs_corr_shyps thy' exhaust ["P"] [y, P]
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36744
diff changeset
   200
      (fold_rev (fn (p, r) => fn prf =>
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36744
diff changeset
   201
          Proofterm.forall_intr_proof' (Logic.varify_global r)
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36744
diff changeset
   202
            (AbsP ("H", SOME (Logic.varify_global p), prf)))
33338
de76079f973a eliminated some old folds;
wenzelm
parents: 33035
diff changeset
   203
        (prems ~~ rs)
37233
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36744
diff changeset
   204
        (Proofterm.proof_combP (prf_of thm', map PBound (length prems - 1 downto 0))));
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36744
diff changeset
   205
    val prf' = Extraction.abs_corr_shyps thy' exhaust []
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36744
diff changeset
   206
      (map Var (Term.add_vars (prop_of exhaust) [])) (prf_of exhaust);
35845
e5980f0ad025 renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
wenzelm
parents: 35625
diff changeset
   207
    val r' = Logic.varify_global (Abs ("y", T,
13725
12404b452034 Changed format of realizers / correctness proofs.
berghofe
parents: 13708
diff changeset
   208
      list_abs (map dest_Free rs, list_comb (r,
12404b452034 Changed format of realizers / correctness proofs.
berghofe
parents: 13708
diff changeset
   209
        map Bound ((length rs - 1 downto 0) @ [length rs])))));
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   210
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   211
  in Extraction.add_realizers_i
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   212
    [(exh_name, (["P"], r', prf)),
37233
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36744
diff changeset
   213
     (exh_name, ([], Extraction.nullt, prf'))] thy'
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   214
  end;
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   215
31668
a616e56a5ec8 datatype packages: record datatype_config for configuration flags; less verbose signatures
haftmann
parents: 31604
diff changeset
   216
fun add_dt_realizers config names thy =
41698
90597e044e5f clarified Proofterm.proofs_enabled;
wenzelm
parents: 41423
diff changeset
   217
  if not (Proofterm.proofs_enabled ()) then thy
38977
e71e2c56479c eliminated really ancient OldGoals operations in favour of Goal.prove_internal (with its minimal checks and without normalization of result);
wenzelm
parents: 38795
diff changeset
   218
  else
e71e2c56479c eliminated really ancient OldGoals operations in favour of Goal.prove_internal (with its minimal checks and without normalization of result);
wenzelm
parents: 38795
diff changeset
   219
    let
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40844
diff changeset
   220
      val _ = Datatype_Aux.message config "Adding realizers for induction and case analysis ...";
38977
e71e2c56479c eliminated really ancient OldGoals operations in favour of Goal.prove_internal (with its minimal checks and without normalization of result);
wenzelm
parents: 38795
diff changeset
   221
      val infos = map (Datatype.the_info thy) names;
e71e2c56479c eliminated really ancient OldGoals operations in favour of Goal.prove_internal (with its minimal checks and without normalization of result);
wenzelm
parents: 38795
diff changeset
   222
      val info :: _ = infos;
e71e2c56479c eliminated really ancient OldGoals operations in favour of Goal.prove_internal (with its minimal checks and without normalization of result);
wenzelm
parents: 38795
diff changeset
   223
    in
e71e2c56479c eliminated really ancient OldGoals operations in favour of Goal.prove_internal (with its minimal checks and without normalization of result);
wenzelm
parents: 38795
diff changeset
   224
      thy
e71e2c56479c eliminated really ancient OldGoals operations in favour of Goal.prove_internal (with its minimal checks and without normalization of result);
wenzelm
parents: 38795
diff changeset
   225
      |> fold_rev (make_ind (#sorts info) info) (subsets 0 (length (#descr info) - 1))
e71e2c56479c eliminated really ancient OldGoals operations in favour of Goal.prove_internal (with its minimal checks and without normalization of result);
wenzelm
parents: 38795
diff changeset
   226
      |> fold_rev (make_casedists (#sorts info)) infos
e71e2c56479c eliminated really ancient OldGoals operations in favour of Goal.prove_internal (with its minimal checks and without normalization of result);
wenzelm
parents: 38795
diff changeset
   227
    end;
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   228
31723
f5cafe803b55 discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents: 31668
diff changeset
   229
val setup = Datatype.interpretation add_dt_realizers;
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   230
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   231
end;