src/Pure/Proof/extraction.ML
author haftmann
Wed, 13 Jul 2005 11:30:37 +0200
changeset 16790 be2780f435e1
parent 16787 b6b6e2faaa41
child 16800 90eff1b52428
permissions -rw-r--r--
(fix for an accidental commit)
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
     1
(*  Title:      Pure/Proof/extraction.ML
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
     2
    ID:         $Id$
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
     3
    Author:     Stefan Berghofer, TU Muenchen
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
     4
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
     5
Extraction of programs from proofs.
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
     6
*)
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
     7
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
     8
signature EXTRACTION =
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
     9
sig
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16363
diff changeset
    10
  val set_preprocessor : (theory -> Proofterm.proof -> Proofterm.proof) -> theory -> theory
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
    11
  val add_realizes_eqns_i : ((term * term) list * (term * term)) list -> theory -> theory
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
    12
  val add_realizes_eqns : string list -> theory -> theory
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
    13
  val add_typeof_eqns_i : ((term * term) list * (term * term)) list -> theory -> theory
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
    14
  val add_typeof_eqns : string list -> theory -> theory
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
    15
  val add_realizers_i : (string * (string list * term * Proofterm.proof)) list
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
    16
    -> theory -> theory
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
    17
  val add_realizers : (thm * (string list * string * string)) list
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
    18
    -> theory -> theory
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
    19
  val add_expand_thms : thm list -> theory -> theory
13732
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
    20
  val add_types : (xstring * ((term -> term option) list *
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
    21
    (term -> typ -> term -> typ -> term) option)) list -> theory -> theory
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
    22
  val extract : (thm * string list) list -> theory -> theory
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
    23
  val nullT : typ
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
    24
  val nullt : term
13714
bdd483321f4b - exported functions etype_of and mk_typ
berghofe
parents: 13609
diff changeset
    25
  val mk_typ : typ -> term
bdd483321f4b - exported functions etype_of and mk_typ
berghofe
parents: 13609
diff changeset
    26
  val etype_of : theory -> string list -> typ list -> term -> typ
bdd483321f4b - exported functions etype_of and mk_typ
berghofe
parents: 13609
diff changeset
    27
  val realizes_of: theory -> string list -> term -> term -> term
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
    28
end;
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
    29
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
    30
structure Extraction : EXTRACTION =
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
    31
struct
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
    32
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
    33
open Proofterm;
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
    34
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
    35
(**** tools ****)
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
    36
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
    37
fun add_syntax thy =
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
    38
  thy
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
    39
  |> Theory.copy
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
    40
  |> Theory.root_path
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
    41
  |> Theory.add_types [("Type", 0, NoSyn), ("Null", 0, NoSyn)]
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
    42
  |> Theory.add_consts
14854
61bdf2ae4dc5 removed obsolete sort 'logic';
wenzelm
parents: 14472
diff changeset
    43
      [("typeof", "'b::{} => Type", NoSyn),
61bdf2ae4dc5 removed obsolete sort 'logic';
wenzelm
parents: 14472
diff changeset
    44
       ("Type", "'a::{} itself => Type", NoSyn),
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
    45
       ("Null", "Null", NoSyn),
14854
61bdf2ae4dc5 removed obsolete sort 'logic';
wenzelm
parents: 14472
diff changeset
    46
       ("realizes", "'a::{} => 'b::{} => 'b", NoSyn)];
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
    47
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
    48
val nullT = Type ("Null", []);
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
    49
val nullt = Const ("Null", nullT);
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
    50
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
    51
fun mk_typ T =
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
    52
  Const ("Type", itselfT T --> Type ("Type", [])) $ Logic.mk_type T;
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
    53
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
    54
fun typeof_proc defaultS vs (Const ("typeof", _) $ u) =
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
    55
      SOME (mk_typ (case strip_comb u of
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
    56
          (Var ((a, i), _), _) =>
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
    57
            if a mem vs then TFree ("'" ^ a ^ ":" ^ string_of_int i, defaultS)
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
    58
            else nullT
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
    59
        | (Free (a, _), _) =>
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
    60
            if a mem vs then TFree ("'" ^ a, defaultS) else nullT
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
    61
        | _ => nullT))
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
    62
  | typeof_proc _ _ _ = NONE;
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
    63
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
    64
fun rlz_proc (Const ("realizes", Type (_, [Type ("Null", []), _])) $ r $ t) = SOME t
13732
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
    65
  | rlz_proc (Const ("realizes", Type (_, [T, _])) $ r $ t) =
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
    66
      (case strip_comb t of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
    67
         (Var (ixn, U), ts) => SOME (list_comb (Var (ixn, T --> U), r :: ts))
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
    68
       | (Free (s, U), ts) => SOME (list_comb (Free (s, T --> U), r :: ts))
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
    69
       | _ => NONE)
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
    70
  | rlz_proc _ = NONE;
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
    71
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
    72
val unpack_ixn = apfst implode o apsnd (fst o read_int o tl) o
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
    73
  take_prefix (not o equal ":") o explode;
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
    74
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
    75
type rules =
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
    76
  {next: int, rs: ((term * term) list * (term * term)) list,
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
    77
   net: (int * ((term * term) list * (term * term))) Net.net};
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
    78
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
    79
val empty_rules : rules = {next = 0, rs = [], net = Net.empty};
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
    80
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
    81
fun add_rule (r as (_, (lhs, _)), {next, rs, net} : rules) =
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
    82
  {next = next - 1, rs = r :: rs, net = Net.insert_term
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
    83
     ((Pattern.eta_contract lhs, (next, r)), net, K false)};
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
    84
13417
12cc77f90811 Tuned type constraint of function merge_rules to make smlnj happy.
berghofe
parents: 13402
diff changeset
    85
fun merge_rules
12cc77f90811 Tuned type constraint of function merge_rules to make smlnj happy.
berghofe
parents: 13402
diff changeset
    86
  ({next, rs = rs1, net} : rules) ({next = next2, rs = rs2, ...} : rules) =
15574
b1d1b5bfc464 Removed practically all references to Library.foldr.
skalberg
parents: 15570
diff changeset
    87
  foldr add_rule {next = next, rs = rs1, net = net} (rs2 \\ rs1);
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
    88
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16363
diff changeset
    89
fun condrew thy rules procs =
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
    90
  let
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16363
diff changeset
    91
    val tsig = Sign.tsig_of thy;
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
    92
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
    93
    fun rew tm =
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
    94
      Pattern.rewrite_term tsig [] (condrew' :: procs) tm
15399
683d83051d6a Added term cache to function condrew in order to speed up rewriting.
berghofe
parents: 14981
diff changeset
    95
    and condrew' tm =
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
    96
      let
15399
683d83051d6a Added term cache to function condrew in order to speed up rewriting.
berghofe
parents: 14981
diff changeset
    97
        val cache = ref ([] : (term * term) list);
683d83051d6a Added term cache to function condrew in order to speed up rewriting.
berghofe
parents: 14981
diff changeset
    98
        fun lookup f x = (case assoc (!cache, x) of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
    99
            NONE =>
15399
683d83051d6a Added term cache to function condrew in order to speed up rewriting.
berghofe
parents: 14981
diff changeset
   100
              let val y = f x
683d83051d6a Added term cache to function condrew in order to speed up rewriting.
berghofe
parents: 14981
diff changeset
   101
              in (cache := (x, y) :: !cache; y) end
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   102
          | SOME y => y);
15399
683d83051d6a Added term cache to function condrew in order to speed up rewriting.
berghofe
parents: 14981
diff changeset
   103
      in
683d83051d6a Added term cache to function condrew in order to speed up rewriting.
berghofe
parents: 14981
diff changeset
   104
        get_first (fn (_, (prems, (tm1, tm2))) =>
683d83051d6a Added term cache to function condrew in order to speed up rewriting.
berghofe
parents: 14981
diff changeset
   105
        let
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   106
          fun ren t = getOpt (Term.rename_abs tm1 tm t, t);
15399
683d83051d6a Added term cache to function condrew in order to speed up rewriting.
berghofe
parents: 14981
diff changeset
   107
          val inc = Logic.incr_indexes ([], maxidx_of_term tm + 1);
683d83051d6a Added term cache to function condrew in order to speed up rewriting.
berghofe
parents: 14981
diff changeset
   108
          val env as (Tenv, tenv) = Pattern.match tsig (inc tm1, tm);
15798
016f3be5a5ec Adapted to new interface of instantiation and unification / matching functions.
berghofe
parents: 15574
diff changeset
   109
          val prems' = map (pairself (Envir.subst_vars env o inc o ren)) prems;
15399
683d83051d6a Added term cache to function condrew in order to speed up rewriting.
berghofe
parents: 14981
diff changeset
   110
          val env' = Envir.Envir
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   111
            {maxidx = Library.foldl Int.max
15399
683d83051d6a Added term cache to function condrew in order to speed up rewriting.
berghofe
parents: 14981
diff changeset
   112
              (~1, map (Int.max o pairself maxidx_of_term) prems'),
15798
016f3be5a5ec Adapted to new interface of instantiation and unification / matching functions.
berghofe
parents: 15574
diff changeset
   113
             iTs = Tenv, asol = tenv};
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   114
          val env'' = Library.foldl (fn (env, p) =>
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16363
diff changeset
   115
            Pattern.unify (thy, env, [pairself (lookup rew) p])) (env', prems')
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   116
        in SOME (Envir.norm_term env'' (inc (ren tm2)))
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   117
        end handle Pattern.MATCH => NONE | Pattern.Unif => NONE)
16486
1a12cdb6ee6b get_thm(s): Name;
wenzelm
parents: 16458
diff changeset
   118
          (sort (int_ord o pairself fst)
15399
683d83051d6a Added term cache to function condrew in order to speed up rewriting.
berghofe
parents: 14981
diff changeset
   119
            (Net.match_term rules (Pattern.eta_contract tm)))
683d83051d6a Added term cache to function condrew in order to speed up rewriting.
berghofe
parents: 14981
diff changeset
   120
      end;
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   121
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   122
  in rew end;
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   123
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   124
val chtype = change_type o SOME;
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   125
16195
wenzelm
parents: 16149
diff changeset
   126
fun extr_name s vs = NameSpace.append "extr" (space_implode "_" (s :: vs));
wenzelm
parents: 16149
diff changeset
   127
fun corr_name s vs = extr_name s vs ^ "_correctness";
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   128
16195
wenzelm
parents: 16149
diff changeset
   129
fun msg d s = priority (Symbol.spaces d ^ s);
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   130
16790
be2780f435e1 (fix for an accidental commit)
haftmann
parents: 16787
diff changeset
   131
fun vars_of t = rev (foldl_aterms
be2780f435e1 (fix for an accidental commit)
haftmann
parents: 16787
diff changeset
   132
  (fn (vs, v as Var _) => v ins vs | (vs, _) => vs) ([], t));
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   133
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   134
fun vfs_of t = vars_of t @ sort (make_ord atless) (term_frees t);
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   135
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   136
fun forall_intr (t, prop) =
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   137
  let val (a, T) = (case t of Var ((a, _), T) => (a, T) | Free p => p)
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   138
  in all T $ Abs (a, T, abstract_over (t, prop)) end;
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   139
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   140
fun forall_intr_prf (t, prf) =
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   141
  let val (a, T) = (case t of Var ((a, _), T) => (a, T) | Free p => p)
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   142
  in Abst (a, SOME T, prf_abstract_over t prf) end;
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   143
15574
b1d1b5bfc464 Removed practically all references to Library.foldr.
skalberg
parents: 15570
diff changeset
   144
val mkabs = foldr (fn (v, t) => Abs ("x", fastype_of v, abstract_over (v, t)));
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   145
13732
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   146
fun strip_abs 0 t = t
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   147
  | strip_abs n (Abs (_, _, t)) = strip_abs (n-1) t
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   148
  | strip_abs _ _ = error "strip_abs: not an abstraction";
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   149
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   150
fun prf_subst_TVars tye =
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   151
  map_proof_terms (subst_TVars tye) (typ_subst_TVars tye);
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   152
15574
b1d1b5bfc464 Removed practically all references to Library.foldr.
skalberg
parents: 15570
diff changeset
   153
fun relevant_vars types prop = foldr (fn
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   154
      (Var ((a, i), T), vs) => (case strip_type T of
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   155
        (_, Type (s, _)) => if s mem types then a :: vs else vs
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   156
      | _ => vs)
15574
b1d1b5bfc464 Removed practically all references to Library.foldr.
skalberg
parents: 15570
diff changeset
   157
    | (_, vs) => vs) [] (vars_of prop);
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   158
13732
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   159
fun tname_of (Type (s, _)) = s
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   160
  | tname_of _ = "";
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   161
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   162
fun get_var_type t =
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   163
  let
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   164
    val vs = Term.add_vars ([], t);
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   165
    val fs = Term.add_frees ([], t)
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   166
  in fn 
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   167
      Var (ixn, _) => (case assoc (Term.add_vars ([], t), ixn) of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   168
          NONE => error "get_var_type: no such variable in term"
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   169
        | SOME T => Var (ixn, T))
13732
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   170
    | Free (s, _) => (case assoc (Term.add_frees ([], t), s) of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   171
          NONE => error "get_var_type: no such variable in term"
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   172
        | SOME T => Free (s, T))
13732
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   173
    | _ => error "get_var_type: not a variable"
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   174
  end;
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   175
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   176
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   177
(**** theory data ****)
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   178
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   179
(* data kind 'Pure/extraction' *)
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   180
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16363
diff changeset
   181
structure ExtractionData = TheoryDataFun
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16363
diff changeset
   182
(struct
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   183
  val name = "Pure/extraction";
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   184
  type T =
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   185
    {realizes_eqns : rules,
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   186
     typeof_eqns : rules,
13732
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   187
     types : (string * ((term -> term option) list *
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   188
       (term -> typ -> term -> typ -> term) option)) list,
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   189
     realizers : (string list * (term * proof)) list Symtab.table,
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   190
     defs : thm list,
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   191
     expand : (string * term) list,
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16363
diff changeset
   192
     prep : (theory -> proof -> proof) option}
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   193
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   194
  val empty =
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   195
    {realizes_eqns = empty_rules,
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   196
     typeof_eqns = empty_rules,
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   197
     types = [],
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   198
     realizers = Symtab.empty,
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   199
     defs = [],
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   200
     expand = [],
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   201
     prep = NONE};
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   202
  val copy = I;
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16363
diff changeset
   203
  val extend = I;
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   204
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16363
diff changeset
   205
  fun merge _
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   206
    (({realizes_eqns = realizes_eqns1, typeof_eqns = typeof_eqns1, types = types1,
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   207
       realizers = realizers1, defs = defs1, expand = expand1, prep = prep1},
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   208
      {realizes_eqns = realizes_eqns2, typeof_eqns = typeof_eqns2, types = types2,
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   209
       realizers = realizers2, defs = defs2, expand = expand2, prep = prep2}) : T * T) =
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   210
    {realizes_eqns = merge_rules realizes_eqns1 realizes_eqns2,
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   211
     typeof_eqns = merge_rules typeof_eqns1 typeof_eqns2,
13732
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   212
     types = merge_alists types1 types2,
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   213
     realizers = Symtab.merge_multi' (eq_set o pairself #1)
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   214
       (realizers1, realizers2),
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   215
     defs = gen_merge_lists eq_thm defs1 defs2,
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   216
     expand = merge_lists expand1 expand2,
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   217
     prep = (case prep1 of NONE => prep2 | _ => prep1)};
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   218
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16363
diff changeset
   219
  fun print _ _ = ();
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16363
diff changeset
   220
end);
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   221
15801
d2f5ca3c048d superceded by Pure.thy and CPure.thy;
wenzelm
parents: 15798
diff changeset
   222
val _ = Context.add_setup [ExtractionData.init];
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   223
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   224
fun read_condeq thy =
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16363
diff changeset
   225
  let val thy' = add_syntax thy
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   226
  in fn s =>
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16363
diff changeset
   227
    let val t = Logic.varify (term_of (read_cterm thy' (s, propT)))
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   228
    in (map Logic.dest_equals (Logic.strip_imp_prems t),
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   229
      Logic.dest_equals (Logic.strip_imp_concl t))
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   230
    end handle TERM _ => error ("Not a (conditional) meta equality:\n" ^ s)
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   231
  end;
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   232
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   233
(** preprocessor **)
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   234
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   235
fun set_preprocessor prep thy =
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   236
  let val {realizes_eqns, typeof_eqns, types, realizers,
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   237
    defs, expand, ...} = ExtractionData.get thy
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   238
  in
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   239
    ExtractionData.put
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   240
      {realizes_eqns = realizes_eqns, typeof_eqns = typeof_eqns, types = types,
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   241
       realizers = realizers, defs = defs, expand = expand, prep = SOME prep} thy
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   242
  end;
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   243
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   244
(** equations characterizing realizability **)
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   245
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   246
fun gen_add_realizes_eqns prep_eq eqns thy =
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   247
  let val {realizes_eqns, typeof_eqns, types, realizers,
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   248
    defs, expand, prep} = ExtractionData.get thy;
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   249
  in
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   250
    ExtractionData.put
15574
b1d1b5bfc464 Removed practically all references to Library.foldr.
skalberg
parents: 15570
diff changeset
   251
      {realizes_eqns = foldr add_rule realizes_eqns (map (prep_eq thy) eqns),
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   252
       typeof_eqns = typeof_eqns, types = types, realizers = realizers,
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   253
       defs = defs, expand = expand, prep = prep} thy
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   254
  end
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   255
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   256
val add_realizes_eqns_i = gen_add_realizes_eqns (K I);
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   257
val add_realizes_eqns = gen_add_realizes_eqns read_condeq;
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   258
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   259
(** equations characterizing type of extracted program **)
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   260
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   261
fun gen_add_typeof_eqns prep_eq eqns thy =
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   262
  let
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   263
    val {realizes_eqns, typeof_eqns, types, realizers,
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   264
      defs, expand, prep} = ExtractionData.get thy;
13732
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   265
    val eqns' = map (prep_eq thy) eqns
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   266
  in
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   267
    ExtractionData.put
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   268
      {realizes_eqns = realizes_eqns, realizers = realizers,
15574
b1d1b5bfc464 Removed practically all references to Library.foldr.
skalberg
parents: 15570
diff changeset
   269
       typeof_eqns = foldr add_rule typeof_eqns eqns',
13732
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   270
       types = types, defs = defs, expand = expand, prep = prep} thy
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   271
  end
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   272
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   273
val add_typeof_eqns_i = gen_add_typeof_eqns (K I);
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   274
val add_typeof_eqns = gen_add_typeof_eqns read_condeq;
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   275
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   276
fun thaw (T as TFree (a, S)) =
16195
wenzelm
parents: 16149
diff changeset
   277
      if exists_string (equal ":") a then TVar (unpack_ixn a, S) else T
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   278
  | thaw (Type (a, Ts)) = Type (a, map thaw Ts)
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   279
  | thaw T = T;
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   280
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   281
fun freeze (TVar ((a, i), S)) = TFree (a ^ ":" ^ string_of_int i, S)
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   282
  | freeze (Type (a, Ts)) = Type (a, map freeze Ts)
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   283
  | freeze T = T;
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   284
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   285
fun freeze_thaw f x =
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   286
  map_term_types thaw (f (map_term_types freeze x));
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   287
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16363
diff changeset
   288
fun etype_of thy vs Ts t =
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   289
  let
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16363
diff changeset
   290
    val {typeof_eqns, ...} = ExtractionData.get thy;
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   291
    fun err () = error ("Unable to determine type of extracted program for\n" ^
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16363
diff changeset
   292
      Sign.string_of_term thy t)
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16363
diff changeset
   293
  in case strip_abs_body (freeze_thaw (condrew thy (#net typeof_eqns)
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16363
diff changeset
   294
    [typeof_proc (Sign.defaultS thy) vs]) (list_abs (map (pair "x") (rev Ts),
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   295
      Const ("typeof", fastype_of1 (Ts, t) --> Type ("Type", [])) $ t))) of
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   296
      Const ("Type", _) $ u => (Logic.dest_type u handle TERM _ => err ())
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   297
    | _ => err ()
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   298
  end;
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   299
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   300
(** realizers for axioms / theorems, together with correctness proofs **)
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   301
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   302
fun gen_add_realizers prep_rlz rs thy =
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   303
  let val {realizes_eqns, typeof_eqns, types, realizers,
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   304
    defs, expand, prep} = ExtractionData.get thy
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   305
  in
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   306
    ExtractionData.put
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   307
      {realizes_eqns = realizes_eqns, typeof_eqns = typeof_eqns, types = types,
15574
b1d1b5bfc464 Removed practically all references to Library.foldr.
skalberg
parents: 15570
diff changeset
   308
       realizers = foldr Symtab.update_multi
b1d1b5bfc464 Removed practically all references to Library.foldr.
skalberg
parents: 15570
diff changeset
   309
         realizers (map (prep_rlz thy) (rev rs)),
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   310
       defs = defs, expand = expand, prep = prep} thy
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   311
  end
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   312
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   313
fun prep_realizer thy =
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   314
  let
13732
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   315
    val {realizes_eqns, typeof_eqns, defs, types, ...} =
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   316
      ExtractionData.get thy;
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   317
    val procs = List.concat (map (fst o snd) types);
13732
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   318
    val rtypes = map fst types;
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   319
    val eqns = Net.merge (#net realizes_eqns, #net typeof_eqns, K false);
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   320
    val thy' = add_syntax thy;
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   321
    val rd = ProofSyntax.read_proof thy' false
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   322
  in fn (thm, (vs, s1, s2)) =>
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   323
    let
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   324
      val name = Thm.name_of_thm thm;
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   325
      val _ = assert (name <> "") "add_realizers: unnamed theorem";
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16363
diff changeset
   326
      val prop = Pattern.rewrite_term (Sign.tsig_of thy')
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   327
        (map (Logic.dest_equals o prop_of) defs) [] (prop_of thm);
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   328
      val vars = vars_of prop;
13732
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   329
      val vars' = filter_out (fn v =>
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   330
        tname_of (body_type (fastype_of v)) mem rtypes) vars;
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16363
diff changeset
   331
      val T = etype_of thy' vs [] prop;
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   332
      val (T', thw) = Type.freeze_thaw_type
13732
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   333
        (if T = nullT then nullT else map fastype_of vars' ---> T);
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16363
diff changeset
   334
      val t = map_term_types thw (term_of (read_cterm thy' (s1, T')));
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16363
diff changeset
   335
      val r' = freeze_thaw (condrew thy' eqns
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16363
diff changeset
   336
        (procs @ [typeof_proc (Sign.defaultS thy') vs, rlz_proc]))
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   337
          (Const ("realizes", T --> propT --> propT) $
13732
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   338
            (if T = nullT then t else list_comb (t, vars')) $ prop);
15574
b1d1b5bfc464 Removed practically all references to Library.foldr.
skalberg
parents: 15570
diff changeset
   339
      val r = foldr forall_intr r' (map (get_var_type r') vars);
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16363
diff changeset
   340
      val prf = Reconstruct.reconstruct_proof thy' r (rd s2);
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   341
    in (name, (vs, (t, prf))) end
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   342
  end;
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   343
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   344
val add_realizers_i = gen_add_realizers
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   345
  (fn _ => fn (name, (vs, t, prf)) => (name, (vs, (t, prf))));
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   346
val add_realizers = gen_add_realizers prep_realizer;
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   347
13714
bdd483321f4b - exported functions etype_of and mk_typ
berghofe
parents: 13609
diff changeset
   348
fun realizes_of thy vs t prop =
bdd483321f4b - exported functions etype_of and mk_typ
berghofe
parents: 13609
diff changeset
   349
  let
bdd483321f4b - exported functions etype_of and mk_typ
berghofe
parents: 13609
diff changeset
   350
    val thy' = add_syntax thy;
13732
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   351
    val {realizes_eqns, typeof_eqns, defs, types, ...} =
13714
bdd483321f4b - exported functions etype_of and mk_typ
berghofe
parents: 13609
diff changeset
   352
      ExtractionData.get thy';
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   353
    val procs = List.concat (map (fst o snd) types);
13714
bdd483321f4b - exported functions etype_of and mk_typ
berghofe
parents: 13609
diff changeset
   354
    val eqns = Net.merge (#net realizes_eqns, #net typeof_eqns, K false);
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16363
diff changeset
   355
    val prop' = Pattern.rewrite_term (Sign.tsig_of thy')
13714
bdd483321f4b - exported functions etype_of and mk_typ
berghofe
parents: 13609
diff changeset
   356
      (map (Logic.dest_equals o prop_of) defs) [] prop;
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16363
diff changeset
   357
  in freeze_thaw (condrew thy' eqns
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16363
diff changeset
   358
    (procs @ [typeof_proc (Sign.defaultS thy') vs, rlz_proc]))
13714
bdd483321f4b - exported functions etype_of and mk_typ
berghofe
parents: 13609
diff changeset
   359
      (Const ("realizes", fastype_of t --> propT --> propT) $ t $ prop')
bdd483321f4b - exported functions etype_of and mk_typ
berghofe
parents: 13609
diff changeset
   360
  end;
bdd483321f4b - exported functions etype_of and mk_typ
berghofe
parents: 13609
diff changeset
   361
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   362
(** expanding theorems / definitions **)
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   363
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   364
fun add_expand_thm (thy, thm) =
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   365
  let
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   366
    val {realizes_eqns, typeof_eqns, types, realizers,
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   367
      defs, expand, prep} = ExtractionData.get thy;
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   368
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   369
    val name = Thm.name_of_thm thm;
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   370
    val _ = assert (name <> "") "add_expand_thms: unnamed theorem";
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   371
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   372
    val is_def =
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   373
      (case strip_comb (fst (Logic.dest_equals (prop_of thm))) of
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   374
         (Const _, ts) => forall is_Var ts andalso null (duplicates ts)
16349
40c5a4d0b3cc can (Thm.get_axiom_i thy) name;
wenzelm
parents: 16287
diff changeset
   375
           andalso can (Thm.get_axiom_i thy) name
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   376
       | _ => false) handle TERM _ => false;
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   377
  in
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   378
    (ExtractionData.put (if is_def then
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   379
        {realizes_eqns = realizes_eqns,
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   380
         typeof_eqns = add_rule (([],
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   381
           Logic.dest_equals (prop_of (Drule.abs_def thm))), typeof_eqns),
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   382
         types = types,
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   383
         realizers = realizers, defs = gen_ins eq_thm (thm, defs),
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   384
         expand = expand, prep = prep}
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   385
      else
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   386
        {realizes_eqns = realizes_eqns, typeof_eqns = typeof_eqns, types = types,
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   387
         realizers = realizers, defs = defs,
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   388
         expand = (name, prop_of thm) ins expand, prep = prep}) thy, thm)
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   389
  end;
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   390
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   391
fun add_expand_thms thms thy = Library.foldl (fst o add_expand_thm) (thy, thms);
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   392
15801
d2f5ca3c048d superceded by Pure.thy and CPure.thy;
wenzelm
parents: 15798
diff changeset
   393
13732
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   394
(** types with computational content **)
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   395
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   396
fun add_types tys thy =
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   397
  let val {realizes_eqns, typeof_eqns, types, realizers,
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   398
    defs, expand, prep} = ExtractionData.get thy;
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   399
  in
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   400
    ExtractionData.put
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   401
      {realizes_eqns = realizes_eqns, typeof_eqns = typeof_eqns,
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16363
diff changeset
   402
       types = map (apfst (Sign.intern_type thy)) tys @ types,
13732
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   403
       realizers = realizers, defs = defs, expand = expand, prep = prep} thy
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   404
  end;
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   405
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   406
15801
d2f5ca3c048d superceded by Pure.thy and CPure.thy;
wenzelm
parents: 15798
diff changeset
   407
(** Pure setup **)
d2f5ca3c048d superceded by Pure.thy and CPure.thy;
wenzelm
parents: 15798
diff changeset
   408
d2f5ca3c048d superceded by Pure.thy and CPure.thy;
wenzelm
parents: 15798
diff changeset
   409
val _ = Context.add_setup
d2f5ca3c048d superceded by Pure.thy and CPure.thy;
wenzelm
parents: 15798
diff changeset
   410
  [add_types [("prop", ([], NONE))],
d2f5ca3c048d superceded by Pure.thy and CPure.thy;
wenzelm
parents: 15798
diff changeset
   411
d2f5ca3c048d superceded by Pure.thy and CPure.thy;
wenzelm
parents: 15798
diff changeset
   412
   add_typeof_eqns
d2f5ca3c048d superceded by Pure.thy and CPure.thy;
wenzelm
parents: 15798
diff changeset
   413
     ["(typeof (PROP P)) == (Type (TYPE(Null))) ==>  \
d2f5ca3c048d superceded by Pure.thy and CPure.thy;
wenzelm
parents: 15798
diff changeset
   414
    \  (typeof (PROP Q)) == (Type (TYPE('Q))) ==>  \
d2f5ca3c048d superceded by Pure.thy and CPure.thy;
wenzelm
parents: 15798
diff changeset
   415
    \    (typeof (PROP P ==> PROP Q)) == (Type (TYPE('Q)))",
d2f5ca3c048d superceded by Pure.thy and CPure.thy;
wenzelm
parents: 15798
diff changeset
   416
d2f5ca3c048d superceded by Pure.thy and CPure.thy;
wenzelm
parents: 15798
diff changeset
   417
      "(typeof (PROP Q)) == (Type (TYPE(Null))) ==>  \
d2f5ca3c048d superceded by Pure.thy and CPure.thy;
wenzelm
parents: 15798
diff changeset
   418
    \    (typeof (PROP P ==> PROP Q)) == (Type (TYPE(Null)))",
d2f5ca3c048d superceded by Pure.thy and CPure.thy;
wenzelm
parents: 15798
diff changeset
   419
d2f5ca3c048d superceded by Pure.thy and CPure.thy;
wenzelm
parents: 15798
diff changeset
   420
      "(typeof (PROP P)) == (Type (TYPE('P))) ==>  \
d2f5ca3c048d superceded by Pure.thy and CPure.thy;
wenzelm
parents: 15798
diff changeset
   421
    \  (typeof (PROP Q)) == (Type (TYPE('Q))) ==>  \
d2f5ca3c048d superceded by Pure.thy and CPure.thy;
wenzelm
parents: 15798
diff changeset
   422
    \    (typeof (PROP P ==> PROP Q)) == (Type (TYPE('P => 'Q)))",
d2f5ca3c048d superceded by Pure.thy and CPure.thy;
wenzelm
parents: 15798
diff changeset
   423
d2f5ca3c048d superceded by Pure.thy and CPure.thy;
wenzelm
parents: 15798
diff changeset
   424
      "(%x. typeof (PROP P (x))) == (%x. Type (TYPE(Null))) ==>  \
d2f5ca3c048d superceded by Pure.thy and CPure.thy;
wenzelm
parents: 15798
diff changeset
   425
    \    (typeof (!!x. PROP P (x))) == (Type (TYPE(Null)))",
d2f5ca3c048d superceded by Pure.thy and CPure.thy;
wenzelm
parents: 15798
diff changeset
   426
d2f5ca3c048d superceded by Pure.thy and CPure.thy;
wenzelm
parents: 15798
diff changeset
   427
      "(%x. typeof (PROP P (x))) == (%x. Type (TYPE('P))) ==>  \
d2f5ca3c048d superceded by Pure.thy and CPure.thy;
wenzelm
parents: 15798
diff changeset
   428
    \    (typeof (!!x::'a. PROP P (x))) == (Type (TYPE('a => 'P)))",
d2f5ca3c048d superceded by Pure.thy and CPure.thy;
wenzelm
parents: 15798
diff changeset
   429
d2f5ca3c048d superceded by Pure.thy and CPure.thy;
wenzelm
parents: 15798
diff changeset
   430
      "(%x. typeof (f (x))) == (%x. Type (TYPE('f))) ==>  \
d2f5ca3c048d superceded by Pure.thy and CPure.thy;
wenzelm
parents: 15798
diff changeset
   431
    \    (typeof (f)) == (Type (TYPE('f)))"],
d2f5ca3c048d superceded by Pure.thy and CPure.thy;
wenzelm
parents: 15798
diff changeset
   432
d2f5ca3c048d superceded by Pure.thy and CPure.thy;
wenzelm
parents: 15798
diff changeset
   433
   add_realizes_eqns
d2f5ca3c048d superceded by Pure.thy and CPure.thy;
wenzelm
parents: 15798
diff changeset
   434
     ["(typeof (PROP P)) == (Type (TYPE(Null))) ==>  \
d2f5ca3c048d superceded by Pure.thy and CPure.thy;
wenzelm
parents: 15798
diff changeset
   435
    \    (realizes (r) (PROP P ==> PROP Q)) ==  \
d2f5ca3c048d superceded by Pure.thy and CPure.thy;
wenzelm
parents: 15798
diff changeset
   436
    \    (PROP realizes (Null) (PROP P) ==> PROP realizes (r) (PROP Q))",
d2f5ca3c048d superceded by Pure.thy and CPure.thy;
wenzelm
parents: 15798
diff changeset
   437
d2f5ca3c048d superceded by Pure.thy and CPure.thy;
wenzelm
parents: 15798
diff changeset
   438
      "(typeof (PROP P)) == (Type (TYPE('P))) ==>  \
d2f5ca3c048d superceded by Pure.thy and CPure.thy;
wenzelm
parents: 15798
diff changeset
   439
    \  (typeof (PROP Q)) == (Type (TYPE(Null))) ==>  \
d2f5ca3c048d superceded by Pure.thy and CPure.thy;
wenzelm
parents: 15798
diff changeset
   440
    \    (realizes (r) (PROP P ==> PROP Q)) ==  \
d2f5ca3c048d superceded by Pure.thy and CPure.thy;
wenzelm
parents: 15798
diff changeset
   441
    \    (!!x::'P. PROP realizes (x) (PROP P) ==> PROP realizes (Null) (PROP Q))",
d2f5ca3c048d superceded by Pure.thy and CPure.thy;
wenzelm
parents: 15798
diff changeset
   442
d2f5ca3c048d superceded by Pure.thy and CPure.thy;
wenzelm
parents: 15798
diff changeset
   443
      "(realizes (r) (PROP P ==> PROP Q)) ==  \
d2f5ca3c048d superceded by Pure.thy and CPure.thy;
wenzelm
parents: 15798
diff changeset
   444
    \  (!!x. PROP realizes (x) (PROP P) ==> PROP realizes (r (x)) (PROP Q))",
d2f5ca3c048d superceded by Pure.thy and CPure.thy;
wenzelm
parents: 15798
diff changeset
   445
d2f5ca3c048d superceded by Pure.thy and CPure.thy;
wenzelm
parents: 15798
diff changeset
   446
      "(%x. typeof (PROP P (x))) == (%x. Type (TYPE(Null))) ==>  \
d2f5ca3c048d superceded by Pure.thy and CPure.thy;
wenzelm
parents: 15798
diff changeset
   447
    \    (realizes (r) (!!x. PROP P (x))) ==  \
d2f5ca3c048d superceded by Pure.thy and CPure.thy;
wenzelm
parents: 15798
diff changeset
   448
    \    (!!x. PROP realizes (Null) (PROP P (x)))",
d2f5ca3c048d superceded by Pure.thy and CPure.thy;
wenzelm
parents: 15798
diff changeset
   449
d2f5ca3c048d superceded by Pure.thy and CPure.thy;
wenzelm
parents: 15798
diff changeset
   450
      "(realizes (r) (!!x. PROP P (x))) ==  \
d2f5ca3c048d superceded by Pure.thy and CPure.thy;
wenzelm
parents: 15798
diff changeset
   451
    \  (!!x. PROP realizes (r (x)) (PROP P (x)))"],
d2f5ca3c048d superceded by Pure.thy and CPure.thy;
wenzelm
parents: 15798
diff changeset
   452
d2f5ca3c048d superceded by Pure.thy and CPure.thy;
wenzelm
parents: 15798
diff changeset
   453
   Attrib.add_attributes
d2f5ca3c048d superceded by Pure.thy and CPure.thy;
wenzelm
parents: 15798
diff changeset
   454
     [("extraction_expand",
d2f5ca3c048d superceded by Pure.thy and CPure.thy;
wenzelm
parents: 15798
diff changeset
   455
       (Attrib.no_args add_expand_thm, K Attrib.undef_local_attribute),
d2f5ca3c048d superceded by Pure.thy and CPure.thy;
wenzelm
parents: 15798
diff changeset
   456
       "specify theorems / definitions to be expanded during extraction")]];
d2f5ca3c048d superceded by Pure.thy and CPure.thy;
wenzelm
parents: 15798
diff changeset
   457
d2f5ca3c048d superceded by Pure.thy and CPure.thy;
wenzelm
parents: 15798
diff changeset
   458
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   459
(**** extract program ****)
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   460
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   461
val dummyt = Const ("dummy", dummyT);
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   462
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   463
fun extract thms thy =
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   464
  let
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16363
diff changeset
   465
    val thy' = add_syntax thy;
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   466
    val {realizes_eqns, typeof_eqns, types, realizers, defs, expand, prep} =
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   467
      ExtractionData.get thy;
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   468
    val procs = List.concat (map (fst o snd) types);
13732
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   469
    val rtypes = map fst types;
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16363
diff changeset
   470
    val typroc = typeof_proc (Sign.defaultS thy');
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16363
diff changeset
   471
    val prep = getOpt (prep, K I) thy' o ProofRewriteRules.elim_defs thy' false defs o
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16363
diff changeset
   472
      Reconstruct.expand_proof thy' (("", NONE) :: map (apsnd SOME) expand);
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   473
    val rrews = Net.merge (#net realizes_eqns, #net typeof_eqns, K false);
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   474
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   475
    fun find_inst prop Ts ts vs =
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   476
      let
13732
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   477
        val rvs = relevant_vars rtypes prop;
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   478
        val vars = vars_of prop;
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   479
        val n = Int.min (length vars, length ts);
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   480
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   481
        fun add_args ((Var ((a, i), _), t), (vs', tye)) =
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   482
          if a mem rvs then
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16363
diff changeset
   483
            let val T = etype_of thy' vs Ts t
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   484
            in if T = nullT then (vs', tye)
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   485
               else (a :: vs', (("'" ^ a, i), T) :: tye)
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   486
            end
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   487
          else (vs', tye)
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   488
15574
b1d1b5bfc464 Removed practically all references to Library.foldr.
skalberg
parents: 15570
diff changeset
   489
      in foldr add_args ([], []) (Library.take (n, vars) ~~ Library.take (n, ts)) end;
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   490
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   491
    fun find vs = Option.map snd o find_first (curry eq_set vs o fst);
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   492
    fun find' s = map snd o List.filter (equal s o fst)
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   493
13732
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   494
    fun app_rlz_rews Ts vs t = strip_abs (length Ts) (freeze_thaw
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16363
diff changeset
   495
      (condrew thy' rrews (procs @ [typroc vs, rlz_proc])) (list_abs
13732
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   496
        (map (pair "x") (rev Ts), t)));
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   497
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   498
    fun realizes_null vs prop = app_rlz_rews [] vs
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   499
      (Const ("realizes", nullT --> propT --> propT) $ nullt $ prop);
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   500
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   501
    fun corr d defs vs ts Ts hs (PBound i) _ _ = (defs, PBound i)
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   502
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   503
      | corr d defs vs ts Ts hs (Abst (s, SOME T, prf)) (Abst (_, _, prf')) t =
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   504
          let val (defs', corr_prf) = corr d defs vs [] (T :: Ts)
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   505
            (dummyt :: hs) prf (incr_pboundvars 1 0 prf')
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   506
            (case t of SOME (Abs (_, _, u)) => SOME u | _ => NONE)
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   507
          in (defs', Abst (s, SOME T, corr_prf)) end
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   508
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   509
      | corr d defs vs ts Ts hs (AbsP (s, SOME prop, prf)) (AbsP (_, _, prf')) t =
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   510
          let
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16363
diff changeset
   511
            val T = etype_of thy' vs Ts prop;
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   512
            val u = if T = nullT then 
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   513
                (case t of SOME u => SOME (incr_boundvars 1 u) | NONE => NONE)
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   514
              else (case t of SOME (Abs (_, _, u)) => SOME u | _ => NONE);
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   515
            val (defs', corr_prf) = corr d defs vs [] (T :: Ts) (prop :: hs)
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   516
              (incr_pboundvars 0 1 prf) (incr_pboundvars 0 1 prf') u;
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   517
            val rlz = Const ("realizes", T --> propT --> propT)
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   518
          in (defs',
13732
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   519
            if T = nullT then AbsP ("R",
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   520
              SOME (app_rlz_rews Ts vs (rlz $ nullt $ prop)),
13732
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   521
                prf_subst_bounds [nullt] corr_prf)
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   522
            else Abst (s, SOME T, AbsP ("R",
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   523
              SOME (app_rlz_rews (T :: Ts) vs
13732
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   524
                (rlz $ Bound 0 $ incr_boundvars 1 prop)), corr_prf)))
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   525
          end
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   526
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   527
      | corr d defs vs ts Ts hs (prf % SOME t) (prf' % _) t' =
13732
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   528
          let
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   529
            val (Us, T) = strip_type (fastype_of1 (Ts, t));
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   530
            val (defs', corr_prf) = corr d defs vs (t :: ts) Ts hs prf prf'
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   531
              (if tname_of T mem rtypes then t'
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   532
               else (case t' of SOME (u $ _) => SOME u | _ => NONE));
13732
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   533
            val u = if not (tname_of T mem rtypes) then t else
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   534
              let
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16363
diff changeset
   535
                val eT = etype_of thy' vs Ts t;
13732
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   536
                val (r, Us') = if eT = nullT then (nullt, Us) else
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   537
                  (Bound (length Us), eT :: Us);
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   538
                val u = list_comb (incr_boundvars (length Us') t,
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   539
                  map Bound (length Us - 1 downto 0));
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   540
                val u' = (case assoc (types, tname_of T) of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   541
                    SOME ((_, SOME f)) => f r eT u T
13732
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   542
                  | _ => Const ("realizes", eT --> T --> T) $ r $ u)
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   543
              in app_rlz_rews Ts vs (list_abs (map (pair "x") Us', u')) end
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   544
          in (defs', corr_prf % SOME u) end
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   545
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   546
      | corr d defs vs ts Ts hs (prf1 %% prf2) (prf1' %% prf2') t =
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   547
          let
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   548
            val prop = Reconstruct.prop_of' hs prf2';
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16363
diff changeset
   549
            val T = etype_of thy' vs Ts prop;
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   550
            val (defs1, f, u) = if T = nullT then (defs, t, NONE) else
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   551
              (case t of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   552
                 SOME (f $ u) => (defs, SOME f, SOME u)
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   553
               | _ =>
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   554
                 let val (defs1, u) = extr d defs vs [] Ts hs prf2'
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   555
                 in (defs1, NONE, SOME u) end)
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   556
            val (defs2, corr_prf1) = corr d defs1 vs [] Ts hs prf1 prf1' f;
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   557
            val (defs3, corr_prf2) = corr d defs2 vs [] Ts hs prf2 prf2' u;
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   558
          in
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   559
            if T = nullT then (defs3, corr_prf1 %% corr_prf2) else
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   560
              (defs3, corr_prf1 % u %% corr_prf2)
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   561
          end
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   562
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   563
      | corr d defs vs ts Ts hs (prf0 as PThm ((name, _), prf, prop, SOME Ts')) _ _ =
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   564
          let
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   565
            val (vs', tye) = find_inst prop Ts ts vs;
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   566
            val tye' = (map fst (term_tvars prop) ~~ Ts') @ tye;
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16363
diff changeset
   567
            val T = etype_of thy' vs' [] prop;
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   568
            val defs' = if T = nullT then defs
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   569
              else fst (extr d defs vs ts Ts hs prf0)
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   570
          in
13609
73c3915553b4 Added check for axioms with "realizes Null A = A".
berghofe
parents: 13417
diff changeset
   571
            if T = nullT andalso realizes_null vs' prop aconv prop then (defs, prf0)
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   572
            else case Symtab.lookup (realizers, name) of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   573
              NONE => (case find vs' (find' name defs') of
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   574
                NONE =>
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   575
                  let
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   576
                    val _ = assert (T = nullT) "corr: internal error";
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   577
                    val _ = msg d ("Building correctness proof for " ^ quote name ^
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   578
                      (if null vs' then ""
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   579
                       else " (relevant variables: " ^ commas_quote vs' ^ ")"));
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16363
diff changeset
   580
                    val prf' = prep (Reconstruct.reconstruct_proof thy' prop prf);
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   581
                    val (defs'', corr_prf) =
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   582
                      corr (d + 1) defs' vs' [] [] [] prf' prf' NONE;
13732
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   583
                    val corr_prop = Reconstruct.prop_of corr_prf;
15574
b1d1b5bfc464 Removed practically all references to Library.foldr.
skalberg
parents: 15570
diff changeset
   584
                    val corr_prf' = foldr forall_intr_prf
b1d1b5bfc464 Removed practically all references to Library.foldr.
skalberg
parents: 15570
diff changeset
   585
                      (proof_combt
13793
c02c81fd11b8 Fixed bug in function corr.
berghofe
parents: 13732
diff changeset
   586
                         (PThm ((corr_name name vs', []), corr_prf, corr_prop,
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   587
                             SOME (map TVar (term_tvars corr_prop))), vfs_of corr_prop))
15574
b1d1b5bfc464 Removed practically all references to Library.foldr.
skalberg
parents: 15570
diff changeset
   588
		      (map (get_var_type corr_prop) (vfs_of prop))
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   589
                  in
13732
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   590
                    ((name, (vs', ((nullt, nullt), (corr_prf, corr_prf')))) :: defs'',
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   591
                     prf_subst_TVars tye' corr_prf')
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   592
                  end
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   593
              | SOME (_, (_, prf')) => (defs', prf_subst_TVars tye' prf'))
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   594
            | SOME rs => (case find vs' rs of
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   595
                SOME (_, prf') => (defs', prf_subst_TVars tye' prf')
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   596
              | NONE => error ("corr: no realizer for instance of theorem " ^
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16363
diff changeset
   597
                  quote name ^ ":\n" ^ Sign.string_of_term thy' (Envir.beta_norm
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   598
                    (Reconstruct.prop_of (proof_combt (prf0, ts))))))
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   599
          end
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   600
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   601
      | corr d defs vs ts Ts hs (prf0 as PAxm (s, prop, SOME Ts')) _ _ =
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   602
          let
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   603
            val (vs', tye) = find_inst prop Ts ts vs;
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   604
            val tye' = (map fst (term_tvars prop) ~~ Ts') @ tye
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   605
          in
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16363
diff changeset
   606
            if etype_of thy' vs' [] prop = nullT andalso
13609
73c3915553b4 Added check for axioms with "realizes Null A = A".
berghofe
parents: 13417
diff changeset
   607
              realizes_null vs' prop aconv prop then (defs, prf0)
73c3915553b4 Added check for axioms with "realizes Null A = A".
berghofe
parents: 13417
diff changeset
   608
            else case find vs' (Symtab.lookup_multi (realizers, s)) of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   609
              SOME (_, prf) => (defs, prf_subst_TVars tye' prf)
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   610
            | NONE => error ("corr: no realizer for instance of axiom " ^
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16363
diff changeset
   611
                quote s ^ ":\n" ^ Sign.string_of_term thy' (Envir.beta_norm
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   612
                  (Reconstruct.prop_of (proof_combt (prf0, ts)))))
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   613
          end
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   614
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   615
      | corr d defs vs ts Ts hs _ _ _ = error "corr: bad proof"
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   616
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   617
    and extr d defs vs ts Ts hs (PBound i) = (defs, Bound i)
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   618
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   619
      | extr d defs vs ts Ts hs (Abst (s, SOME T, prf)) =
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   620
          let val (defs', t) = extr d defs vs []
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   621
            (T :: Ts) (dummyt :: hs) (incr_pboundvars 1 0 prf)
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   622
          in (defs', Abs (s, T, t)) end
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   623
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   624
      | extr d defs vs ts Ts hs (AbsP (s, SOME t, prf)) =
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   625
          let
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16363
diff changeset
   626
            val T = etype_of thy' vs Ts t;
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   627
            val (defs', t) = extr d defs vs [] (T :: Ts) (t :: hs)
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   628
              (incr_pboundvars 0 1 prf)
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   629
          in (defs',
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   630
            if T = nullT then subst_bound (nullt, t) else Abs (s, T, t))
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   631
          end
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   632
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   633
      | extr d defs vs ts Ts hs (prf % SOME t) =
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   634
          let val (defs', u) = extr d defs vs (t :: ts) Ts hs prf
13732
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   635
          in (defs',
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   636
            if tname_of (body_type (fastype_of1 (Ts, t))) mem rtypes then u
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   637
            else u $ t)
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   638
          end
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   639
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   640
      | extr d defs vs ts Ts hs (prf1 %% prf2) =
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   641
          let
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   642
            val (defs', f) = extr d defs vs [] Ts hs prf1;
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   643
            val prop = Reconstruct.prop_of' hs prf2;
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16363
diff changeset
   644
            val T = etype_of thy' vs Ts prop
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   645
          in
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   646
            if T = nullT then (defs', f) else
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   647
              let val (defs'', t) = extr d defs' vs [] Ts hs prf2
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   648
              in (defs'', f $ t) end
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   649
          end
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   650
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   651
      | extr d defs vs ts Ts hs (prf0 as PThm ((s, _), prf, prop, SOME Ts')) =
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   652
          let
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   653
            val (vs', tye) = find_inst prop Ts ts vs;
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   654
            val tye' = (map fst (term_tvars prop) ~~ Ts') @ tye
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   655
          in
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   656
            case Symtab.lookup (realizers, s) of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   657
              NONE => (case find vs' (find' s defs) of
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   658
                NONE =>
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   659
                  let
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   660
                    val _ = msg d ("Extracting " ^ quote s ^
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   661
                      (if null vs' then ""
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   662
                       else " (relevant variables: " ^ commas_quote vs' ^ ")"));
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16363
diff changeset
   663
                    val prf' = prep (Reconstruct.reconstruct_proof thy' prop prf);
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   664
                    val (defs', t) = extr (d + 1) defs vs' [] [] [] prf';
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   665
                    val (defs'', corr_prf) =
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   666
                      corr (d + 1) defs' vs' [] [] [] prf' prf' (SOME t);
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   667
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   668
                    val nt = Envir.beta_norm t;
13732
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   669
                    val args = filter_out (fn v => tname_of (body_type
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   670
                      (fastype_of v)) mem rtypes) (vfs_of prop);
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   671
                    val args' = List.filter (fn v => Logic.occs (v, nt)) args;
15574
b1d1b5bfc464 Removed practically all references to Library.foldr.
skalberg
parents: 15570
diff changeset
   672
                    val t' = mkabs nt args';
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   673
                    val T = fastype_of t';
13732
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   674
                    val cname = extr_name s vs';
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   675
                    val c = Const (cname, T);
15574
b1d1b5bfc464 Removed practically all references to Library.foldr.
skalberg
parents: 15570
diff changeset
   676
                    val u = mkabs (list_comb (c, args')) args;
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   677
                    val eqn = Logic.mk_equals (c, t');
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   678
                    val rlz =
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   679
                      Const ("realizes", fastype_of nt --> propT --> propT);
13732
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   680
                    val lhs = app_rlz_rews [] vs' (rlz $ nt $ prop);
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   681
                    val rhs = app_rlz_rews [] vs' (rlz $ list_comb (c, args') $ prop);
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   682
                    val f = app_rlz_rews [] vs'
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   683
                      (Abs ("x", T, rlz $ list_comb (Bound 0, args') $ prop));
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   684
13732
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   685
                    val corr_prf' =
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   686
                      chtype [] equal_elim_axm %> lhs %> rhs %%
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   687
                       (chtype [propT] symmetric_axm %> rhs %> lhs %%
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   688
                         (chtype [propT, T] combination_axm %> f %> f %> c %> t' %%
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   689
                           (chtype [T --> propT] reflexive_axm %> f) %%
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   690
                           PAxm (cname ^ "_def", eqn,
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   691
                             SOME (map TVar (term_tvars eqn))))) %% corr_prf;
13732
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   692
                    val corr_prop = Reconstruct.prop_of corr_prf';
15574
b1d1b5bfc464 Removed practically all references to Library.foldr.
skalberg
parents: 15570
diff changeset
   693
                    val corr_prf'' = foldr forall_intr_prf
b1d1b5bfc464 Removed practically all references to Library.foldr.
skalberg
parents: 15570
diff changeset
   694
                      (proof_combt
13732
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   695
                        (PThm ((corr_name s vs', []), corr_prf', corr_prop,
15574
b1d1b5bfc464 Removed practically all references to Library.foldr.
skalberg
parents: 15570
diff changeset
   696
                          SOME (map TVar (term_tvars corr_prop))),  vfs_of corr_prop))
b1d1b5bfc464 Removed practically all references to Library.foldr.
skalberg
parents: 15570
diff changeset
   697
		      (map (get_var_type corr_prop) (vfs_of prop));
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   698
                  in
13732
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   699
                    ((s, (vs', ((t', u), (corr_prf', corr_prf'')))) :: defs'',
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   700
                     subst_TVars tye' u)
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   701
                  end
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   702
              | SOME ((_, u), _) => (defs, subst_TVars tye' u))
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   703
            | SOME rs => (case find vs' rs of
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   704
                SOME (t, _) => (defs, subst_TVars tye' t)
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   705
              | NONE => error ("extr: no realizer for instance of theorem " ^
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16363
diff changeset
   706
                  quote s ^ ":\n" ^ Sign.string_of_term thy' (Envir.beta_norm
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   707
                    (Reconstruct.prop_of (proof_combt (prf0, ts))))))
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   708
          end
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   709
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   710
      | extr d defs vs ts Ts hs (prf0 as PAxm (s, prop, SOME Ts')) =
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   711
          let
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   712
            val (vs', tye) = find_inst prop Ts ts vs;
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   713
            val tye' = (map fst (term_tvars prop) ~~ Ts') @ tye
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   714
          in
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   715
            case find vs' (Symtab.lookup_multi (realizers, s)) of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   716
              SOME (t, _) => (defs, subst_TVars tye' t)
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   717
            | NONE => error ("extr: no realizer for instance of axiom " ^
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16363
diff changeset
   718
                quote s ^ ":\n" ^ Sign.string_of_term thy' (Envir.beta_norm
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   719
                  (Reconstruct.prop_of (proof_combt (prf0, ts)))))
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   720
          end
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   721
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   722
      | extr d defs vs ts Ts hs _ = error "extr: bad proof";
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   723
13732
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   724
    fun prep_thm (thm, vs) =
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   725
      let
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   726
        val {prop, der = (_, prf), sign, ...} = rep_thm thm;
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   727
        val name = Thm.name_of_thm thm;
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   728
        val _ = assert (name <> "") "extraction: unnamed theorem";
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16363
diff changeset
   729
        val _ = assert (etype_of thy' vs [] prop <> nullT) ("theorem " ^
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   730
          quote name ^ " has no computational content")
13732
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   731
      in (Reconstruct.reconstruct_proof sign prop prf, vs) end;
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   732
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   733
    val defs = Library.foldl (fn (defs, (prf, vs)) =>
13732
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   734
      fst (extr 0 defs vs [] [] [] prf)) ([], map prep_thm thms);
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   735
16149
d8cac577493c Theory.restore_naming;
wenzelm
parents: 15801
diff changeset
   736
    fun add_def (s, (vs, ((t, u), (prf, _)))) thy =
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16363
diff changeset
   737
      (case Sign.const_type thy (extr_name s vs) of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   738
         NONE =>
13732
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   739
           let
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   740
             val corr_prop = Reconstruct.prop_of prf;
16287
7a03b4b4df67 Type.freeze;
wenzelm
parents: 16195
diff changeset
   741
             val ft = Type.freeze t;
7a03b4b4df67 Type.freeze;
wenzelm
parents: 16195
diff changeset
   742
             val fu = Type.freeze u;
13732
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   743
             val thy' = if t = nullt then thy else thy |>
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   744
               Theory.add_consts_i [(extr_name s vs, fastype_of ft, NoSyn)] |>
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   745
               fst o PureThy.add_defs_i false [((extr_name s vs ^ "_def",
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   746
                 Logic.mk_equals (head_of (strip_abs_body fu), ft)), [])];
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   747
           in
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   748
             fst (PureThy.store_thm ((corr_name s vs,
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   749
               Thm.varifyT (funpow (length (term_vars corr_prop))
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   750
                 (forall_elim_var 0) (forall_intr_frees
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   751
                   (ProofChecker.thm_of_proof thy'
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   752
                     (fst (Proofterm.freeze_thaw_prf prf)))))), []) thy')
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   753
           end
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   754
       | SOME _ => thy);
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   755
16149
d8cac577493c Theory.restore_naming;
wenzelm
parents: 15801
diff changeset
   756
  in
d8cac577493c Theory.restore_naming;
wenzelm
parents: 15801
diff changeset
   757
    thy
d8cac577493c Theory.restore_naming;
wenzelm
parents: 15801
diff changeset
   758
    |> Theory.absolute_path
d8cac577493c Theory.restore_naming;
wenzelm
parents: 15801
diff changeset
   759
    |> fold_rev add_def defs
d8cac577493c Theory.restore_naming;
wenzelm
parents: 15801
diff changeset
   760
    |> Theory.restore_naming thy
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   761
  end;
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   762
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   763
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   764
(**** interface ****)
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   765
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   766
structure P = OuterParse and K = OuterSyntax.Keyword;
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   767
13732
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   768
val parse_vars = Scan.optional (P.$$$ "(" |-- P.list1 P.name --| P.$$$ ")") [];
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   769
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   770
val realizersP =
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   771
  OuterSyntax.command "realizers"
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   772
  "specify realizers for primitive axioms / theorems, together with correctness proof"
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   773
  K.thy_decl
13732
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   774
    (Scan.repeat1 (P.xname -- parse_vars --| P.$$$ ":" -- P.string -- P.string) >>
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   775
     (fn xs => Toplevel.theory (fn thy => add_realizers
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   776
       (map (fn (((a, vs), s1), s2) =>
16486
1a12cdb6ee6b get_thm(s): Name;
wenzelm
parents: 16458
diff changeset
   777
         (PureThy.get_thm thy (Name a), (vs, s1, s2))) xs) thy)));
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   778
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   779
val realizabilityP =
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   780
  OuterSyntax.command "realizability"
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   781
  "add equations characterizing realizability" K.thy_decl
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   782
  (Scan.repeat1 P.string >> (Toplevel.theory o add_realizes_eqns));
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   783
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   784
val typeofP =
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   785
  OuterSyntax.command "extract_type"
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   786
  "add equations characterizing type of extracted program" K.thy_decl
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   787
  (Scan.repeat1 P.string >> (Toplevel.theory o add_typeof_eqns));
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   788
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   789
val extractP =
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   790
  OuterSyntax.command "extract" "extract terms from proofs" K.thy_decl
13732
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   791
    (Scan.repeat1 (P.xname -- parse_vars) >> (fn xs => Toplevel.theory
16486
1a12cdb6ee6b get_thm(s): Name;
wenzelm
parents: 16458
diff changeset
   792
      (fn thy => extract (map (apfst (PureThy.get_thm thy o Name)) xs) thy)));
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   793
15801
d2f5ca3c048d superceded by Pure.thy and CPure.thy;
wenzelm
parents: 15798
diff changeset
   794
val _ = OuterSyntax.add_parsers [realizersP, realizabilityP, typeofP, extractP];
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   795
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16363
diff changeset
   796
val etype_of = etype_of o add_syntax;
13714
bdd483321f4b - exported functions etype_of and mk_typ
berghofe
parents: 13609
diff changeset
   797
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   798
end;