src/HOL/Tools/datatype_realizer.ML
author wenzelm
Sat, 18 Jul 2015 20:47:08 +0200
changeset 60752 b48830b670a1
parent 59642 929984c529d3
child 60781 2da59cdf531c
permissions -rw-r--r--
prefer tactics with explicit context;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
58275
280ede57a6a9 renamed ML file and module
blanchet
parents: 58274
diff changeset
     1
(*  Title:      HOL/Tools/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
58275
280ede57a6a9 renamed ML file and module
blanchet
parents: 58274
diff changeset
     8
signature DATATYPE_REALIZER =
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
     9
sig
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
    10
  val add_dt_realizers: Old_Datatype_Aux.config -> string list -> theory -> theory
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    11
end;
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    12
58275
280ede57a6a9 renamed ML file and module
blanchet
parents: 58274
diff changeset
    13
structure Datatype_Realizer : DATATYPE_REALIZER =
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    14
struct
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    15
35364
b8c62d60195c more antiquotations;
wenzelm
parents: 33969
diff changeset
    16
fun subsets i j =
b8c62d60195c more antiquotations;
wenzelm
parents: 33969
diff changeset
    17
  if i <= j then
b8c62d60195c more antiquotations;
wenzelm
parents: 33969
diff changeset
    18
    let val is = subsets (i+1) j
b8c62d60195c more antiquotations;
wenzelm
parents: 33969
diff changeset
    19
    in map (fn ks => i::ks) is @ is end
b8c62d60195c more antiquotations;
wenzelm
parents: 33969
diff changeset
    20
  else [[]];
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    21
40844
5895c525739d more direct use of binder_types/body_type;
wenzelm
parents: 39557
diff changeset
    22
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
    23
13725
12404b452034 Changed format of realizers / correctness proofs.
berghofe
parents: 13708
diff changeset
    24
fun tname_of (Type (s, _)) = s
12404b452034 Changed format of realizers / correctness proofs.
berghofe
parents: 13708
diff changeset
    25
  | tname_of _ = "";
12404b452034 Changed format of realizers / correctness proofs.
berghofe
parents: 13708
diff changeset
    26
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
    27
fun make_ind ({descr, rec_names, rec_rewrites, induct, ...} : Old_Datatype_Aux.info) is thy =
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    28
  let
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 52487
diff changeset
    29
    val ctxt = Proof_Context.init_global thy;
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 52487
diff changeset
    30
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
    31
    val recTs = Old_Datatype_Aux.get_rec_types descr;
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
    32
    val pnames =
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
    33
      if length descr = 1 then ["P"]
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    34
      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
    35
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    36
    val rec_result_Ts = map (fn ((i, _), P) =>
56254
a2dd9200854d more antiquotations;
wenzelm
parents: 54895
diff changeset
    37
        if member (op =) is i then TFree ("'" ^ P, @{sort type}) else HOLogic.unitT)
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
    38
      (descr ~~ pnames);
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    39
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    40
    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
    41
      if member (op =) is i then
42364
8c674b3b8e44 eliminated old List.nth;
wenzelm
parents: 41698
diff changeset
    42
        Free (nth pnames i, T --> U --> HOLogic.boolT) $ r $ x
8c674b3b8e44 eliminated old List.nth;
wenzelm
parents: 41698
diff changeset
    43
      else Free (nth pnames i, U --> HOLogic.boolT) $ x;
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    44
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    45
    fun mk_all i s T t =
46215
0da9433f959e discontinued old-style Term.list_all_free in favour of plain Logic.all;
wenzelm
parents: 45896
diff changeset
    46
      if member (op =) is i then Logic.all (Free (s, T)) t else t;
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    47
31458
b1cf26f2919b avoid Library.foldl_map
haftmann
parents: 30435
diff changeset
    48
    val (prems, rec_fns) = split_list (flat (fst (fold_map
b1cf26f2919b avoid Library.foldl_map
haftmann
parents: 30435
diff changeset
    49
      (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
    50
        let
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
    51
          val Ts = map (Old_Datatype_Aux.typ_of_dtyp descr) cargs;
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
    52
          val tnames = Name.variant_list pnames (Old_Datatype_Prop.make_tnames Ts);
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
    53
          val recs = filter (Old_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
    54
          val frees = tnames ~~ Ts;
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    55
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
    56
          fun mk_prems vs [] =
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    57
                let
31458
b1cf26f2919b avoid Library.foldl_map
haftmann
parents: 30435
diff changeset
    58
                  val rT = nth (rec_result_Ts) i;
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    59
                  val vs' = filter_out is_unit vs;
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
    60
                  val f = Old_Datatype_Aux.mk_Free "f" (map fastype_of vs' ---> rT) j;
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
    61
                  val f' =
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
    62
                    Envir.eta_contract (fold_rev (absfree o dest_Free) vs
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
    63
                      (if member (op =) is i then list_comb (f, vs') else HOLogic.unit));
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
    64
                in
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
    65
                  (HOLogic.mk_Trueprop (make_pred i rT T (list_comb (f, vs'))
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
    66
                    (list_comb (Const (cname, Ts ---> T), map Free frees))), f')
13467
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
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
    70
                  val k = Old_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);
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
    75
                  val (p, f) = mk_prems (vs @ [r]) ds;
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
    76
                in
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
    77
                  (mk_all k ("r" ^ s) (Us ---> rT) (Logic.mk_implies
46218
ecf6375e2abb renamed Term.list_all to Logic.list_all, in accordance to HOLogic.list_all;
wenzelm
parents: 46215
diff changeset
    78
                    (Logic.list_all (map (pair "x") Us, HOLogic.mk_Trueprop
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
    79
                      (make_pred k rT U (Old_Datatype_Aux.app_bnds r i)
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
    80
                        (Old_Datatype_Aux.app_bnds (Free (s, T)) i))), p)), f)
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
    81
                end;
46215
0da9433f959e discontinued old-style Term.list_all_free in favour of plain Logic.all;
wenzelm
parents: 45896
diff changeset
    82
        in (apfst (fold_rev (Logic.all o Free) frees) (mk_prems (map Free frees) recs), j + 1) end)
31458
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
58111
82db9ad610b9 tuned structure inclusion
blanchet
parents: 56254
diff changeset
    85
    fun mk_proj _ [] t = t
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
    86
      | mk_proj j (i :: is) t =
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
    87
          if null is then t
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
    88
          else 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
    89
          else mk_proj j is (HOLogic.mk_snd t);
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    90
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
    91
    val tnames = Old_Datatype_Prop.make_tnames recTs;
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    92
    val fTs = map fastype_of rec_fns;
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
    93
    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
    94
      (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
    95
        (descr ~~ recTs ~~ rec_result_Ts ~~ rec_names);
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
    96
    val r =
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
    97
      if null is then Extraction.nullt
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
    98
      else
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
    99
        foldr1 HOLogic.mk_prod (map_filter (fn (((((i, _), T), U), s), tname) =>
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   100
          if member (op =) is i then SOME
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   101
            (list_comb (Const (s, fTs ---> T --> U), rec_fns) $ Free (tname, T))
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   102
          else NONE) (descr ~~ recTs ~~ rec_result_Ts ~~ rec_names ~~ tnames));
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   103
    val concl =
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   104
      HOLogic.mk_Trueprop (foldr1 (HOLogic.mk_binop @{const_name HOL.conj})
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   105
        (map (fn ((((i, _), T), U), tname) =>
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   106
          make_pred i U T (mk_proj i is r) (Free (tname, T)))
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   107
            (descr ~~ recTs ~~ rec_result_Ts ~~ tnames)));
59642
929984c529d3 clarified context;
wenzelm
parents: 59621
diff changeset
   108
    val inst = map (apply2 (Thm.cterm_of ctxt)) (map head_of (HOLogic.dest_conj
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   109
      (HOLogic.dest_Trueprop (Thm.concl_of induct))) ~~ ps);
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   110
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   111
    val thm =
59642
929984c529d3 clarified context;
wenzelm
parents: 59621
diff changeset
   112
      Goal.prove_internal ctxt (map (Thm.cterm_of ctxt) prems) (Thm.cterm_of ctxt concl)
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   113
        (fn prems =>
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   114
           EVERY [
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 52487
diff changeset
   115
            rewrite_goals_tac ctxt (map mk_meta_eq [@{thm fst_conv}, @{thm snd_conv}]),
60752
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 59642
diff changeset
   116
            resolve_tac ctxt [cterm_instantiate inst induct] 1,
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 52487
diff changeset
   117
            ALLGOALS (Object_Logic.atomize_prems_tac ctxt),
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 52487
diff changeset
   118
            rewrite_goals_tac ctxt (@{thm o_def} :: map mk_meta_eq rec_rewrites),
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
   119
            REPEAT ((resolve_tac ctxt prems THEN_ALL_NEW (fn i =>
60752
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 59642
diff changeset
   120
              REPEAT (eresolve_tac ctxt [allE] i) THEN assume_tac ctxt i)) 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
   121
      |> Drule.export_without_context;
13467
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;
42364
8c674b3b8e44 eliminated old List.nth;
wenzelm
parents: 41698
diff changeset
   124
    val vs = map (nth pnames) 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
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
   127
      |> Global_Theory.store_thm
30435
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
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
   131
    val ivs = rev (Term.add_vars (Logic.varify_global (Old_Datatype_Prop.make_ind [descr])) []);
22691
290454649b8c Thm.fold_terms;
wenzelm
parents: 22596
diff changeset
   132
    val rvs = rev (Thm.fold_terms Term.add_vars thm' []);
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   133
    val ivs1 = map Var (filter_out (fn (_, T) => @{type_name bool} = tname_of (body_type T)) ivs);
33035
15eab423e573 standardized basic operations on type option;
wenzelm
parents: 32952
diff changeset
   134
    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
   135
33338
de76079f973a eliminated some old folds;
wenzelm
parents: 33035
diff changeset
   136
    val prf =
37233
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36744
diff changeset
   137
      Extraction.abs_corr_shyps thy' induct vs ivs2
33338
de76079f973a eliminated some old folds;
wenzelm
parents: 33035
diff changeset
   138
        (fold_rev (fn (f, p) => fn prf =>
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   139
            (case head_of (strip_abs_body f) of
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   140
              Free (s, T) =>
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   141
                let val T' = Logic.varifyT_global T in
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   142
                  Abst (s, SOME T', Proofterm.prf_abstract_over
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   143
                    (Var ((s, 0), T')) (AbsP ("H", SOME p, prf)))
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   144
                end
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   145
            | _ => AbsP ("H", SOME p, prf)))
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   146
          (rec_fns ~~ Thm.prems_of thm)
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   147
          (Proofterm.proof_combP
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   148
            (Reconstruct.proof_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
   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
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   152
      else
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   153
        Logic.varify_global (fold_rev lambda
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   154
          (map Logic.unvarify_global ivs1 @ filter_out is_unit
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   155
              (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
   156
58277
0dcd3a623a6e made realizer more robust in the face of nesting through functions
blanchet
parents: 58275
diff changeset
   157
  in Extraction.add_realizers_i [(ind_name, (vs, r', prf))] thy' end;
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   158
58182
82478e6c60cb tweaked setup for datatype realizer
blanchet
parents: 58112
diff changeset
   159
fun make_casedists ({index, descr, case_name, case_rewrites, exhaust, ...} : Old_Datatype_Aux.info) thy =
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   160
  let
54883
dd04a8b654fc proper context for norm_hhf and derived operations;
wenzelm
parents: 54742
diff changeset
   161
    val ctxt = Proof_Context.init_global thy;
56254
a2dd9200854d more antiquotations;
wenzelm
parents: 54895
diff changeset
   162
    val rT = TFree ("'P", @{sort type});
a2dd9200854d more antiquotations;
wenzelm
parents: 54895
diff changeset
   163
    val rT' = TVar (("'P", 0), @{sort type});
13467
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
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
   167
        val Ts = map (Old_Datatype_Aux.typ_of_dtyp descr) cargs;
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
   168
        val frees = Name.variant_list ["P", "y"] (Old_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)
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   171
      in
46215
0da9433f959e discontinued old-style Term.list_all_free in favour of plain Logic.all;
wenzelm
parents: 45896
diff changeset
   172
        (r, fold_rev Logic.all free_ts
0da9433f959e discontinued old-style Term.list_all_free in favour of plain Logic.all;
wenzelm
parents: 45896
diff changeset
   173
          (Logic.mk_implies (HOLogic.mk_Trueprop
0da9433f959e discontinued old-style Term.list_all_free in favour of plain Logic.all;
wenzelm
parents: 45896
diff changeset
   174
            (HOLogic.mk_eq (Free ("y", T), list_comb (Const (cname, Ts ---> T), free_ts))),
0da9433f959e discontinued old-style Term.list_all_free in favour of plain Logic.all;
wenzelm
parents: 45896
diff changeset
   175
              HOLogic.mk_Trueprop (Free ("P", rT --> HOLogic.boolT) $
0da9433f959e discontinued old-style Term.list_all_free in favour of plain Logic.all;
wenzelm
parents: 45896
diff changeset
   176
                list_comb (r, free_ts)))))
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   177
      end;
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   178
17521
0f1c48de39f5 introduced AList module in favor of assoc etc.
haftmann
parents: 16123
diff changeset
   179
    val SOME (_, _, constrs) = AList.lookup (op =) descr index;
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
   180
    val T = nth (Old_Datatype_Aux.get_rec_types descr) index;
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   181
    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
   182
    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
   183
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
   184
    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
   185
    val y' = Free ("y", T);
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   186
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   187
    val thm =
59642
929984c529d3 clarified context;
wenzelm
parents: 59621
diff changeset
   188
      Goal.prove_internal ctxt (map (Thm.cterm_of ctxt) prems)
929984c529d3 clarified context;
wenzelm
parents: 59621
diff changeset
   189
        (Thm.cterm_of ctxt
59617
b60e65ad13df tuned -- more explicit use of context;
wenzelm
parents: 59582
diff changeset
   190
          (HOLogic.mk_Trueprop (Free ("P", rT --> HOLogic.boolT) $ list_comb (r, rs @ [y']))))
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   191
        (fn prems =>
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   192
           EVERY [
60752
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 59642
diff changeset
   193
            resolve_tac ctxt [cterm_instantiate [apply2 (Thm.cterm_of ctxt) (y, y')] exhaust] 1,
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   194
            ALLGOALS (EVERY'
54895
515630483010 clarified simplifier context;
wenzelm
parents: 54883
diff changeset
   195
              [asm_simp_tac (put_simpset HOL_basic_ss ctxt addsimps case_rewrites),
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
   196
               resolve_tac ctxt prems, asm_simp_tac (put_simpset HOL_basic_ss ctxt)])])
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
   197
      |> Drule.export_without_context;
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   198
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
   199
    val exh_name = Thm.derivation_name exhaust;
18358
0a733e11021a re-oriented some result tuples in PureThy
haftmann
parents: 17959
diff changeset
   200
    val (thm', thy') = thy
30435
e62d6ecab6ad explicit Binding.qualified_name -- prevents implicitly qualified bstring;
wenzelm
parents: 30364
diff changeset
   201
      |> 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
   202
      |> 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
   203
      ||> Sign.restore_naming thy;
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   204
13725
12404b452034 Changed format of realizers / correctness proofs.
berghofe
parents: 13708
diff changeset
   205
    val P = Var (("P", 0), rT' --> HOLogic.boolT);
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   206
    val prf =
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   207
      Extraction.abs_corr_shyps thy' exhaust ["P"] [y, P]
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   208
        (fold_rev (fn (p, r) => fn prf =>
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   209
            Proofterm.forall_intr_proof' (Logic.varify_global r)
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   210
              (AbsP ("H", SOME (Logic.varify_global p), prf)))
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   211
          (prems ~~ rs)
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   212
          (Proofterm.proof_combP
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   213
            (Reconstruct.proof_of thm', map PBound (length prems - 1 downto 0))));
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   214
    val prf' =
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   215
      Extraction.abs_corr_shyps thy' exhaust []
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   216
        (map Var (Term.add_vars (Thm.prop_of exhaust) [])) (Reconstruct.proof_of exhaust);
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   217
    val r' =
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   218
      Logic.varify_global (Abs ("y", T,
46219
426ed18eba43 discontinued old-style Term.list_abs in favour of plain Term.abs;
wenzelm
parents: 46218
diff changeset
   219
        (fold_rev (Term.abs o dest_Free) rs
426ed18eba43 discontinued old-style Term.list_abs in favour of plain Term.abs;
wenzelm
parents: 46218
diff changeset
   220
          (list_comb (r, map Bound ((length rs - 1 downto 0) @ [length rs]))))));
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   221
  in
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   222
    Extraction.add_realizers_i
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   223
      [(exh_name, (["P"], r', prf)),
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   224
       (exh_name, ([], Extraction.nullt, prf'))] thy'
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   225
  end;
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   226
31668
a616e56a5ec8 datatype packages: record datatype_config for configuration flags; less verbose signatures
haftmann
parents: 31604
diff changeset
   227
fun add_dt_realizers config names thy =
52487
48bc24467008 backout dedd7952a62c: static "proofs" value within theory prevents later inferencing with different configuration;
wenzelm
parents: 52470
diff changeset
   228
  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
   229
  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
   230
    let
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
   231
      val _ = Old_Datatype_Aux.message config "Adding realizers for induction and case analysis ...";
58354
04ac60da613e support (finite values of) codatatypes in Quickcheck
blanchet
parents: 58277
diff changeset
   232
      val infos = map (BNF_LFP_Compat.the_info thy []) names;
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
   233
      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
   234
    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
   235
      thy
58277
0dcd3a623a6e made realizer more robust in the face of nesting through functions
blanchet
parents: 58275
diff changeset
   236
      |> fold_rev (perhaps o try o make_ind info) (subsets 0 (length (#descr info) - 1))
0dcd3a623a6e made realizer more robust in the face of nesting through functions
blanchet
parents: 58275
diff changeset
   237
      |> fold_rev (perhaps o try o make_casedists) infos
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
   238
    end;
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   239
58659
6c9821c32dd5 Local_Interpretation is superseded by Plugin with formal Plugin_Name management, avoiding undeclared strings;
wenzelm
parents: 58354
diff changeset
   240
val _ = Theory.setup (BNF_LFP_Compat.interpretation @{plugin extraction} [] add_dt_realizers);
13467
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   241
d66b526192bf Module for defining realizers for induction and case analysis theorems
berghofe
parents:
diff changeset
   242
end;