src/HOL/Tools/Datatype/datatype_realizer.ML
author wenzelm
Sat, 08 May 2010 16:53:53 +0200
changeset 36744 6e1f3d609a68
parent 36692 54b64d4ad524
child 37136 e0c9d3e49e15
child 37233 b78f31ca4675
permissions -rw-r--r--
renamed Thm.get_name -> Thm.derivation_name and Thm.put_name -> Thm.name_derivation, to emphasize the true nature of these operations;
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
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
    17
open Datatype_Aux;
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    18
35364
b8c62d60195c more antiquotations;
wenzelm
parents: 33969
diff changeset
    19
fun subsets i j =
b8c62d60195c more antiquotations;
wenzelm
parents: 33969
diff changeset
    20
  if i <= j then
b8c62d60195c more antiquotations;
wenzelm
parents: 33969
diff changeset
    21
    let val is = subsets (i+1) j
b8c62d60195c more antiquotations;
wenzelm
parents: 33969
diff changeset
    22
    in map (fn ks => i::ks) is @ is end
b8c62d60195c more antiquotations;
wenzelm
parents: 33969
diff changeset
    23
  else [[]];
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    24
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    25
fun forall_intr_prf (t, prf) =
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    26
  let val (a, T) = (case t of Var ((a, _), T) => (a, T) | Free p => p)
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15256
diff changeset
    27
  in Abst (a, SOME T, Proofterm.prf_abstract_over t prf) end;
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    28
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    29
fun prf_of thm =
28814
463c9e9111ae clarified Thm.proof_body_of vs. Thm.proof_of;
wenzelm
parents: 28799
diff changeset
    30
  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
    31
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    32
fun is_unit t = snd (strip_type (fastype_of t)) = HOLogic.unitT;
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    33
13725
12404b452034 Changed format of realizers / correctness proofs.
berghofe
parents: 13708
diff changeset
    34
fun tname_of (Type (s, _)) = s
12404b452034 Changed format of realizers / correctness proofs.
berghofe
parents: 13708
diff changeset
    35
  | tname_of _ = "";
12404b452034 Changed format of realizers / correctness proofs.
berghofe
parents: 13708
diff changeset
    36
32727
9072201cd69d avoid compound fields in datatype info record
haftmann
parents: 32712
diff changeset
    37
fun make_ind sorts ({descr, rec_names, rec_rewrites, induct, ...} : info) is thy =
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    38
  let
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    39
    val recTs = get_rec_types descr sorts;
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    40
    val pnames = if length descr = 1 then ["P"]
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    41
      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
    42
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    43
    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
    44
      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
    45
        (descr ~~ pnames);
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    46
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    47
    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
    48
      if member (op =) is i then
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
    49
        Free (List.nth (pnames, i), T --> U --> HOLogic.boolT) $ r $ x
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
    50
      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
    51
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    52
    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
    53
      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
    54
31458
b1cf26f2919b avoid Library.foldl_map
haftmann
parents: 30435
diff changeset
    55
    val (prems, rec_fns) = split_list (flat (fst (fold_map
b1cf26f2919b avoid Library.foldl_map
haftmann
parents: 30435
diff changeset
    56
      (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
    57
        let
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    58
          val Ts = map (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
    59
          val tnames = Name.variant_list pnames (Datatype_Prop.make_tnames Ts);
31458
b1cf26f2919b avoid Library.foldl_map
haftmann
parents: 30435
diff changeset
    60
          val recs = filter (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
    61
          val frees = tnames ~~ Ts;
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    62
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    63
          fun mk_prems vs [] = 
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    64
                let
31458
b1cf26f2919b avoid Library.foldl_map
haftmann
parents: 30435
diff changeset
    65
                  val rT = nth (rec_result_Ts) i;
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    66
                  val vs' = filter_out is_unit vs;
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    67
                  val f = mk_Free "f" (map fastype_of vs' ---> rT) j;
18929
d81435108688 Envir.(beta_)eta_contract;
wenzelm
parents: 18358
diff changeset
    68
                  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
    69
                    (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
    70
                      else HOLogic.unit));
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    71
                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
    72
                  (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
    73
                end
13641
63d1790a43ed Reimplemented parts of datatype package dealing with datatypes involving
berghofe
parents: 13467
diff changeset
    74
            | mk_prems vs (((dt, s), T) :: ds) = 
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    75
                let
13641
63d1790a43ed Reimplemented parts of datatype package dealing with datatypes involving
berghofe
parents: 13467
diff changeset
    76
                  val k = body_index dt;
63d1790a43ed Reimplemented parts of datatype package dealing with datatypes involving
berghofe
parents: 13467
diff changeset
    77
                  val (Us, U) = strip_type T;
63d1790a43ed Reimplemented parts of datatype package dealing with datatypes involving
berghofe
parents: 13467
diff changeset
    78
                  val i = length Us;
31458
b1cf26f2919b avoid Library.foldl_map
haftmann
parents: 30435
diff changeset
    79
                  val rT = nth (rec_result_Ts) k;
13641
63d1790a43ed Reimplemented parts of datatype package dealing with datatypes involving
berghofe
parents: 13467
diff changeset
    80
                  val r = Free ("r" ^ s, Us ---> rT);
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    81
                  val (p, f) = mk_prems (vs @ [r]) ds
13641
63d1790a43ed Reimplemented parts of datatype package dealing with datatypes involving
berghofe
parents: 13467
diff changeset
    82
                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
    83
                  (list_all (map (pair "x") Us, HOLogic.mk_Trueprop
63d1790a43ed Reimplemented parts of datatype package dealing with datatypes involving
berghofe
parents: 13467
diff changeset
    84
                    (make_pred k rT U (app_bnds r i)
63d1790a43ed Reimplemented parts of datatype package dealing with datatypes involving
berghofe
parents: 13467
diff changeset
    85
                      (app_bnds (Free (s, T)) i))), p)), f)
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    86
                end
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    87
31458
b1cf26f2919b avoid Library.foldl_map
haftmann
parents: 30435
diff changeset
    88
        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
    89
          constrs) (descr ~~ recTs) 1)));
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    90
 
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    91
    fun mk_proj j [] t = t
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    92
      | mk_proj j (i :: is) t = if null is then t else
23577
c5b93c69afd3 avoid polymorphic equality;
wenzelm
parents: 22691
diff changeset
    93
          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
    94
          else mk_proj j is (HOLogic.mk_snd t);
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    95
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
    96
    val tnames = Datatype_Prop.make_tnames recTs;
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    97
    val fTs = map fastype_of rec_fns;
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    98
    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
    99
      (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
   100
        (descr ~~ recTs ~~ rec_result_Ts ~~ rec_names);
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   101
    val r = if null is then Extraction.nullt else
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32727
diff changeset
   102
      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
   103
        if member (op =) is i then SOME
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   104
          (list_comb (Const (s, fTs ---> T --> U), rec_fns) $ Free (tname, T))
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15256
diff changeset
   105
        else NONE) (descr ~~ recTs ~~ rec_result_Ts ~~ rec_names ~~ tnames));
35364
b8c62d60195c more antiquotations;
wenzelm
parents: 33969
diff changeset
   106
    val concl = HOLogic.mk_Trueprop (foldr1 (HOLogic.mk_binop @{const_name "op &"})
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   107
      (map (fn ((((i, _), T), U), tname) =>
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   108
        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
   109
          (descr ~~ recTs ~~ rec_result_Ts ~~ tnames)));
22578
b0eb5652f210 removed obsolete sign_of/sign_of_thm;
wenzelm
parents: 21646
diff changeset
   110
    val cert = cterm_of thy;
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   111
    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
   112
      (HOLogic.dest_Trueprop (concl_of induct))) ~~ ps);
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   113
17959
8db36a108213 OldGoals;
wenzelm
parents: 17521
diff changeset
   114
    val thm = OldGoals.simple_prove_goal_cterm (cert (Logic.list_implies (prems, concl)))
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   115
      (fn prems =>
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   116
         [rewrite_goals_tac (map mk_meta_eq [fst_conv, snd_conv]),
32712
ec5976f4d3d8 registering split rules and projected induction rules; ML identifiers more close to Isar theorem names
haftmann
parents: 32013
diff changeset
   117
          rtac (cterm_instantiate inst induct) 1,
35625
9c818cab0dd0 modernized structure Object_Logic;
wenzelm
parents: 35364
diff changeset
   118
          ALLGOALS Object_Logic.atomize_prems_tac,
26359
6d437bde2f1d more antiquotations
haftmann
parents: 25223
diff changeset
   119
          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
   120
          REPEAT ((resolve_tac prems THEN_ALL_NEW (fn i =>
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   121
            REPEAT (etac allE i) THEN atac i)) 1)]);
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   122
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
   123
    val ind_name = Thm.derivation_name induct;
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   124
    val vs = map (fn i => List.nth (pnames, i)) is;
18358
0a733e11021a re-oriented some result tuples in PureThy
haftmann
parents: 17959
diff changeset
   125
    val (thm', thy') = thy
30435
e62d6ecab6ad explicit Binding.qualified_name -- prevents implicitly qualified bstring;
wenzelm
parents: 30364
diff changeset
   126
      |> Sign.root_path
e62d6ecab6ad explicit Binding.qualified_name -- prevents implicitly qualified bstring;
wenzelm
parents: 30364
diff changeset
   127
      |> PureThy.store_thm
e62d6ecab6ad explicit Binding.qualified_name -- prevents implicitly qualified bstring;
wenzelm
parents: 30364
diff changeset
   128
        (Binding.qualified_name (space_implode "_" (ind_name :: vs @ ["correctness"])), thm)
24712
64ed05609568 proper Sign operations instead of Theory aliases;
wenzelm
parents: 24711
diff changeset
   129
      ||> Sign.restore_naming thy;
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   130
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
   131
    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
   132
    val rvs = rev (Thm.fold_terms Term.add_vars thm' []);
35364
b8c62d60195c more antiquotations;
wenzelm
parents: 33969
diff changeset
   133
    val ivs1 = map Var (filter_out (fn (_, T) =>  (* FIXME set (!??) *)
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
   134
      member (op =) [@{type_abbrev set}, @{type_name bool}] (tname_of (body_type T))) ivs);
33035
15eab423e573 standardized basic operations on type option;
wenzelm
parents: 32952
diff changeset
   135
    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
   136
33338
de76079f973a eliminated some old folds;
wenzelm
parents: 33035
diff changeset
   137
    val prf =
de76079f973a eliminated some old folds;
wenzelm
parents: 33035
diff changeset
   138
      List.foldr forall_intr_prf
de76079f973a eliminated some old folds;
wenzelm
parents: 33035
diff changeset
   139
        (fold_rev (fn (f, p) => fn prf =>
de76079f973a eliminated some old folds;
wenzelm
parents: 33035
diff changeset
   140
          (case head_of (strip_abs_body f) of
de76079f973a eliminated some old folds;
wenzelm
parents: 33035
diff changeset
   141
             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
   142
               let val T' = Logic.varifyT_global T
33338
de76079f973a eliminated some old folds;
wenzelm
parents: 33035
diff changeset
   143
               in Abst (s, SOME T', Proofterm.prf_abstract_over
de76079f973a eliminated some old folds;
wenzelm
parents: 33035
diff changeset
   144
                 (Var ((s, 0), T')) (AbsP ("H", SOME p, prf)))
de76079f973a eliminated some old folds;
wenzelm
parents: 33035
diff changeset
   145
               end
de76079f973a eliminated some old folds;
wenzelm
parents: 33035
diff changeset
   146
          | _ => AbsP ("H", SOME p, prf)))
de76079f973a eliminated some old folds;
wenzelm
parents: 33035
diff changeset
   147
            (rec_fns ~~ prems_of thm)
de76079f973a eliminated some old folds;
wenzelm
parents: 33035
diff changeset
   148
            (Proofterm.proof_combP (prf_of thm', map PBound (length prems - 1 downto 0)))) ivs2;
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   149
33338
de76079f973a eliminated some old folds;
wenzelm
parents: 33035
diff changeset
   150
    val r' =
de76079f973a eliminated some old folds;
wenzelm
parents: 33035
diff changeset
   151
      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
   152
      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
   153
        (map Logic.unvarify_global ivs1 @ filter_out is_unit
33338
de76079f973a eliminated some old folds;
wenzelm
parents: 33035
diff changeset
   154
            (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
   155
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   156
  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
   157
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   158
32712
ec5976f4d3d8 registering split rules and projected induction rules; ML identifiers more close to Isar theorem names
haftmann
parents: 32013
diff changeset
   159
fun make_casedists sorts ({index, descr, case_name, case_rewrites, exhaust, ...} : info) thy =
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   160
  let
19806
f860b7a98445 renamed Type.(un)varifyT to Logic.(un)varifyT;
wenzelm
parents: 18929
diff changeset
   161
    val cert = cterm_of thy;
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   162
    val rT = TFree ("'P", HOLogic.typeS);
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   163
    val rT' = TVar (("'P", 0), HOLogic.typeS);
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   164
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   165
    fun make_casedist_prem T (cname, cargs) =
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   166
      let
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   167
        val Ts = map (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
   168
        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
   169
        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
   170
        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
   171
      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
   172
        (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
   173
          HOLogic.mk_Trueprop (Free ("P", rT --> HOLogic.boolT) $
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   174
            list_comb (r, free_ts)))))
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   175
      end;
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   176
17521
0f1c48de39f5 introduced AList module in favor of assoc etc.
haftmann
parents: 16123
diff changeset
   177
    val SOME (_, _, constrs) = AList.lookup (op =) descr index;
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   178
    val T = List.nth (get_rec_types descr sorts, index);
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   179
    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
   180
    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
   181
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
   182
    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
   183
    val y' = Free ("y", T);
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   184
17959
8db36a108213 OldGoals;
wenzelm
parents: 17521
diff changeset
   185
    val thm = OldGoals.prove_goalw_cterm [] (cert (Logic.list_implies (prems,
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   186
      HOLogic.mk_Trueprop (Free ("P", rT --> HOLogic.boolT) $
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   187
        list_comb (r, rs @ [y'])))))
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   188
      (fn prems =>
32712
ec5976f4d3d8 registering split rules and projected induction rules; ML identifiers more close to Isar theorem names
haftmann
parents: 32013
diff changeset
   189
         [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
   190
          ALLGOALS (EVERY'
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   191
            [asm_simp_tac (HOL_basic_ss addsimps case_rewrites),
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   192
             resolve_tac prems, asm_simp_tac HOL_basic_ss])]);
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   193
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
   194
    val exh_name = Thm.derivation_name exhaust;
18358
0a733e11021a re-oriented some result tuples in PureThy
haftmann
parents: 17959
diff changeset
   195
    val (thm', thy') = thy
30435
e62d6ecab6ad explicit Binding.qualified_name -- prevents implicitly qualified bstring;
wenzelm
parents: 30364
diff changeset
   196
      |> Sign.root_path
e62d6ecab6ad explicit Binding.qualified_name -- prevents implicitly qualified bstring;
wenzelm
parents: 30364
diff changeset
   197
      |> PureThy.store_thm (Binding.qualified_name (exh_name ^ "_P_correctness"), thm)
24712
64ed05609568 proper Sign operations instead of Theory aliases;
wenzelm
parents: 24711
diff changeset
   198
      ||> Sign.restore_naming thy;
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   199
13725
12404b452034 Changed format of realizers / correctness proofs.
berghofe
parents: 13708
diff changeset
   200
    val P = Var (("P", 0), rT' --> HOLogic.boolT);
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   201
    val prf = forall_intr_prf (y, forall_intr_prf (P,
33338
de76079f973a eliminated some old folds;
wenzelm
parents: 33035
diff changeset
   202
      fold_rev (fn (p, r) => fn prf =>
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
   203
          forall_intr_prf (Logic.varify_global r, AbsP ("H", SOME (Logic.varify_global p), prf)))
33338
de76079f973a eliminated some old folds;
wenzelm
parents: 33035
diff changeset
   204
        (prems ~~ rs)
de76079f973a eliminated some old folds;
wenzelm
parents: 33035
diff changeset
   205
        (Proofterm.proof_combP (prf_of thm', map PBound (length prems - 1 downto 0)))));
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
   206
    val r' = Logic.varify_global (Abs ("y", T,
13725
12404b452034 Changed format of realizers / correctness proofs.
berghofe
parents: 13708
diff changeset
   207
      list_abs (map dest_Free rs, list_comb (r,
12404b452034 Changed format of realizers / correctness proofs.
berghofe
parents: 13708
diff changeset
   208
        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
   209
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   210
  in Extraction.add_realizers_i
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   211
    [(exh_name, (["P"], r', prf)),
32712
ec5976f4d3d8 registering split rules and projected induction rules; ML identifiers more close to Isar theorem names
haftmann
parents: 32013
diff changeset
   212
     (exh_name, ([], Extraction.nullt, prf_of exhaust))] thy'
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   213
  end;
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   214
31668
a616e56a5ec8 datatype packages: record datatype_config for configuration flags; less verbose signatures
haftmann
parents: 31604
diff changeset
   215
fun add_dt_realizers config names thy =
25223
7463251e7273 qualified Proofterm.proofs;
wenzelm
parents: 24712
diff changeset
   216
  if ! Proofterm.proofs < 2 then thy
24699
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 23590
diff changeset
   217
  else let
31668
a616e56a5ec8 datatype packages: record datatype_config for configuration flags; less verbose signatures
haftmann
parents: 31604
diff changeset
   218
    val _ = message config "Adding realizers for induction and case analysis ..."
31784
bd3486c57ba3 tuned interfaces of datatype module
haftmann
parents: 31775
diff changeset
   219
    val infos = map (Datatype.the_info thy) names;
24699
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 23590
diff changeset
   220
    val info :: _ = infos;
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 23590
diff changeset
   221
  in
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 23590
diff changeset
   222
    thy
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 23590
diff changeset
   223
    |> fold_rev (make_ind (#sorts info) info) (subsets 0 (length (#descr info) - 1))
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 23590
diff changeset
   224
    |> fold_rev (make_casedists (#sorts info)) infos
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 23590
diff changeset
   225
  end;
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   226
31723
f5cafe803b55 discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents: 31668
diff changeset
   227
val setup = Datatype.interpretation add_dt_realizers;
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   228
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   229
end;