src/Pure/Proof/extraction.ML
author wenzelm
Sat, 04 Sep 2021 22:17:15 +0200
changeset 74235 dbaed92fd8af
parent 71465 910a081cca74
child 74561 8e6c973003c8
permissions -rw-r--r--
tuned signature;
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
    Author:     Stefan Berghofer, TU Muenchen
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
     3
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
     4
Extraction of programs from proofs.
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
     5
*)
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
     6
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
     7
signature EXTRACTION =
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
     8
sig
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16363
diff changeset
     9
  val set_preprocessor : (theory -> Proofterm.proof -> Proofterm.proof) -> theory -> theory
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
    10
  val add_realizes_eqns_i : ((term * term) list * (term * term)) list -> theory -> theory
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
    11
  val add_realizes_eqns : string list -> theory -> theory
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
    12
  val add_typeof_eqns_i : ((term * term) list * (term * term)) list -> theory -> theory
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
    13
  val add_typeof_eqns : string list -> theory -> theory
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
    14
  val add_realizers_i : (string * (string list * term * Proofterm.proof)) list
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
    15
    -> theory -> theory
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
    16
  val add_realizers : (thm * (string list * string * string)) list
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
    17
    -> theory -> theory
33704
6aeb8454efc1 add_expand_thm: explicit indication of is_def instead of fragile heuristic, tuned signature;
wenzelm
parents: 33522
diff changeset
    18
  val add_expand_thm : bool -> thm -> theory -> theory
13732
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
    19
  val add_types : (xstring * ((term -> term option) list *
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
    20
    (term -> typ -> term -> typ -> term) option)) list -> theory -> theory
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
    21
  val extract : (thm * string list) list -> theory -> theory
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
    22
  val nullT : typ
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
    23
  val nullt : term
13714
bdd483321f4b - exported functions etype_of and mk_typ
berghofe
parents: 13609
diff changeset
    24
  val mk_typ : typ -> term
bdd483321f4b - exported functions etype_of and mk_typ
berghofe
parents: 13609
diff changeset
    25
  val etype_of : theory -> string list -> typ list -> term -> typ
bdd483321f4b - exported functions etype_of and mk_typ
berghofe
parents: 13609
diff changeset
    26
  val realizes_of: theory -> string list -> term -> term -> term
37233
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36953
diff changeset
    27
  val abs_corr_shyps: theory -> thm -> string list -> term list -> Proofterm.proof -> Proofterm.proof
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
(**** tools ****)
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
    34
42407
5b9dd52f5dca prefer internal types, via Simple_Syntax.read_typ;
wenzelm
parents: 42406
diff changeset
    35
val typ = Simple_Syntax.read_typ;
5b9dd52f5dca prefer internal types, via Simple_Syntax.read_typ;
wenzelm
parents: 42406
diff changeset
    36
5b9dd52f5dca prefer internal types, via Simple_Syntax.read_typ;
wenzelm
parents: 42406
diff changeset
    37
val add_syntax =
52788
da1fdbfebd39 type theory is purely value-oriented;
wenzelm
parents: 49960
diff changeset
    38
  Sign.root_path
56436
30ccec1e82fb more source positions;
wenzelm
parents: 56239
diff changeset
    39
  #> Sign.add_types_global
64556
851ae0e7b09c more symbols;
wenzelm
parents: 63239
diff changeset
    40
    [(Binding.make ("Type", \<^here>), 0, NoSyn),
851ae0e7b09c more symbols;
wenzelm
parents: 63239
diff changeset
    41
     (Binding.make ("Null", \<^here>), 0, NoSyn)]
56239
17df7145a871 tuned signature;
wenzelm
parents: 56206
diff changeset
    42
  #> Sign.add_consts
67721
5348bea4accd eliminated ASCII syntax from Pure bootstrap;
wenzelm
parents: 67522
diff changeset
    43
    [(Binding.make ("typeof", \<^here>), typ "'b \<Rightarrow> Type", NoSyn),
5348bea4accd eliminated ASCII syntax from Pure bootstrap;
wenzelm
parents: 67522
diff changeset
    44
     (Binding.make ("Type", \<^here>), typ "'a itself \<Rightarrow> Type", NoSyn),
64556
851ae0e7b09c more symbols;
wenzelm
parents: 63239
diff changeset
    45
     (Binding.make ("Null", \<^here>), typ "Null", NoSyn),
67721
5348bea4accd eliminated ASCII syntax from Pure bootstrap;
wenzelm
parents: 67522
diff changeset
    46
     (Binding.make ("realizes", \<^here>), typ "'a \<Rightarrow> 'b \<Rightarrow> '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 =
19391
4812d28c90a6 Term.itselfT;
wenzelm
parents: 19305
diff changeset
    52
  Const ("Type", Term.itselfT T --> Type ("Type", [])) $ Logic.mk_type T;
13402
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), _), _) =>
20664
ffbc5a57191a member (op =);
wenzelm
parents: 20548
diff changeset
    57
            if member (op =) vs a then TFree ("'" ^ a ^ ":" ^ string_of_int i, defaultS)
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
    58
            else nullT
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
    59
        | (Free (a, _), _) =>
20664
ffbc5a57191a member (op =);
wenzelm
parents: 20548
diff changeset
    60
            if member (op =) vs a then TFree ("'" ^ a, defaultS) else nullT
13402
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
49960
1167c1157a5b more conventional argument order;
haftmann
parents: 48704
diff changeset
    64
fun rlz_proc (Const ("realizes", Type (_, [Type ("Null", []), _])) $ _ $ 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
67522
9e712280cc37 clarified take/drop/chop prefix/suffix;
wenzelm
parents: 67147
diff changeset
    73
  chop_prefix (fn s => s <> ":") o raw_explode;
13402
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
33337
9c3b9bf81e8b eliminated some old folds;
wenzelm
parents: 33317
diff changeset
    81
fun add_rule (r as (_, (lhs, _))) ({next, rs, net} : rules) =
16800
90eff1b52428 improved Net interface;
wenzelm
parents: 16790
diff changeset
    82
  {next = next - 1, rs = r :: rs, net = Net.insert_term (K false)
18956
c050ae1f8f52 Envir.(beta_)eta_contract;
wenzelm
parents: 18928
diff changeset
    83
     (Envir.eta_contract lhs, (next, r)) net};
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
    84
32784
1a5dde5079ac eliminated redundant bindings;
wenzelm
parents: 32738
diff changeset
    85
fun merge_rules ({next, rs = rs1, net} : rules) ({rs = rs2, ...} : rules) =
33337
9c3b9bf81e8b eliminated some old folds;
wenzelm
parents: 33317
diff changeset
    86
  fold_rev add_rule (subtract (op =) rs1 rs2) {next = next, rs = rs1, net = net};
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
    87
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16363
diff changeset
    88
fun condrew thy rules procs =
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
    89
  let
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
    90
    fun rew tm =
17203
29b2563f5c11 refer to theory instead of low-level tsig;
wenzelm
parents: 17057
diff changeset
    91
      Pattern.rewrite_term thy [] (condrew' :: procs) tm
15399
683d83051d6a Added term cache to function condrew in order to speed up rewriting.
berghofe
parents: 14981
diff changeset
    92
    and condrew' tm =
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
    93
      let
32738
15bb09ca0378 explicit indication of Unsynchronized.ref;
wenzelm
parents: 32035
diff changeset
    94
        val cache = Unsynchronized.ref ([] : (term * term) list);
17232
148c241d2492 some 'assoc' etc. refactoring
haftmann
parents: 17223
diff changeset
    95
        fun lookup f x = (case AList.lookup (op =) (!cache) x of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
    96
            NONE =>
15399
683d83051d6a Added term cache to function condrew in order to speed up rewriting.
berghofe
parents: 14981
diff changeset
    97
              let val y = f x
683d83051d6a Added term cache to function condrew in order to speed up rewriting.
berghofe
parents: 14981
diff changeset
    98
              in (cache := (x, y) :: !cache; y) end
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
    99
          | SOME y => y);
15399
683d83051d6a Added term cache to function condrew in order to speed up rewriting.
berghofe
parents: 14981
diff changeset
   100
      in
683d83051d6a Added term cache to function condrew in order to speed up rewriting.
berghofe
parents: 14981
diff changeset
   101
        get_first (fn (_, (prems, (tm1, tm2))) =>
683d83051d6a Added term cache to function condrew in order to speed up rewriting.
berghofe
parents: 14981
diff changeset
   102
        let
19466
wenzelm
parents: 19391
diff changeset
   103
          fun ren t = the_default t (Term.rename_abs tm1 tm t);
59787
6e2a20486897 local fixes may depend on goal params;
wenzelm
parents: 59582
diff changeset
   104
          val inc = Logic.incr_indexes ([], [], maxidx_of_term tm + 1);
18184
43c4589a9a78 tuned Pattern.match/unify;
wenzelm
parents: 17412
diff changeset
   105
          val env as (Tenv, tenv) = Pattern.match thy (inc tm1, tm) (Vartab.empty, Vartab.empty);
59058
a78612c67ec0 renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents: 58950
diff changeset
   106
          val prems' = map (apply2 (Envir.subst_term 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
   107
          val env' = Envir.Envir
32032
a6a6e8031c14 tuned/modernized Envir operations;
wenzelm
parents: 30718
diff changeset
   108
            {maxidx = fold (fn (t, u) => Term.maxidx_term t #> Term.maxidx_term u) prems' ~1,
a6a6e8031c14 tuned/modernized Envir operations;
wenzelm
parents: 30718
diff changeset
   109
             tenv = tenv, tyenv = Tenv};
59058
a78612c67ec0 renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents: 58950
diff changeset
   110
          val env'' = fold (Pattern.unify (Context.Theory thy) o apply2 (lookup rew)) prems' env';
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   111
        in SOME (Envir.norm_term env'' (inc (ren tm2)))
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   112
        end handle Pattern.MATCH => NONE | Pattern.Unif => NONE)
59058
a78612c67ec0 renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents: 58950
diff changeset
   113
          (sort (int_ord o apply2 fst)
18956
c050ae1f8f52 Envir.(beta_)eta_contract;
wenzelm
parents: 18928
diff changeset
   114
            (Net.match_term rules (Envir.eta_contract tm)))
15399
683d83051d6a Added term cache to function condrew in order to speed up rewriting.
berghofe
parents: 14981
diff changeset
   115
      end;
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   116
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   117
  in rew end;
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   118
70806
wenzelm
parents: 70589
diff changeset
   119
val change_types = Proofterm.change_types o SOME;
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   120
30364
577edc39b501 moved basic algebra of long names from structure NameSpace to Long_Name;
wenzelm
parents: 30344
diff changeset
   121
fun extr_name s vs = Long_Name.append "extr" (space_implode "_" (s :: vs));
16195
wenzelm
parents: 16149
diff changeset
   122
fun corr_name s vs = extr_name s vs ^ "_correctness";
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   123
61865
6dcc9e4f1aa6 tuned signature;
wenzelm
parents: 61094
diff changeset
   124
fun msg d s = writeln (Symbol.spaces d ^ s);
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   125
74235
dbaed92fd8af tuned signature;
wenzelm
parents: 71465
diff changeset
   126
fun vars_of t = map Var (build_rev (Term.add_vars t));
dbaed92fd8af tuned signature;
wenzelm
parents: 71465
diff changeset
   127
fun frees_of t = map Free (build_rev (Term.add_frees t));
28812
413695e07bd4 Frees in PThms are now quantified in the order of their appearance in the
berghofe
parents: 28805
diff changeset
   128
fun vfs_of t = vars_of t @ frees_of t;
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   129
37233
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36953
diff changeset
   130
val mkabs = fold_rev (fn v => fn t => Abs ("x", fastype_of v, abstract_over (v, t)));
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   131
37233
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36953
diff changeset
   132
val mkabsp = fold_rev (fn t => fn prf => AbsP ("H", SOME t, prf));
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   133
13732
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   134
fun strip_abs 0 t = t
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   135
  | strip_abs n (Abs (_, _, t)) = strip_abs (n-1) t
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   136
  | strip_abs _ _ = error "strip_abs: not an abstraction";
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   137
37310
96e2b9a6f074 do not open Proofterm, which is very ould style;
wenzelm
parents: 37237
diff changeset
   138
val prf_subst_TVars = Proofterm.map_proof_types o typ_subst_TVars;
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   139
40844
5895c525739d more direct use of binder_types/body_type;
wenzelm
parents: 40627
diff changeset
   140
fun relevant_vars types prop =
5895c525739d more direct use of binder_types/body_type;
wenzelm
parents: 40627
diff changeset
   141
  List.foldr
5895c525739d more direct use of binder_types/body_type;
wenzelm
parents: 40627
diff changeset
   142
    (fn (Var ((a, _), T), vs) =>
5895c525739d more direct use of binder_types/body_type;
wenzelm
parents: 40627
diff changeset
   143
        (case body_type T of
5895c525739d more direct use of binder_types/body_type;
wenzelm
parents: 40627
diff changeset
   144
          Type (s, _) => if member (op =) types s then a :: vs else vs
5895c525739d more direct use of binder_types/body_type;
wenzelm
parents: 40627
diff changeset
   145
        | _ => vs)
5895c525739d more direct use of binder_types/body_type;
wenzelm
parents: 40627
diff changeset
   146
      | (_, vs) => vs) [] (vars_of prop);
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   147
13732
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   148
fun tname_of (Type (s, _)) = s
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   149
  | tname_of _ = "";
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   150
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   151
fun get_var_type t =
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   152
  let
16865
fb39dcfc1c24 tuned fold on terms;
wenzelm
parents: 16800
diff changeset
   153
    val vs = Term.add_vars t [];
fb39dcfc1c24 tuned fold on terms;
wenzelm
parents: 16800
diff changeset
   154
    val fs = Term.add_frees t [];
42406
05f2468d6b36 eliminated obsolete Proof_Syntax.strip_sorts_consttypes;
wenzelm
parents: 42375
diff changeset
   155
  in
05f2468d6b36 eliminated obsolete Proof_Syntax.strip_sorts_consttypes;
wenzelm
parents: 42375
diff changeset
   156
    fn Var (ixn, _) =>
05f2468d6b36 eliminated obsolete Proof_Syntax.strip_sorts_consttypes;
wenzelm
parents: 42375
diff changeset
   157
        (case AList.lookup (op =) vs ixn of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   158
          NONE => error "get_var_type: no such variable in term"
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   159
        | SOME T => Var (ixn, T))
42406
05f2468d6b36 eliminated obsolete Proof_Syntax.strip_sorts_consttypes;
wenzelm
parents: 42375
diff changeset
   160
     | Free (s, _) =>
05f2468d6b36 eliminated obsolete Proof_Syntax.strip_sorts_consttypes;
wenzelm
parents: 42375
diff changeset
   161
        (case AList.lookup (op =) fs s of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   162
          NONE => error "get_var_type: no such variable in term"
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   163
        | SOME T => Free (s, T))
13732
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   164
    | _ => error "get_var_type: not a variable"
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   165
  end;
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   166
62922
96691631c1eb clarified context;
wenzelm
parents: 61865
diff changeset
   167
fun read_term ctxt T s =
37233
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36953
diff changeset
   168
  let
62922
96691631c1eb clarified context;
wenzelm
parents: 61865
diff changeset
   169
    val ctxt' = ctxt
62958
b41c1cb5e251 Type_Infer.object_logic controls improvement of type inference result;
wenzelm
parents: 62922
diff changeset
   170
      |> Proof_Context.set_defsort []
b41c1cb5e251 Type_Infer.object_logic controls improvement of type inference result;
wenzelm
parents: 62922
diff changeset
   171
      |> Config.put Type_Infer.object_logic false
b41c1cb5e251 Type_Infer.object_logic controls improvement of type inference result;
wenzelm
parents: 62922
diff changeset
   172
      |> Config.put Type_Infer_Context.const_sorts false;
37233
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36953
diff changeset
   173
    val parse = if T = propT then Syntax.parse_prop else Syntax.parse_term;
62922
96691631c1eb clarified context;
wenzelm
parents: 61865
diff changeset
   174
  in parse ctxt' s |> Type.constraint T |> Syntax.check_term ctxt' end;
37233
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36953
diff changeset
   175
70589
00b67aaa4010 clarified modules;
wenzelm
parents: 70577
diff changeset
   176
fun make_proof_body prf =
00b67aaa4010 clarified modules;
wenzelm
parents: 70577
diff changeset
   177
  let
00b67aaa4010 clarified modules;
wenzelm
parents: 70577
diff changeset
   178
    val (oracles, thms) =
00b67aaa4010 clarified modules;
wenzelm
parents: 70577
diff changeset
   179
      ([prf], ([], [])) |-> Proofterm.fold_proof_atoms false
71465
910a081cca74 more position information for oracles (e.g. "skip_proof" for 'sorry'), requires Proofterm.proofs := 1;
wenzelm
parents: 71214
diff changeset
   180
        (fn Oracle (name, prop, _) => apfst (cons ((name, Position.none), SOME prop))
70589
00b67aaa4010 clarified modules;
wenzelm
parents: 70577
diff changeset
   181
          | PThm (header, thm_body) => apsnd (cons (Proofterm.make_thm header thm_body))
00b67aaa4010 clarified modules;
wenzelm
parents: 70577
diff changeset
   182
          | _ => I);
00b67aaa4010 clarified modules;
wenzelm
parents: 70577
diff changeset
   183
    val body =
00b67aaa4010 clarified modules;
wenzelm
parents: 70577
diff changeset
   184
      PBody
00b67aaa4010 clarified modules;
wenzelm
parents: 70577
diff changeset
   185
       {oracles = Ord_List.make Proofterm.oracle_ord oracles,
00b67aaa4010 clarified modules;
wenzelm
parents: 70577
diff changeset
   186
        thms = Ord_List.make Proofterm.thm_ord thms,
00b67aaa4010 clarified modules;
wenzelm
parents: 70577
diff changeset
   187
        proof = prf};
00b67aaa4010 clarified modules;
wenzelm
parents: 70577
diff changeset
   188
  in Proofterm.thm_body body end;
00b67aaa4010 clarified modules;
wenzelm
parents: 70577
diff changeset
   189
00b67aaa4010 clarified modules;
wenzelm
parents: 70577
diff changeset
   190
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   191
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   192
(**** theory data ****)
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   193
22846
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22796
diff changeset
   194
(* theory data *)
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   195
33522
737589bb9bb8 adapted Theory_Data;
wenzelm
parents: 33388
diff changeset
   196
structure ExtractionData = Theory_Data
22846
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22796
diff changeset
   197
(
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   198
  type T =
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   199
    {realizes_eqns : rules,
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   200
     typeof_eqns : rules,
13732
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   201
     types : (string * ((term -> term option) list *
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   202
       (term -> typ -> term -> typ -> term) option)) list,
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   203
     realizers : (string list * (term * proof)) list Symtab.table,
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   204
     defs : thm list,
37233
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36953
diff changeset
   205
     expand : string list,
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16363
diff changeset
   206
     prep : (theory -> proof -> proof) option}
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   207
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   208
  val empty =
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   209
    {realizes_eqns = empty_rules,
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   210
     typeof_eqns = empty_rules,
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   211
     types = [],
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   212
     realizers = Symtab.empty,
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   213
     defs = [],
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   214
     expand = [],
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   215
     prep = NONE};
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16363
diff changeset
   216
  val extend = I;
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   217
33522
737589bb9bb8 adapted Theory_Data;
wenzelm
parents: 33388
diff changeset
   218
  fun merge
737589bb9bb8 adapted Theory_Data;
wenzelm
parents: 33388
diff changeset
   219
    ({realizes_eqns = realizes_eqns1, typeof_eqns = typeof_eqns1, types = types1,
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   220
       realizers = realizers1, defs = defs1, expand = expand1, prep = prep1},
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   221
      {realizes_eqns = realizes_eqns2, typeof_eqns = typeof_eqns2, types = types2,
33522
737589bb9bb8 adapted Theory_Data;
wenzelm
parents: 33388
diff changeset
   222
       realizers = realizers2, defs = defs2, expand = expand2, prep = prep2}) : T =
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   223
    {realizes_eqns = merge_rules realizes_eqns1 realizes_eqns2,
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   224
     typeof_eqns = merge_rules typeof_eqns1 typeof_eqns2,
22717
74dbc7696083 canonical merge operations
haftmann
parents: 22675
diff changeset
   225
     types = AList.merge (op =) (K true) (types1, types2),
59058
a78612c67ec0 renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents: 58950
diff changeset
   226
     realizers = Symtab.merge_list (eq_set (op =) o apply2 #1) (realizers1, realizers2),
22662
3e492ba59355 canonical merge operations
haftmann
parents: 22596
diff changeset
   227
     defs = Library.merge Thm.eq_thm (defs1, defs2),
37233
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36953
diff changeset
   228
     expand = Library.merge (op =) (expand1, expand2),
38761
b32975d3db3e theory data merge: prefer left side uniformly;
wenzelm
parents: 37310
diff changeset
   229
     prep = if is_some prep1 then prep1 else prep2};
22846
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22796
diff changeset
   230
);
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   231
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   232
fun read_condeq thy =
62922
96691631c1eb clarified context;
wenzelm
parents: 61865
diff changeset
   233
  let val ctxt' = Proof_Context.init_global (add_syntax thy)
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   234
  in fn s =>
62922
96691631c1eb clarified context;
wenzelm
parents: 61865
diff changeset
   235
    let val t = Logic.varify_global (read_term ctxt' propT s)
35424
08c37d7bd2ad more precise scope of exception handler;
wenzelm
parents: 33957
diff changeset
   236
    in
08c37d7bd2ad more precise scope of exception handler;
wenzelm
parents: 33957
diff changeset
   237
      (map Logic.dest_equals (Logic.strip_imp_prems t),
08c37d7bd2ad more precise scope of exception handler;
wenzelm
parents: 33957
diff changeset
   238
        Logic.dest_equals (Logic.strip_imp_concl t))
08c37d7bd2ad more precise scope of exception handler;
wenzelm
parents: 33957
diff changeset
   239
      handle TERM _ => error ("Not a (conditional) meta equality:\n" ^ s)
08c37d7bd2ad more precise scope of exception handler;
wenzelm
parents: 33957
diff changeset
   240
    end
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   241
  end;
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   242
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   243
(** preprocessor **)
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   244
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   245
fun set_preprocessor prep thy =
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   246
  let val {realizes_eqns, typeof_eqns, types, realizers,
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   247
    defs, expand, ...} = ExtractionData.get thy
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   248
  in
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   249
    ExtractionData.put
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   250
      {realizes_eqns = realizes_eqns, typeof_eqns = typeof_eqns, types = types,
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   251
       realizers = realizers, defs = defs, expand = expand, prep = SOME prep} thy
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   252
  end;
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   253
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   254
(** equations characterizing realizability **)
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   255
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   256
fun gen_add_realizes_eqns prep_eq eqns thy =
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   257
  let val {realizes_eqns, typeof_eqns, types, realizers,
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   258
    defs, expand, prep} = ExtractionData.get thy;
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   259
  in
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   260
    ExtractionData.put
33337
9c3b9bf81e8b eliminated some old folds;
wenzelm
parents: 33317
diff changeset
   261
      {realizes_eqns = fold_rev add_rule (map (prep_eq thy) eqns) realizes_eqns,
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   262
       typeof_eqns = typeof_eqns, types = types, realizers = realizers,
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   263
       defs = defs, expand = expand, prep = prep} thy
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   264
  end
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   265
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   266
val add_realizes_eqns_i = gen_add_realizes_eqns (K I);
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   267
val add_realizes_eqns = gen_add_realizes_eqns read_condeq;
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   268
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   269
(** equations characterizing type of extracted program **)
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   270
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   271
fun gen_add_typeof_eqns prep_eq eqns thy =
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   272
  let
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   273
    val {realizes_eqns, typeof_eqns, types, realizers,
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   274
      defs, expand, prep} = ExtractionData.get thy;
13732
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   275
    val eqns' = map (prep_eq thy) eqns
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   276
  in
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   277
    ExtractionData.put
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   278
      {realizes_eqns = realizes_eqns, realizers = realizers,
33337
9c3b9bf81e8b eliminated some old folds;
wenzelm
parents: 33317
diff changeset
   279
       typeof_eqns = fold_rev add_rule eqns' typeof_eqns,
13732
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   280
       types = types, defs = defs, expand = expand, prep = prep} thy
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   281
  end
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   282
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   283
val add_typeof_eqns_i = gen_add_typeof_eqns (K I);
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   284
val add_typeof_eqns = gen_add_typeof_eqns read_condeq;
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   285
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   286
fun thaw (T as TFree (a, S)) =
28375
c879d88d038a eliminated polymorphic equality;
wenzelm
parents: 28370
diff changeset
   287
      if exists_string (fn s => s = ":") a then TVar (unpack_ixn a, S) else T
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   288
  | thaw (Type (a, Ts)) = Type (a, map thaw Ts)
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   289
  | thaw T = T;
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   290
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   291
fun freeze (TVar ((a, i), S)) = TFree (a ^ ":" ^ string_of_int i, S)
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   292
  | freeze (Type (a, Ts)) = Type (a, map freeze Ts)
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   293
  | freeze T = T;
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   294
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   295
fun freeze_thaw f x =
20548
8ef25fe585a8 renamed Term.map_term_types to Term.map_types (cf. Term.fold_types);
wenzelm
parents: 19482
diff changeset
   296
  map_types thaw (f (map_types freeze x));
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   297
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16363
diff changeset
   298
fun etype_of thy vs Ts t =
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   299
  let
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16363
diff changeset
   300
    val {typeof_eqns, ...} = ExtractionData.get thy;
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   301
    fun err () = error ("Unable to determine type of extracted program for\n" ^
26939
1035c89b4c02 moved global pretty/string_of functions from Sign to Syntax;
wenzelm
parents: 26653
diff changeset
   302
      Syntax.string_of_term_global thy t)
46219
426ed18eba43 discontinued old-style Term.list_abs in favour of plain Term.abs;
wenzelm
parents: 44057
diff changeset
   303
  in
426ed18eba43 discontinued old-style Term.list_abs in favour of plain Term.abs;
wenzelm
parents: 44057
diff changeset
   304
    (case
426ed18eba43 discontinued old-style Term.list_abs in favour of plain Term.abs;
wenzelm
parents: 44057
diff changeset
   305
      strip_abs_body
426ed18eba43 discontinued old-style Term.list_abs in favour of plain Term.abs;
wenzelm
parents: 44057
diff changeset
   306
        (freeze_thaw (condrew thy (#net typeof_eqns) [typeof_proc [] vs])
426ed18eba43 discontinued old-style Term.list_abs in favour of plain Term.abs;
wenzelm
parents: 44057
diff changeset
   307
          (fold (Term.abs o pair "x") Ts
426ed18eba43 discontinued old-style Term.list_abs in favour of plain Term.abs;
wenzelm
parents: 44057
diff changeset
   308
            (Const ("typeof", fastype_of1 (Ts, t) --> Type ("Type", [])) $ t))) of
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   309
      Const ("Type", _) $ u => (Logic.dest_type u handle TERM _ => err ())
46219
426ed18eba43 discontinued old-style Term.list_abs in favour of plain Term.abs;
wenzelm
parents: 44057
diff changeset
   310
    | _ => err ())
13402
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
(** realizers for axioms / theorems, together with correctness proofs **)
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   314
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   315
fun gen_add_realizers prep_rlz rs thy =
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   316
  let val {realizes_eqns, typeof_eqns, types, realizers,
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   317
    defs, expand, prep} = ExtractionData.get thy
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   318
  in
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   319
    ExtractionData.put
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   320
      {realizes_eqns = realizes_eqns, typeof_eqns = typeof_eqns, types = types,
25389
3e58c7cb5a73 renamed Symtab.update_list to Symtab.cons_list;
wenzelm
parents: 24867
diff changeset
   321
       realizers = fold (Symtab.cons_list o prep_rlz thy) rs realizers,
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   322
       defs = defs, expand = expand, prep = prep} thy
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   323
  end
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   324
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   325
fun prep_realizer thy =
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   326
  let
13732
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   327
    val {realizes_eqns, typeof_eqns, defs, types, ...} =
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   328
      ExtractionData.get thy;
19482
9f11af8f7ef9 tuned basic list operators (flat, maps, map_filter);
wenzelm
parents: 19466
diff changeset
   329
    val procs = maps (fst o snd) types;
13732
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   330
    val rtypes = map fst types;
16800
90eff1b52428 improved Net interface;
wenzelm
parents: 16790
diff changeset
   331
    val eqns = Net.merge (K false) (#net realizes_eqns, #net typeof_eqns);
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   332
    val thy' = add_syntax thy;
62922
96691631c1eb clarified context;
wenzelm
parents: 61865
diff changeset
   333
    val ctxt' = Proof_Context.init_global thy';
37233
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36953
diff changeset
   334
    val rd = Proof_Syntax.read_proof thy' true false;
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   335
  in fn (thm, (vs, s1, s2)) =>
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   336
    let
36744
6e1f3d609a68 renamed Thm.get_name -> Thm.derivation_name and Thm.put_name -> Thm.name_derivation, to emphasize the true nature of these operations;
wenzelm
parents: 36620
diff changeset
   337
      val name = Thm.derivation_name thm;
21858
05f57309170c avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
wenzelm
parents: 21646
diff changeset
   338
      val _ = name <> "" orelse error "add_realizers: unnamed theorem";
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59058
diff changeset
   339
      val prop = Thm.unconstrainT thm |> Thm.prop_of |>
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59058
diff changeset
   340
        Pattern.rewrite_term thy' (map (Logic.dest_equals o Thm.prop_of) defs) [];
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   341
      val vars = vars_of prop;
13732
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   342
      val vars' = filter_out (fn v =>
20664
ffbc5a57191a member (op =);
wenzelm
parents: 20548
diff changeset
   343
        member (op =) rtypes (tname_of (body_type (fastype_of v)))) vars;
37233
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36953
diff changeset
   344
      val shyps = maps (fn Var ((x, i), _) =>
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36953
diff changeset
   345
        if member (op =) vs x then Logic.mk_of_sort
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36953
diff changeset
   346
          (TVar (("'" ^ x, i), []), Sign.defaultS thy')
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36953
diff changeset
   347
        else []) vars;
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16363
diff changeset
   348
      val T = etype_of thy' vs [] prop;
33832
cff42395c246 explicitly mark some legacy freeze operations;
wenzelm
parents: 33704
diff changeset
   349
      val (T', thw) = Type.legacy_freeze_thaw_type
13732
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   350
        (if T = nullT then nullT else map fastype_of vars' ---> T);
62922
96691631c1eb clarified context;
wenzelm
parents: 61865
diff changeset
   351
      val t = map_types thw (read_term ctxt' T' s1);
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16363
diff changeset
   352
      val r' = freeze_thaw (condrew thy' eqns
37233
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36953
diff changeset
   353
        (procs @ [typeof_proc [] vs, rlz_proc]))
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   354
          (Const ("realizes", T --> propT --> propT) $
13732
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   355
            (if T = nullT then t else list_comb (t, vars')) $ prop);
37233
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36953
diff changeset
   356
      val r = Logic.list_implies (shyps,
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36953
diff changeset
   357
        fold_rev Logic.all (map (get_var_type r') vars) r');
70449
6e34025981be clarified global theory context;
wenzelm
parents: 70447
diff changeset
   358
      val prf = Proofterm.reconstruct_proof thy' r (rd s2);
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   359
    in (name, (vs, (t, prf))) end
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   360
  end;
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   361
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   362
val add_realizers_i = gen_add_realizers
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   363
  (fn _ => fn (name, (vs, t, prf)) => (name, (vs, (t, prf))));
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   364
val add_realizers = gen_add_realizers prep_realizer;
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   365
13714
bdd483321f4b - exported functions etype_of and mk_typ
berghofe
parents: 13609
diff changeset
   366
fun realizes_of thy vs t prop =
bdd483321f4b - exported functions etype_of and mk_typ
berghofe
parents: 13609
diff changeset
   367
  let
bdd483321f4b - exported functions etype_of and mk_typ
berghofe
parents: 13609
diff changeset
   368
    val thy' = add_syntax thy;
13732
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   369
    val {realizes_eqns, typeof_eqns, defs, types, ...} =
13714
bdd483321f4b - exported functions etype_of and mk_typ
berghofe
parents: 13609
diff changeset
   370
      ExtractionData.get thy';
22717
74dbc7696083 canonical merge operations
haftmann
parents: 22675
diff changeset
   371
    val procs = maps (rev o fst o snd) types;
16800
90eff1b52428 improved Net interface;
wenzelm
parents: 16790
diff changeset
   372
    val eqns = Net.merge (K false) (#net realizes_eqns, #net typeof_eqns);
17203
29b2563f5c11 refer to theory instead of low-level tsig;
wenzelm
parents: 17057
diff changeset
   373
    val prop' = Pattern.rewrite_term thy'
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59058
diff changeset
   374
      (map (Logic.dest_equals o Thm.prop_of) defs) [] prop;
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16363
diff changeset
   375
  in freeze_thaw (condrew thy' eqns
37233
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36953
diff changeset
   376
    (procs @ [typeof_proc [] vs, rlz_proc]))
13714
bdd483321f4b - exported functions etype_of and mk_typ
berghofe
parents: 13609
diff changeset
   377
      (Const ("realizes", fastype_of t --> propT --> propT) $ t $ prop')
bdd483321f4b - exported functions etype_of and mk_typ
berghofe
parents: 13609
diff changeset
   378
  end;
bdd483321f4b - exported functions etype_of and mk_typ
berghofe
parents: 13609
diff changeset
   379
37233
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36953
diff changeset
   380
fun abs_corr_shyps thy thm vs xs prf =
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36953
diff changeset
   381
  let
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36953
diff changeset
   382
    val S = Sign.defaultS thy;
70436
251f1fb44ccd clarified signature;
wenzelm
parents: 70435
diff changeset
   383
    val (ucontext, prop') =
61039
80f40d89dab6 tuned signature;
wenzelm
parents: 60826
diff changeset
   384
      Logic.unconstrainT (Thm.shyps_of thm) (Thm.prop_of thm);
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59058
diff changeset
   385
    val atyps = fold_types (fold_atyps (insert (op =))) (Thm.prop_of thm) [];
37233
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36953
diff changeset
   386
    val Ts = map_filter (fn ((v, i), _) => if member (op =) vs v then
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36953
diff changeset
   387
        SOME (TVar (("'" ^ v, i), [])) else NONE)
74235
dbaed92fd8af tuned signature;
wenzelm
parents: 71465
diff changeset
   388
      (build_rev (Term.add_vars prop'));
37233
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36953
diff changeset
   389
    val cs = maps (fn T => map (pair T) S) Ts;
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36953
diff changeset
   390
    val constraints' = map Logic.mk_of_class cs;
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36953
diff changeset
   391
    fun typ_map T = Type.strip_sorts
70436
251f1fb44ccd clarified signature;
wenzelm
parents: 70435
diff changeset
   392
      (map_atyps (fn U => if member (op =) atyps U then (#atyp_map ucontext) U else U) T);
37233
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36953
diff changeset
   393
    fun mk_hyp (T, c) = Hyp (Logic.mk_of_class (typ_map T, c));
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36953
diff changeset
   394
    val xs' = map (map_types typ_map) xs
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36953
diff changeset
   395
  in
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36953
diff changeset
   396
    prf |>
37310
96e2b9a6f074 do not open Proofterm, which is very ould style;
wenzelm
parents: 37237
diff changeset
   397
    Same.commit (Proofterm.map_proof_same (map_types typ_map) typ_map mk_hyp) |>
70436
251f1fb44ccd clarified signature;
wenzelm
parents: 70435
diff changeset
   398
    fold_rev Proofterm.implies_intr_proof' (map snd (#constraints ucontext)) |>
37310
96e2b9a6f074 do not open Proofterm, which is very ould style;
wenzelm
parents: 37237
diff changeset
   399
    fold_rev Proofterm.forall_intr_proof' xs' |>
96e2b9a6f074 do not open Proofterm, which is very ould style;
wenzelm
parents: 37237
diff changeset
   400
    fold_rev Proofterm.implies_intr_proof' constraints'
37233
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36953
diff changeset
   401
  end;
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36953
diff changeset
   402
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   403
(** expanding theorems / definitions **)
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   404
33704
6aeb8454efc1 add_expand_thm: explicit indication of is_def instead of fragile heuristic, tuned signature;
wenzelm
parents: 33522
diff changeset
   405
fun add_expand_thm is_def thm thy =
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   406
  let
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   407
    val {realizes_eqns, typeof_eqns, types, realizers,
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   408
      defs, expand, prep} = ExtractionData.get thy;
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   409
36744
6e1f3d609a68 renamed Thm.get_name -> Thm.derivation_name and Thm.put_name -> Thm.name_derivation, to emphasize the true nature of these operations;
wenzelm
parents: 36620
diff changeset
   410
    val name = Thm.derivation_name thm;
33704
6aeb8454efc1 add_expand_thm: explicit indication of is_def instead of fragile heuristic, tuned signature;
wenzelm
parents: 33522
diff changeset
   411
    val _ = name <> "" orelse error "add_expand_thm: unnamed theorem";
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   412
  in
33704
6aeb8454efc1 add_expand_thm: explicit indication of is_def instead of fragile heuristic, tuned signature;
wenzelm
parents: 33522
diff changeset
   413
    thy |> ExtractionData.put
6aeb8454efc1 add_expand_thm: explicit indication of is_def instead of fragile heuristic, tuned signature;
wenzelm
parents: 33522
diff changeset
   414
      (if is_def then
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   415
        {realizes_eqns = realizes_eqns,
37233
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36953
diff changeset
   416
         typeof_eqns = add_rule ([], Logic.dest_equals (map_types
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59058
diff changeset
   417
           Type.strip_sorts (Thm.prop_of (Drule.abs_def thm)))) typeof_eqns,
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   418
         types = types,
61094
3d88cd531abe trim context for persistent storage;
wenzelm
parents: 61039
diff changeset
   419
         realizers = realizers, defs = insert Thm.eq_thm_prop (Thm.trim_context thm) defs,
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   420
         expand = expand, prep = prep}
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   421
      else
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   422
        {realizes_eqns = realizes_eqns, typeof_eqns = typeof_eqns, types = types,
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   423
         realizers = realizers, defs = defs,
37233
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36953
diff changeset
   424
         expand = insert (op =) name expand, prep = prep})
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   425
  end;
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   426
33704
6aeb8454efc1 add_expand_thm: explicit indication of is_def instead of fragile heuristic, tuned signature;
wenzelm
parents: 33522
diff changeset
   427
fun extraction_expand is_def =
6aeb8454efc1 add_expand_thm: explicit indication of is_def instead of fragile heuristic, tuned signature;
wenzelm
parents: 33522
diff changeset
   428
  Thm.declaration_attribute (fn th => Context.mapping (add_expand_thm is_def th) I);
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   429
15801
d2f5ca3c048d superceded by Pure.thy and CPure.thy;
wenzelm
parents: 15798
diff changeset
   430
13732
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   431
(** types with computational content **)
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   432
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   433
fun add_types tys thy =
22717
74dbc7696083 canonical merge operations
haftmann
parents: 22675
diff changeset
   434
  ExtractionData.map
74dbc7696083 canonical merge operations
haftmann
parents: 22675
diff changeset
   435
    (fn {realizes_eqns, typeof_eqns, types, realizers, defs, expand, prep} =>
13732
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   436
      {realizes_eqns = realizes_eqns, typeof_eqns = typeof_eqns,
22717
74dbc7696083 canonical merge operations
haftmann
parents: 22675
diff changeset
   437
       types = fold (AList.update (op =) o apfst (Sign.intern_type thy)) tys types,
74dbc7696083 canonical merge operations
haftmann
parents: 22675
diff changeset
   438
       realizers = realizers, defs = defs, expand = expand, prep = prep})
74dbc7696083 canonical merge operations
haftmann
parents: 22675
diff changeset
   439
    thy;
13732
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   440
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   441
15801
d2f5ca3c048d superceded by Pure.thy and CPure.thy;
wenzelm
parents: 15798
diff changeset
   442
(** Pure setup **)
d2f5ca3c048d superceded by Pure.thy and CPure.thy;
wenzelm
parents: 15798
diff changeset
   443
53171
a5e54d4d9081 added Theory.setup convenience;
wenzelm
parents: 52788
diff changeset
   444
val _ = Theory.setup
18708
4b3dadb4fe33 setup: theory -> theory;
wenzelm
parents: 18358
diff changeset
   445
  (add_types [("prop", ([], NONE))] #>
15801
d2f5ca3c048d superceded by Pure.thy and CPure.thy;
wenzelm
parents: 15798
diff changeset
   446
d2f5ca3c048d superceded by Pure.thy and CPure.thy;
wenzelm
parents: 15798
diff changeset
   447
   add_typeof_eqns
67721
5348bea4accd eliminated ASCII syntax from Pure bootstrap;
wenzelm
parents: 67522
diff changeset
   448
     ["(typeof (PROP P)) \<equiv> (Type (TYPE(Null))) \<Longrightarrow>  \
5348bea4accd eliminated ASCII syntax from Pure bootstrap;
wenzelm
parents: 67522
diff changeset
   449
    \  (typeof (PROP Q)) \<equiv> (Type (TYPE('Q))) \<Longrightarrow>  \
5348bea4accd eliminated ASCII syntax from Pure bootstrap;
wenzelm
parents: 67522
diff changeset
   450
    \    (typeof (PROP P \<Longrightarrow> PROP Q)) \<equiv> (Type (TYPE('Q)))",
15801
d2f5ca3c048d superceded by Pure.thy and CPure.thy;
wenzelm
parents: 15798
diff changeset
   451
67721
5348bea4accd eliminated ASCII syntax from Pure bootstrap;
wenzelm
parents: 67522
diff changeset
   452
      "(typeof (PROP Q)) \<equiv> (Type (TYPE(Null))) \<Longrightarrow>  \
5348bea4accd eliminated ASCII syntax from Pure bootstrap;
wenzelm
parents: 67522
diff changeset
   453
    \    (typeof (PROP P \<Longrightarrow> PROP Q)) \<equiv> (Type (TYPE(Null)))",
15801
d2f5ca3c048d superceded by Pure.thy and CPure.thy;
wenzelm
parents: 15798
diff changeset
   454
67721
5348bea4accd eliminated ASCII syntax from Pure bootstrap;
wenzelm
parents: 67522
diff changeset
   455
      "(typeof (PROP P)) \<equiv> (Type (TYPE('P))) \<Longrightarrow>  \
5348bea4accd eliminated ASCII syntax from Pure bootstrap;
wenzelm
parents: 67522
diff changeset
   456
    \  (typeof (PROP Q)) \<equiv> (Type (TYPE('Q))) \<Longrightarrow>  \
5348bea4accd eliminated ASCII syntax from Pure bootstrap;
wenzelm
parents: 67522
diff changeset
   457
    \    (typeof (PROP P \<Longrightarrow> PROP Q)) \<equiv> (Type (TYPE('P \<Rightarrow> 'Q)))",
15801
d2f5ca3c048d superceded by Pure.thy and CPure.thy;
wenzelm
parents: 15798
diff changeset
   458
67721
5348bea4accd eliminated ASCII syntax from Pure bootstrap;
wenzelm
parents: 67522
diff changeset
   459
      "(\<lambda>x. typeof (PROP P (x))) \<equiv> (\<lambda>x. Type (TYPE(Null))) \<Longrightarrow>  \
5348bea4accd eliminated ASCII syntax from Pure bootstrap;
wenzelm
parents: 67522
diff changeset
   460
    \    (typeof (\<And>x. PROP P (x))) \<equiv> (Type (TYPE(Null)))",
15801
d2f5ca3c048d superceded by Pure.thy and CPure.thy;
wenzelm
parents: 15798
diff changeset
   461
67721
5348bea4accd eliminated ASCII syntax from Pure bootstrap;
wenzelm
parents: 67522
diff changeset
   462
      "(\<lambda>x. typeof (PROP P (x))) \<equiv> (\<lambda>x. Type (TYPE('P))) \<Longrightarrow>  \
5348bea4accd eliminated ASCII syntax from Pure bootstrap;
wenzelm
parents: 67522
diff changeset
   463
    \    (typeof (\<And>x::'a. PROP P (x))) \<equiv> (Type (TYPE('a \<Rightarrow> 'P)))",
15801
d2f5ca3c048d superceded by Pure.thy and CPure.thy;
wenzelm
parents: 15798
diff changeset
   464
67721
5348bea4accd eliminated ASCII syntax from Pure bootstrap;
wenzelm
parents: 67522
diff changeset
   465
      "(\<lambda>x. typeof (f (x))) \<equiv> (\<lambda>x. Type (TYPE('f))) \<Longrightarrow>  \
5348bea4accd eliminated ASCII syntax from Pure bootstrap;
wenzelm
parents: 67522
diff changeset
   466
    \    (typeof (f)) \<equiv> (Type (TYPE('f)))"] #>
15801
d2f5ca3c048d superceded by Pure.thy and CPure.thy;
wenzelm
parents: 15798
diff changeset
   467
d2f5ca3c048d superceded by Pure.thy and CPure.thy;
wenzelm
parents: 15798
diff changeset
   468
   add_realizes_eqns
67721
5348bea4accd eliminated ASCII syntax from Pure bootstrap;
wenzelm
parents: 67522
diff changeset
   469
     ["(typeof (PROP P)) \<equiv> (Type (TYPE(Null))) \<Longrightarrow>  \
5348bea4accd eliminated ASCII syntax from Pure bootstrap;
wenzelm
parents: 67522
diff changeset
   470
    \    (realizes (r) (PROP P \<Longrightarrow> PROP Q)) \<equiv>  \
5348bea4accd eliminated ASCII syntax from Pure bootstrap;
wenzelm
parents: 67522
diff changeset
   471
    \    (PROP realizes (Null) (PROP P) \<Longrightarrow> PROP realizes (r) (PROP Q))",
15801
d2f5ca3c048d superceded by Pure.thy and CPure.thy;
wenzelm
parents: 15798
diff changeset
   472
67721
5348bea4accd eliminated ASCII syntax from Pure bootstrap;
wenzelm
parents: 67522
diff changeset
   473
      "(typeof (PROP P)) \<equiv> (Type (TYPE('P))) \<Longrightarrow>  \
5348bea4accd eliminated ASCII syntax from Pure bootstrap;
wenzelm
parents: 67522
diff changeset
   474
    \  (typeof (PROP Q)) \<equiv> (Type (TYPE(Null))) \<Longrightarrow>  \
5348bea4accd eliminated ASCII syntax from Pure bootstrap;
wenzelm
parents: 67522
diff changeset
   475
    \    (realizes (r) (PROP P \<Longrightarrow> PROP Q)) \<equiv>  \
5348bea4accd eliminated ASCII syntax from Pure bootstrap;
wenzelm
parents: 67522
diff changeset
   476
    \    (\<And>x::'P. PROP realizes (x) (PROP P) \<Longrightarrow> PROP realizes (Null) (PROP Q))",
15801
d2f5ca3c048d superceded by Pure.thy and CPure.thy;
wenzelm
parents: 15798
diff changeset
   477
67721
5348bea4accd eliminated ASCII syntax from Pure bootstrap;
wenzelm
parents: 67522
diff changeset
   478
      "(realizes (r) (PROP P \<Longrightarrow> PROP Q)) \<equiv>  \
5348bea4accd eliminated ASCII syntax from Pure bootstrap;
wenzelm
parents: 67522
diff changeset
   479
    \  (\<And>x. PROP realizes (x) (PROP P) \<Longrightarrow> PROP realizes (r (x)) (PROP Q))",
15801
d2f5ca3c048d superceded by Pure.thy and CPure.thy;
wenzelm
parents: 15798
diff changeset
   480
67721
5348bea4accd eliminated ASCII syntax from Pure bootstrap;
wenzelm
parents: 67522
diff changeset
   481
      "(\<lambda>x. typeof (PROP P (x))) \<equiv> (\<lambda>x. Type (TYPE(Null))) \<Longrightarrow>  \
5348bea4accd eliminated ASCII syntax from Pure bootstrap;
wenzelm
parents: 67522
diff changeset
   482
    \    (realizes (r) (\<And>x. PROP P (x))) \<equiv>  \
5348bea4accd eliminated ASCII syntax from Pure bootstrap;
wenzelm
parents: 67522
diff changeset
   483
    \    (\<And>x. PROP realizes (Null) (PROP P (x)))",
15801
d2f5ca3c048d superceded by Pure.thy and CPure.thy;
wenzelm
parents: 15798
diff changeset
   484
67721
5348bea4accd eliminated ASCII syntax from Pure bootstrap;
wenzelm
parents: 67522
diff changeset
   485
      "(realizes (r) (\<And>x. PROP P (x))) \<equiv>  \
5348bea4accd eliminated ASCII syntax from Pure bootstrap;
wenzelm
parents: 67522
diff changeset
   486
    \  (\<And>x. PROP realizes (r (x)) (PROP P (x)))"] #>
15801
d2f5ca3c048d superceded by Pure.thy and CPure.thy;
wenzelm
parents: 15798
diff changeset
   487
67147
dea94b1aabc3 prefer control symbol antiquotations;
wenzelm
parents: 66251
diff changeset
   488
   Attrib.setup \<^binding>\<open>extraction_expand\<close> (Scan.succeed (extraction_expand false))
33704
6aeb8454efc1 add_expand_thm: explicit indication of is_def instead of fragile heuristic, tuned signature;
wenzelm
parents: 33522
diff changeset
   489
     "specify theorems to be expanded during extraction" #>
67147
dea94b1aabc3 prefer control symbol antiquotations;
wenzelm
parents: 66251
diff changeset
   490
   Attrib.setup \<^binding>\<open>extraction_expand_def\<close> (Scan.succeed (extraction_expand true))
53171
a5e54d4d9081 added Theory.setup convenience;
wenzelm
parents: 52788
diff changeset
   491
     "specify definitions to be expanded during extraction");
15801
d2f5ca3c048d superceded by Pure.thy and CPure.thy;
wenzelm
parents: 15798
diff changeset
   492
d2f5ca3c048d superceded by Pure.thy and CPure.thy;
wenzelm
parents: 15798
diff changeset
   493
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   494
(**** extract program ****)
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   495
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   496
val dummyt = Const ("dummy", dummyT);
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   497
58436
haftmann
parents: 56436
diff changeset
   498
fun extract thm_vss thy =
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   499
  let
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16363
diff changeset
   500
    val thy' = add_syntax thy;
62922
96691631c1eb clarified context;
wenzelm
parents: 61865
diff changeset
   501
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   502
    val {realizes_eqns, typeof_eqns, types, realizers, defs, expand, prep} =
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   503
      ExtractionData.get thy;
22717
74dbc7696083 canonical merge operations
haftmann
parents: 22675
diff changeset
   504
    val procs = maps (rev o fst o snd) types;
13732
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   505
    val rtypes = map fst types;
37233
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36953
diff changeset
   506
    val typroc = typeof_proc [];
70915
bd4d37edfee4 clarified expand_proof/expand_name: allow more detailed control via thm_header;
wenzelm
parents: 70840
diff changeset
   507
    fun expand_name ({name, ...}: Proofterm.thm_header) =
bd4d37edfee4 clarified expand_proof/expand_name: allow more detailed control via thm_header;
wenzelm
parents: 70840
diff changeset
   508
      if name = "" orelse member (op =) expand name then SOME "" else NONE;
61094
3d88cd531abe trim context for persistent storage;
wenzelm
parents: 61039
diff changeset
   509
    val prep = the_default (K I) prep thy' o
70840
5b80eb4fd0f3 proper ML names;
wenzelm
parents: 70836
diff changeset
   510
      Proof_Rewrite_Rules.elim_defs thy' false (map (Thm.transfer thy) defs) o
70915
bd4d37edfee4 clarified expand_proof/expand_name: allow more detailed control via thm_header;
wenzelm
parents: 70840
diff changeset
   511
      Proofterm.expand_proof thy' expand_name;
16800
90eff1b52428 improved Net interface;
wenzelm
parents: 16790
diff changeset
   512
    val rrews = Net.merge (K false) (#net realizes_eqns, #net typeof_eqns);
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   513
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   514
    fun find_inst prop Ts ts vs =
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   515
      let
13732
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   516
        val rvs = relevant_vars rtypes prop;
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   517
        val vars = vars_of prop;
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   518
        val n = Int.min (length vars, length ts);
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   519
33337
9c3b9bf81e8b eliminated some old folds;
wenzelm
parents: 33317
diff changeset
   520
        fun add_args (Var ((a, i), _), t) (vs', tye) =
20664
ffbc5a57191a member (op =);
wenzelm
parents: 20548
diff changeset
   521
          if member (op =) rvs a then
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16363
diff changeset
   522
            let val T = etype_of thy' vs Ts t
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   523
            in if T = nullT then (vs', tye)
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   524
               else (a :: vs', (("'" ^ a, i), T) :: tye)
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   525
            end
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   526
          else (vs', tye)
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   527
33957
e9afca2118d4 normalized uncurry take/drop
haftmann
parents: 33955
diff changeset
   528
      in fold_rev add_args (take n vars ~~ take n ts) ([], []) end;
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   529
37233
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36953
diff changeset
   530
    fun mk_shyps tye = maps (fn (ixn, _) =>
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36953
diff changeset
   531
      Logic.mk_of_sort (TVar (ixn, []), Sign.defaultS thy)) tye;
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36953
diff changeset
   532
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36953
diff changeset
   533
    fun mk_sprfs cs tye = maps (fn (_, T) =>
70840
5b80eb4fd0f3 proper ML names;
wenzelm
parents: 70836
diff changeset
   534
      Proof_Rewrite_Rules.expand_of_sort_proof thy' (map SOME cs)
37233
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36953
diff changeset
   535
        (T, Sign.defaultS thy)) tye;
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36953
diff changeset
   536
33038
8f9594c31de4 dropped redundant gen_ prefix
haftmann
parents: 32784
diff changeset
   537
    fun find (vs: string list) = Option.map snd o find_first (curry (eq_set (op =)) vs o fst);
28375
c879d88d038a eliminated polymorphic equality;
wenzelm
parents: 28370
diff changeset
   538
    fun find' (s: string) = map_filter (fn (s', x) => if s = s' then SOME x else NONE);
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   539
46219
426ed18eba43 discontinued old-style Term.list_abs in favour of plain Term.abs;
wenzelm
parents: 44057
diff changeset
   540
    fun app_rlz_rews Ts vs t =
426ed18eba43 discontinued old-style Term.list_abs in favour of plain Term.abs;
wenzelm
parents: 44057
diff changeset
   541
      strip_abs (length Ts)
426ed18eba43 discontinued old-style Term.list_abs in favour of plain Term.abs;
wenzelm
parents: 44057
diff changeset
   542
        (freeze_thaw (condrew thy' rrews (procs @ [typroc vs, rlz_proc]))
426ed18eba43 discontinued old-style Term.list_abs in favour of plain Term.abs;
wenzelm
parents: 44057
diff changeset
   543
          (fold (Term.abs o pair "x") Ts t));
13732
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   544
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   545
    fun realizes_null vs prop = app_rlz_rews [] vs
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   546
      (Const ("realizes", nullT --> propT --> propT) $ nullt $ prop);
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   547
49960
1167c1157a5b more conventional argument order;
haftmann
parents: 48704
diff changeset
   548
    fun corr d vs ts Ts hs cs _ (PBound i) _ defs = (PBound i, defs)
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   549
49960
1167c1157a5b more conventional argument order;
haftmann
parents: 48704
diff changeset
   550
      | corr d vs ts Ts hs cs t (Abst (s, SOME T, prf)) (Abst (_, _, prf')) defs =
1167c1157a5b more conventional argument order;
haftmann
parents: 48704
diff changeset
   551
          let val (corr_prf, defs') = corr d vs [] (T :: Ts)
1167c1157a5b more conventional argument order;
haftmann
parents: 48704
diff changeset
   552
            (dummyt :: hs) cs (case t of SOME (Abs (_, _, u)) => SOME u | _ => NONE)
1167c1157a5b more conventional argument order;
haftmann
parents: 48704
diff changeset
   553
            prf (Proofterm.incr_pboundvars 1 0 prf') defs
1167c1157a5b more conventional argument order;
haftmann
parents: 48704
diff changeset
   554
          in (Abst (s, SOME T, corr_prf), defs') end
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   555
49960
1167c1157a5b more conventional argument order;
haftmann
parents: 48704
diff changeset
   556
      | corr d vs ts Ts hs cs t (AbsP (s, SOME prop, prf)) (AbsP (_, _, prf')) defs =
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   557
          let
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16363
diff changeset
   558
            val T = etype_of thy' vs Ts prop;
42407
5b9dd52f5dca prefer internal types, via Simple_Syntax.read_typ;
wenzelm
parents: 42406
diff changeset
   559
            val u = if T = nullT then
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   560
                (case t of SOME u => SOME (incr_boundvars 1 u) | NONE => NONE)
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   561
              else (case t of SOME (Abs (_, _, u)) => SOME u | _ => NONE);
49960
1167c1157a5b more conventional argument order;
haftmann
parents: 48704
diff changeset
   562
            val (corr_prf, defs') =
1167c1157a5b more conventional argument order;
haftmann
parents: 48704
diff changeset
   563
              corr d vs [] (T :: Ts) (prop :: hs)
1167c1157a5b more conventional argument order;
haftmann
parents: 48704
diff changeset
   564
                (prop :: cs) u (Proofterm.incr_pboundvars 0 1 prf)
1167c1157a5b more conventional argument order;
haftmann
parents: 48704
diff changeset
   565
                (Proofterm.incr_pboundvars 0 1 prf') defs;
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   566
            val rlz = Const ("realizes", T --> propT --> propT)
49960
1167c1157a5b more conventional argument order;
haftmann
parents: 48704
diff changeset
   567
          in (
13732
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   568
            if T = nullT then AbsP ("R",
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   569
              SOME (app_rlz_rews Ts vs (rlz $ nullt $ prop)),
37310
96e2b9a6f074 do not open Proofterm, which is very ould style;
wenzelm
parents: 37237
diff changeset
   570
                Proofterm.prf_subst_bounds [nullt] corr_prf)
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   571
            else Abst (s, SOME T, AbsP ("R",
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   572
              SOME (app_rlz_rews (T :: Ts) vs
49960
1167c1157a5b more conventional argument order;
haftmann
parents: 48704
diff changeset
   573
                (rlz $ Bound 0 $ incr_boundvars 1 prop)), corr_prf)), defs')
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   574
          end
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   575
49960
1167c1157a5b more conventional argument order;
haftmann
parents: 48704
diff changeset
   576
      | corr d vs ts Ts hs cs t' (prf % SOME t) (prf' % _) defs =
13732
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   577
          let
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   578
            val (Us, T) = strip_type (fastype_of1 (Ts, t));
49960
1167c1157a5b more conventional argument order;
haftmann
parents: 48704
diff changeset
   579
            val (corr_prf, defs') = corr d vs (t :: ts) Ts hs cs
20664
ffbc5a57191a member (op =);
wenzelm
parents: 20548
diff changeset
   580
              (if member (op =) rtypes (tname_of T) then t'
49960
1167c1157a5b more conventional argument order;
haftmann
parents: 48704
diff changeset
   581
               else (case t' of SOME (u $ _) => SOME u | _ => NONE))
1167c1157a5b more conventional argument order;
haftmann
parents: 48704
diff changeset
   582
               prf prf' defs;
20664
ffbc5a57191a member (op =);
wenzelm
parents: 20548
diff changeset
   583
            val u = if not (member (op =) rtypes (tname_of T)) then t else
13732
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   584
              let
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16363
diff changeset
   585
                val eT = etype_of thy' vs Ts t;
13732
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   586
                val (r, Us') = if eT = nullT then (nullt, Us) else
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   587
                  (Bound (length Us), eT :: Us);
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   588
                val u = list_comb (incr_boundvars (length Us') t,
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   589
                  map Bound (length Us - 1 downto 0));
17271
2756a73f63a5 introduced some new-style AList operations
haftmann
parents: 17232
diff changeset
   590
                val u' = (case AList.lookup (op =) types (tname_of T) of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   591
                    SOME ((_, SOME f)) => f r eT u T
13732
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   592
                  | _ => Const ("realizes", eT --> T --> T) $ r $ u)
46219
426ed18eba43 discontinued old-style Term.list_abs in favour of plain Term.abs;
wenzelm
parents: 44057
diff changeset
   593
              in app_rlz_rews Ts vs (fold_rev (Term.abs o pair "x") Us' u') end
49960
1167c1157a5b more conventional argument order;
haftmann
parents: 48704
diff changeset
   594
          in (corr_prf % SOME u, defs') end
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   595
49960
1167c1157a5b more conventional argument order;
haftmann
parents: 48704
diff changeset
   596
      | corr d vs ts Ts hs cs t (prf1 %% prf2) (prf1' %% prf2') defs =
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   597
          let
70447
755d58b48cec clarified modules: provide reconstruct_proof / expand_proof at the bottom of proof term construction;
wenzelm
parents: 70436
diff changeset
   598
            val prop = Proofterm.prop_of' hs prf2';
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16363
diff changeset
   599
            val T = etype_of thy' vs Ts prop;
49960
1167c1157a5b more conventional argument order;
haftmann
parents: 48704
diff changeset
   600
            val (f, u, defs1) = if T = nullT then (t, NONE, defs) else
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   601
              (case t of
49960
1167c1157a5b more conventional argument order;
haftmann
parents: 48704
diff changeset
   602
                 SOME (f $ u) => (SOME f, SOME u, defs)
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   603
               | _ =>
49960
1167c1157a5b more conventional argument order;
haftmann
parents: 48704
diff changeset
   604
                 let val (u, defs1) = extr d vs [] Ts hs prf2' defs
1167c1157a5b more conventional argument order;
haftmann
parents: 48704
diff changeset
   605
                 in (NONE, SOME u, defs1) end)
1167c1157a5b more conventional argument order;
haftmann
parents: 48704
diff changeset
   606
            val ((corr_prf1, corr_prf2), defs2) =
1167c1157a5b more conventional argument order;
haftmann
parents: 48704
diff changeset
   607
              defs1
1167c1157a5b more conventional argument order;
haftmann
parents: 48704
diff changeset
   608
              |> corr d vs [] Ts hs cs f prf1 prf1'
1167c1157a5b more conventional argument order;
haftmann
parents: 48704
diff changeset
   609
              ||>> corr d vs [] Ts hs cs u prf2 prf2';
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   610
          in
49960
1167c1157a5b more conventional argument order;
haftmann
parents: 48704
diff changeset
   611
            if T = nullT then (corr_prf1 %% corr_prf2, defs2) else
1167c1157a5b more conventional argument order;
haftmann
parents: 48704
diff changeset
   612
              (corr_prf1 % u %% corr_prf2, defs2)
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   613
          end
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   614
70494
41108e3e9ca5 formal position for PThm nodes;
wenzelm
parents: 70493
diff changeset
   615
      | corr d vs ts Ts hs cs _ (prf0 as PThm (thm_header as {types = SOME Ts', ...}, thm_body)) _ defs =
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   616
          let
70538
fc9ba6fe367f clarified PThm: theory_name simplifies retrieval from exports;
wenzelm
parents: 70501
diff changeset
   617
            val {pos, theory_name, name, prop, ...} = thm_header;
70493
a9053fa30909 clarified ML types;
wenzelm
parents: 70449
diff changeset
   618
            val prf = Proofterm.thm_body_proof_open thm_body;
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   619
            val (vs', tye) = find_inst prop Ts ts vs;
37233
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36953
diff changeset
   620
            val shyps = mk_shyps tye;
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36953
diff changeset
   621
            val sprfs = mk_sprfs cs tye;
36042
85efdadee8ae switched PThm/PAxm etc. to use canonical order of type variables (term variables unchanged)
krauss
parents: 35985
diff changeset
   622
            val tye' = (map fst (Term.add_tvars prop [] |> rev) ~~ Ts') @ tye;
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16363
diff changeset
   623
            val T = etype_of thy' vs' [] prop;
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   624
            val defs' = if T = nullT then defs
49960
1167c1157a5b more conventional argument order;
haftmann
parents: 48704
diff changeset
   625
              else snd (extr d vs ts Ts hs prf0 defs)
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   626
          in
49960
1167c1157a5b more conventional argument order;
haftmann
parents: 48704
diff changeset
   627
            if T = nullT andalso realizes_null vs' prop aconv prop then (prf0, defs)
1167c1157a5b more conventional argument order;
haftmann
parents: 48704
diff changeset
   628
            else (case Symtab.lookup realizers name of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   629
              NONE => (case find vs' (find' name defs') of
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   630
                NONE =>
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   631
                  let
21858
05f57309170c avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
wenzelm
parents: 21646
diff changeset
   632
                    val _ = T = nullT orelse error "corr: internal error";
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   633
                    val _ = msg d ("Building correctness proof for " ^ quote name ^
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   634
                      (if null vs' then ""
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   635
                       else " (relevant variables: " ^ commas_quote vs' ^ ")"));
70449
6e34025981be clarified global theory context;
wenzelm
parents: 70447
diff changeset
   636
                    val prf' = prep (Proofterm.reconstruct_proof thy' prop prf);
49960
1167c1157a5b more conventional argument order;
haftmann
parents: 48704
diff changeset
   637
                    val (corr_prf0, defs'') = corr (d + 1) vs' [] [] []
1167c1157a5b more conventional argument order;
haftmann
parents: 48704
diff changeset
   638
                      (rev shyps) NONE prf' prf' defs';
37233
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36953
diff changeset
   639
                    val corr_prf = mkabsp shyps corr_prf0;
70447
755d58b48cec clarified modules: provide reconstruct_proof / expand_proof at the bottom of proof term construction;
wenzelm
parents: 70436
diff changeset
   640
                    val corr_prop = Proofterm.prop_of corr_prf;
70493
a9053fa30909 clarified ML types;
wenzelm
parents: 70449
diff changeset
   641
                    val corr_header =
70577
ed651a58afe4 back to uniform serial (reverting 913b4afb6ac2): this allows to treat derivation id like name space entity id;
wenzelm
parents: 70538
diff changeset
   642
                      Proofterm.thm_header (serial ()) pos theory_name
70538
fc9ba6fe367f clarified PThm: theory_name simplifies retrieval from exports;
wenzelm
parents: 70501
diff changeset
   643
                        (corr_name name vs') corr_prop
fc9ba6fe367f clarified PThm: theory_name simplifies retrieval from exports;
wenzelm
parents: 70501
diff changeset
   644
                        (SOME (map TVar (Term.add_tvars corr_prop [] |> rev)));
37233
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36953
diff changeset
   645
                    val corr_prf' =
70493
a9053fa30909 clarified ML types;
wenzelm
parents: 70449
diff changeset
   646
                      Proofterm.proof_combP
a9053fa30909 clarified ML types;
wenzelm
parents: 70449
diff changeset
   647
                        (Proofterm.proof_combt
70589
00b67aaa4010 clarified modules;
wenzelm
parents: 70577
diff changeset
   648
                          (PThm (corr_header, make_proof_body corr_prf), vfs_of corr_prop),
70493
a9053fa30909 clarified ML types;
wenzelm
parents: 70449
diff changeset
   649
                            map PBound (length shyps - 1 downto 0)) |>
37310
96e2b9a6f074 do not open Proofterm, which is very ould style;
wenzelm
parents: 37237
diff changeset
   650
                      fold_rev Proofterm.forall_intr_proof'
96e2b9a6f074 do not open Proofterm, which is very ould style;
wenzelm
parents: 37237
diff changeset
   651
                        (map (get_var_type corr_prop) (vfs_of prop)) |>
37233
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36953
diff changeset
   652
                      mkabsp shyps
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   653
                  in
49960
1167c1157a5b more conventional argument order;
haftmann
parents: 48704
diff changeset
   654
                    (Proofterm.proof_combP (prf_subst_TVars tye' corr_prf', sprfs),
1167c1157a5b more conventional argument order;
haftmann
parents: 48704
diff changeset
   655
                      (name, (vs', ((nullt, nullt), (corr_prf, corr_prf')))) :: defs'')
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   656
                  end
37310
96e2b9a6f074 do not open Proofterm, which is very ould style;
wenzelm
parents: 37237
diff changeset
   657
              | SOME (_, (_, prf')) =>
49960
1167c1157a5b more conventional argument order;
haftmann
parents: 48704
diff changeset
   658
                  (Proofterm.proof_combP (prf_subst_TVars tye' prf', sprfs), defs'))
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   659
            | SOME rs => (case find vs' rs of
49960
1167c1157a5b more conventional argument order;
haftmann
parents: 48704
diff changeset
   660
                SOME (_, prf') => (Proofterm.proof_combP (prf_subst_TVars tye' prf', sprfs), defs')
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   661
              | NONE => error ("corr: no realizer for instance of theorem " ^
70449
6e34025981be clarified global theory context;
wenzelm
parents: 70447
diff changeset
   662
                  quote name ^ ":\n" ^ Syntax.string_of_term_global thy' (Envir.beta_norm
70447
755d58b48cec clarified modules: provide reconstruct_proof / expand_proof at the bottom of proof term construction;
wenzelm
parents: 70436
diff changeset
   663
                    (Proofterm.prop_of (Proofterm.proof_combt (prf0, ts)))))))
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   664
          end
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   665
49960
1167c1157a5b more conventional argument order;
haftmann
parents: 48704
diff changeset
   666
      | corr d vs ts Ts hs cs _ (prf0 as PAxm (s, prop, SOME Ts')) _ defs =
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   667
          let
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   668
            val (vs', tye) = find_inst prop Ts ts vs;
36042
85efdadee8ae switched PThm/PAxm etc. to use canonical order of type variables (term variables unchanged)
krauss
parents: 35985
diff changeset
   669
            val tye' = (map fst (Term.add_tvars prop [] |> rev) ~~ Ts') @ tye
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   670
          in
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16363
diff changeset
   671
            if etype_of thy' vs' [] prop = nullT andalso
49960
1167c1157a5b more conventional argument order;
haftmann
parents: 48704
diff changeset
   672
              realizes_null vs' prop aconv prop then (prf0, defs)
18956
c050ae1f8f52 Envir.(beta_)eta_contract;
wenzelm
parents: 18928
diff changeset
   673
            else case find vs' (Symtab.lookup_list realizers s) of
49960
1167c1157a5b more conventional argument order;
haftmann
parents: 48704
diff changeset
   674
              SOME (_, prf) => (Proofterm.proof_combP (prf_subst_TVars tye' prf, mk_sprfs cs tye),
1167c1157a5b more conventional argument order;
haftmann
parents: 48704
diff changeset
   675
                defs)
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   676
            | NONE => error ("corr: no realizer for instance of axiom " ^
70449
6e34025981be clarified global theory context;
wenzelm
parents: 70447
diff changeset
   677
                quote s ^ ":\n" ^ Syntax.string_of_term_global thy' (Envir.beta_norm
70447
755d58b48cec clarified modules: provide reconstruct_proof / expand_proof at the bottom of proof term construction;
wenzelm
parents: 70436
diff changeset
   678
                  (Proofterm.prop_of (Proofterm.proof_combt (prf0, ts)))))
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   679
          end
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   680
49960
1167c1157a5b more conventional argument order;
haftmann
parents: 48704
diff changeset
   681
      | corr d vs ts Ts hs _ _ _ _ defs = error "corr: bad proof"
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   682
49960
1167c1157a5b more conventional argument order;
haftmann
parents: 48704
diff changeset
   683
    and extr d vs ts Ts hs (PBound i) defs = (Bound i, defs)
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   684
49960
1167c1157a5b more conventional argument order;
haftmann
parents: 48704
diff changeset
   685
      | extr d vs ts Ts hs (Abst (s, SOME T, prf)) defs =
1167c1157a5b more conventional argument order;
haftmann
parents: 48704
diff changeset
   686
          let val (t, defs') = extr d vs []
1167c1157a5b more conventional argument order;
haftmann
parents: 48704
diff changeset
   687
            (T :: Ts) (dummyt :: hs) (Proofterm.incr_pboundvars 1 0 prf) defs
1167c1157a5b more conventional argument order;
haftmann
parents: 48704
diff changeset
   688
          in (Abs (s, T, t), defs') end
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   689
49960
1167c1157a5b more conventional argument order;
haftmann
parents: 48704
diff changeset
   690
      | extr d vs ts Ts hs (AbsP (s, SOME t, prf)) defs =
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   691
          let
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16363
diff changeset
   692
            val T = etype_of thy' vs Ts t;
49960
1167c1157a5b more conventional argument order;
haftmann
parents: 48704
diff changeset
   693
            val (t, defs') =
1167c1157a5b more conventional argument order;
haftmann
parents: 48704
diff changeset
   694
              extr d vs [] (T :: Ts) (t :: hs) (Proofterm.incr_pboundvars 0 1 prf) defs
1167c1157a5b more conventional argument order;
haftmann
parents: 48704
diff changeset
   695
          in
1167c1157a5b more conventional argument order;
haftmann
parents: 48704
diff changeset
   696
            (if T = nullT then subst_bound (nullt, t) else Abs (s, T, t), defs')
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   697
          end
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   698
49960
1167c1157a5b more conventional argument order;
haftmann
parents: 48704
diff changeset
   699
      | extr d vs ts Ts hs (prf % SOME t) defs =
1167c1157a5b more conventional argument order;
haftmann
parents: 48704
diff changeset
   700
          let val (u, defs') = extr d vs (t :: ts) Ts hs prf defs
1167c1157a5b more conventional argument order;
haftmann
parents: 48704
diff changeset
   701
          in (if member (op =) rtypes (tname_of (body_type (fastype_of1 (Ts, t)))) then u
1167c1157a5b more conventional argument order;
haftmann
parents: 48704
diff changeset
   702
            else u $ t, defs')
13732
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   703
          end
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   704
49960
1167c1157a5b more conventional argument order;
haftmann
parents: 48704
diff changeset
   705
      | extr d vs ts Ts hs (prf1 %% prf2) defs =
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   706
          let
49960
1167c1157a5b more conventional argument order;
haftmann
parents: 48704
diff changeset
   707
            val (f, defs') = extr d vs [] Ts hs prf1 defs;
70447
755d58b48cec clarified modules: provide reconstruct_proof / expand_proof at the bottom of proof term construction;
wenzelm
parents: 70436
diff changeset
   708
            val prop = Proofterm.prop_of' hs prf2;
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16363
diff changeset
   709
            val T = etype_of thy' vs Ts prop
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   710
          in
49960
1167c1157a5b more conventional argument order;
haftmann
parents: 48704
diff changeset
   711
            if T = nullT then (f, defs') else
1167c1157a5b more conventional argument order;
haftmann
parents: 48704
diff changeset
   712
              let val (t, defs'') = extr d vs [] Ts hs prf2 defs'
1167c1157a5b more conventional argument order;
haftmann
parents: 48704
diff changeset
   713
              in (f $ t, defs'') end
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   714
          end
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   715
70494
41108e3e9ca5 formal position for PThm nodes;
wenzelm
parents: 70493
diff changeset
   716
      | extr d vs ts Ts hs (prf0 as PThm (thm_header as {types = SOME Ts', ...}, thm_body)) defs =
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   717
          let
70538
fc9ba6fe367f clarified PThm: theory_name simplifies retrieval from exports;
wenzelm
parents: 70501
diff changeset
   718
            val {pos, theory_name, name = s, prop, ...} = thm_header;
70493
a9053fa30909 clarified ML types;
wenzelm
parents: 70449
diff changeset
   719
            val prf = Proofterm.thm_body_proof_open thm_body;
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   720
            val (vs', tye) = find_inst prop Ts ts vs;
37233
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36953
diff changeset
   721
            val shyps = mk_shyps tye;
36042
85efdadee8ae switched PThm/PAxm etc. to use canonical order of type variables (term variables unchanged)
krauss
parents: 35985
diff changeset
   722
            val tye' = (map fst (Term.add_tvars prop [] |> rev) ~~ Ts') @ tye
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   723
          in
17412
e26cb20ef0cc TableFun/Symtab: curried lookup and update;
wenzelm
parents: 17271
diff changeset
   724
            case Symtab.lookup realizers s of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   725
              NONE => (case find vs' (find' s defs) of
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   726
                NONE =>
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   727
                  let
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   728
                    val _ = msg d ("Extracting " ^ quote s ^
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   729
                      (if null vs' then ""
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   730
                       else " (relevant variables: " ^ commas_quote vs' ^ ")"));
70449
6e34025981be clarified global theory context;
wenzelm
parents: 70447
diff changeset
   731
                    val prf' = prep (Proofterm.reconstruct_proof thy' prop prf);
49960
1167c1157a5b more conventional argument order;
haftmann
parents: 48704
diff changeset
   732
                    val (t, defs') = extr (d + 1) vs' [] [] [] prf' defs;
1167c1157a5b more conventional argument order;
haftmann
parents: 48704
diff changeset
   733
                    val (corr_prf, defs'') = corr (d + 1) vs' [] [] []
1167c1157a5b more conventional argument order;
haftmann
parents: 48704
diff changeset
   734
                      (rev shyps) (SOME t) prf' prf' defs';
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   735
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   736
                    val nt = Envir.beta_norm t;
20664
ffbc5a57191a member (op =);
wenzelm
parents: 20548
diff changeset
   737
                    val args = filter_out (fn v => member (op =) rtypes
ffbc5a57191a member (op =);
wenzelm
parents: 20548
diff changeset
   738
                      (tname_of (body_type (fastype_of v)))) (vfs_of prop);
33317
b4534348b8fd standardized filter/filter_out;
wenzelm
parents: 33245
diff changeset
   739
                    val args' = filter (fn v => Logic.occs (v, nt)) args;
37233
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36953
diff changeset
   740
                    val t' = mkabs args' nt;
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   741
                    val T = fastype_of t';
13732
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   742
                    val cname = extr_name s vs';
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   743
                    val c = Const (cname, T);
37233
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36953
diff changeset
   744
                    val u = mkabs args (list_comb (c, args'));
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   745
                    val eqn = Logic.mk_equals (c, t');
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   746
                    val rlz =
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   747
                      Const ("realizes", fastype_of nt --> propT --> propT);
13732
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   748
                    val lhs = app_rlz_rews [] vs' (rlz $ nt $ prop);
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   749
                    val rhs = app_rlz_rews [] vs' (rlz $ list_comb (c, args') $ prop);
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   750
                    val f = app_rlz_rews [] vs'
f8badfef5cf2 Correctness proofs are now modular, too.
berghofe
parents: 13719
diff changeset
   751
                      (Abs ("x", T, rlz $ list_comb (Bound 0, args') $ prop));
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   752
37233
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36953
diff changeset
   753
                    val corr_prf' = mkabsp shyps
70806
wenzelm
parents: 70589
diff changeset
   754
                      (change_types [] Proofterm.equal_elim_axm %> lhs %> rhs %%
wenzelm
parents: 70589
diff changeset
   755
                       (change_types [propT] Proofterm.symmetric_axm %> rhs %> lhs %%
wenzelm
parents: 70589
diff changeset
   756
                         (change_types [T, propT] Proofterm.combination_axm %> f %> f %> c %> t' %%
wenzelm
parents: 70589
diff changeset
   757
                           (change_types [T --> propT] Proofterm.reflexive_axm %> f) %%
46909
3c73a121a387 more explicit indication of def names;
wenzelm
parents: 46219
diff changeset
   758
                           PAxm (Thm.def_name cname, eqn,
37233
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36953
diff changeset
   759
                             SOME (map TVar (Term.add_tvars eqn [] |> rev))))) %% corr_prf);
70447
755d58b48cec clarified modules: provide reconstruct_proof / expand_proof at the bottom of proof term construction;
wenzelm
parents: 70436
diff changeset
   760
                    val corr_prop = Proofterm.prop_of corr_prf';
70493
a9053fa30909 clarified ML types;
wenzelm
parents: 70449
diff changeset
   761
                    val corr_header =
70577
ed651a58afe4 back to uniform serial (reverting 913b4afb6ac2): this allows to treat derivation id like name space entity id;
wenzelm
parents: 70538
diff changeset
   762
                      Proofterm.thm_header (serial ()) pos theory_name
70538
fc9ba6fe367f clarified PThm: theory_name simplifies retrieval from exports;
wenzelm
parents: 70501
diff changeset
   763
                        (corr_name s vs') corr_prop
fc9ba6fe367f clarified PThm: theory_name simplifies retrieval from exports;
wenzelm
parents: 70501
diff changeset
   764
                        (SOME (map TVar (Term.add_tvars corr_prop [] |> rev)));
37233
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36953
diff changeset
   765
                    val corr_prf'' =
37310
96e2b9a6f074 do not open Proofterm, which is very ould style;
wenzelm
parents: 37237
diff changeset
   766
                      Proofterm.proof_combP (Proofterm.proof_combt
70589
00b67aaa4010 clarified modules;
wenzelm
parents: 70577
diff changeset
   767
                        (PThm (corr_header, make_proof_body corr_prf), vfs_of corr_prop),
37233
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36953
diff changeset
   768
                             map PBound (length shyps - 1 downto 0)) |>
37310
96e2b9a6f074 do not open Proofterm, which is very ould style;
wenzelm
parents: 37237
diff changeset
   769
                      fold_rev Proofterm.forall_intr_proof'
96e2b9a6f074 do not open Proofterm, which is very ould style;
wenzelm
parents: 37237
diff changeset
   770
                        (map (get_var_type corr_prop) (vfs_of prop)) |>
37233
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36953
diff changeset
   771
                      mkabsp shyps
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   772
                  in
49960
1167c1157a5b more conventional argument order;
haftmann
parents: 48704
diff changeset
   773
                    (subst_TVars tye' u,
1167c1157a5b more conventional argument order;
haftmann
parents: 48704
diff changeset
   774
                      (s, (vs', ((t', u), (corr_prf', corr_prf'')))) :: defs'')
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   775
                  end
49960
1167c1157a5b more conventional argument order;
haftmann
parents: 48704
diff changeset
   776
              | SOME ((_, u), _) => (subst_TVars tye' u, defs))
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   777
            | SOME rs => (case find vs' rs of
49960
1167c1157a5b more conventional argument order;
haftmann
parents: 48704
diff changeset
   778
                SOME (t, _) => (subst_TVars tye' t, defs)
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   779
              | NONE => error ("extr: no realizer for instance of theorem " ^
70449
6e34025981be clarified global theory context;
wenzelm
parents: 70447
diff changeset
   780
                  quote s ^ ":\n" ^ Syntax.string_of_term_global thy' (Envir.beta_norm
70447
755d58b48cec clarified modules: provide reconstruct_proof / expand_proof at the bottom of proof term construction;
wenzelm
parents: 70436
diff changeset
   781
                    (Proofterm.prop_of (Proofterm.proof_combt (prf0, ts))))))
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   782
          end
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   783
49960
1167c1157a5b more conventional argument order;
haftmann
parents: 48704
diff changeset
   784
      | extr d vs ts Ts hs (prf0 as PAxm (s, prop, SOME Ts')) defs =
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   785
          let
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   786
            val (vs', tye) = find_inst prop Ts ts vs;
36042
85efdadee8ae switched PThm/PAxm etc. to use canonical order of type variables (term variables unchanged)
krauss
parents: 35985
diff changeset
   787
            val tye' = (map fst (Term.add_tvars prop [] |> rev) ~~ Ts') @ tye
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   788
          in
18956
c050ae1f8f52 Envir.(beta_)eta_contract;
wenzelm
parents: 18928
diff changeset
   789
            case find vs' (Symtab.lookup_list realizers s) of
49960
1167c1157a5b more conventional argument order;
haftmann
parents: 48704
diff changeset
   790
              SOME (t, _) => (subst_TVars tye' t, defs)
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   791
            | NONE => error ("extr: no realizer for instance of axiom " ^
70449
6e34025981be clarified global theory context;
wenzelm
parents: 70447
diff changeset
   792
                quote s ^ ":\n" ^ Syntax.string_of_term_global thy' (Envir.beta_norm
70447
755d58b48cec clarified modules: provide reconstruct_proof / expand_proof at the bottom of proof term construction;
wenzelm
parents: 70436
diff changeset
   793
                  (Proofterm.prop_of (Proofterm.proof_combt (prf0, ts)))))
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   794
          end
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   795
49960
1167c1157a5b more conventional argument order;
haftmann
parents: 48704
diff changeset
   796
      | extr d vs ts Ts hs _ defs = error "extr: bad proof";
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   797
60826
695cf1fab6cc clarified context;
wenzelm
parents: 59787
diff changeset
   798
    fun prep_thm vs raw_thm =
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   799
      let
60826
695cf1fab6cc clarified context;
wenzelm
parents: 59787
diff changeset
   800
        val thm = Thm.transfer thy raw_thm;
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26481
diff changeset
   801
        val prop = Thm.prop_of thm;
28814
463c9e9111ae clarified Thm.proof_body_of vs. Thm.proof_of;
wenzelm
parents: 28812
diff changeset
   802
        val prf = Thm.proof_of thm;
36744
6e1f3d609a68 renamed Thm.get_name -> Thm.derivation_name and Thm.put_name -> Thm.name_derivation, to emphasize the true nature of these operations;
wenzelm
parents: 36620
diff changeset
   803
        val name = Thm.derivation_name thm;
21858
05f57309170c avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
wenzelm
parents: 21646
diff changeset
   804
        val _ = name <> "" orelse error "extraction: unnamed theorem";
05f57309170c avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
wenzelm
parents: 21646
diff changeset
   805
        val _ = etype_of thy' vs [] prop <> nullT orelse error ("theorem " ^
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   806
          quote name ^ " has no computational content")
70449
6e34025981be clarified global theory context;
wenzelm
parents: 70447
diff changeset
   807
      in Proofterm.reconstruct_proof thy' prop prf end;
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   808
33245
65232054ffd0 eliminated some old folds;
wenzelm
parents: 33038
diff changeset
   809
    val defs =
58436
haftmann
parents: 56436
diff changeset
   810
      fold (fn (thm, vs) => snd o (extr 0 vs [] [] [] o prep_thm vs) thm)
haftmann
parents: 56436
diff changeset
   811
        thm_vss [];
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   812
66145
haftmann
parents: 64556
diff changeset
   813
    fun add_def (s, (vs, ((t, u)))) thy =
haftmann
parents: 64556
diff changeset
   814
      let
haftmann
parents: 64556
diff changeset
   815
        val ft = Type.legacy_freeze t;
haftmann
parents: 64556
diff changeset
   816
        val fu = Type.legacy_freeze u;
66146
fd3e128b174f register equations stemming from extracted proofs as specification rules
haftmann
parents: 66145
diff changeset
   817
        val head = head_of (strip_abs_body fu);
71179
592e2afdd50c more informative spec rules: optional name;
wenzelm
parents: 70915
diff changeset
   818
        val b = Binding.qualified_name (extr_name s vs);
66145
haftmann
parents: 64556
diff changeset
   819
      in
haftmann
parents: 64556
diff changeset
   820
        thy
71179
592e2afdd50c more informative spec rules: optional name;
wenzelm
parents: 70915
diff changeset
   821
        |> Sign.add_consts [(b, fastype_of ft, NoSyn)]
66145
haftmann
parents: 64556
diff changeset
   822
        |> Global_Theory.add_defs false
haftmann
parents: 64556
diff changeset
   823
           [((Binding.qualified_name (Thm.def_name (extr_name s vs)),
66146
fd3e128b174f register equations stemming from extracted proofs as specification rules
haftmann
parents: 66145
diff changeset
   824
             Logic.mk_equals (head, ft)), [])]
fd3e128b174f register equations stemming from extracted proofs as specification rules
haftmann
parents: 66145
diff changeset
   825
        |-> (fn [def_thm] =>
71214
5727bcc3c47c proper spec_rule name via naming/binding/Morphism.binding;
wenzelm
parents: 71206
diff changeset
   826
           Spec_Rules.add_global b Spec_Rules.equational
71206
20dce31fe7f4 proper spec rules via resulting def_thm, e.g. relevant for "isabelle build -o export_theory";
wenzelm
parents: 71179
diff changeset
   827
             [Thm.term_of (Thm.lhs_of def_thm)] [def_thm]
66251
cd935b7cb3fb proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents: 66146
diff changeset
   828
           #> Code.declare_default_eqns_global [(def_thm, true)])
66145
haftmann
parents: 64556
diff changeset
   829
      end;
haftmann
parents: 64556
diff changeset
   830
haftmann
parents: 64556
diff changeset
   831
    fun add_corr (s, (vs, prf)) thy =
haftmann
parents: 64556
diff changeset
   832
      let
70447
755d58b48cec clarified modules: provide reconstruct_proof / expand_proof at the bottom of proof term construction;
wenzelm
parents: 70436
diff changeset
   833
        val corr_prop = Proofterm.prop_of prf;
66145
haftmann
parents: 64556
diff changeset
   834
      in
haftmann
parents: 64556
diff changeset
   835
        thy
haftmann
parents: 64556
diff changeset
   836
        |> Global_Theory.store_thm (Binding.qualified_name (corr_name s vs),
haftmann
parents: 64556
diff changeset
   837
            Thm.varifyT_global (funpow (length (vars_of corr_prop))
haftmann
parents: 64556
diff changeset
   838
              (Thm.forall_elim_var 0) (Thm.forall_intr_frees
haftmann
parents: 64556
diff changeset
   839
                (Proof_Checker.thm_of_proof thy
haftmann
parents: 64556
diff changeset
   840
                  (fst (Proofterm.freeze_thaw_prf prf))))))
haftmann
parents: 64556
diff changeset
   841
        |> snd
haftmann
parents: 64556
diff changeset
   842
      end;
haftmann
parents: 64556
diff changeset
   843
haftmann
parents: 64556
diff changeset
   844
    fun add_def_and_corr (s, (vs, ((t, u), (prf, _)))) thy =
haftmann
parents: 64556
diff changeset
   845
      if is_none (Sign.const_type thy (extr_name s vs))
haftmann
parents: 64556
diff changeset
   846
      then
haftmann
parents: 64556
diff changeset
   847
        thy
haftmann
parents: 64556
diff changeset
   848
        |> not (t = nullt) ? add_def (s, (vs, ((t, u))))
haftmann
parents: 64556
diff changeset
   849
        |> add_corr (s, (vs, prf))
haftmann
parents: 64556
diff changeset
   850
      else
haftmann
parents: 64556
diff changeset
   851
        thy;
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   852
16149
d8cac577493c Theory.restore_naming;
wenzelm
parents: 15801
diff changeset
   853
  in
d8cac577493c Theory.restore_naming;
wenzelm
parents: 15801
diff changeset
   854
    thy
30435
e62d6ecab6ad explicit Binding.qualified_name -- prevents implicitly qualified bstring;
wenzelm
parents: 30364
diff changeset
   855
    |> Sign.root_path
66145
haftmann
parents: 64556
diff changeset
   856
    |> fold_rev add_def_and_corr defs
22796
34c316d7b630 renamed some old names Theory.xxx to Sign.xxx;
wenzelm
parents: 22750
diff changeset
   857
    |> Sign.restore_naming thy
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   858
  end;
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   859
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16363
diff changeset
   860
val etype_of = etype_of o add_syntax;
13714
bdd483321f4b - exported functions etype_of and mk_typ
berghofe
parents: 13609
diff changeset
   861
13402
e6e826bb8c3c Added program extraction module.
berghofe
parents:
diff changeset
   862
end;