| author | bulwahn | 
| Thu, 12 Jul 2012 16:22:33 +0200 | |
| changeset 48253 | 4410a709913c | 
| parent 47455 | 26315a545e26 | 
| child 52788 | da1fdbfebd39 | 
| permissions | -rw-r--r-- | 
| 47455 | 1  | 
(* Title: HOL/Matrix_LP/Compute_Oracle/compute.ML  | 
| 23174 | 2  | 
Author: Steven Obua  | 
3  | 
*)  | 
|
4  | 
||
5  | 
signature COMPUTE = sig  | 
|
6  | 
||
7  | 
type computer  | 
|
| 25217 | 8  | 
type theorem  | 
9  | 
type naming = int -> string  | 
|
| 23174 | 10  | 
|
| 23663 | 11  | 
datatype machine = BARRAS | BARRAS_COMPILED | HASKELL | SML  | 
| 23174 | 12  | 
|
| 25217 | 13  | 
(* Functions designated with a ! in front of them actually update the computer parameter *)  | 
14  | 
||
| 23663 | 15  | 
exception Make of string  | 
16  | 
val make : machine -> theory -> thm list -> computer  | 
|
| 25520 | 17  | 
val make_with_cache : machine -> theory -> term list -> thm list -> computer  | 
| 23174 | 18  | 
val theory_of : computer -> theory  | 
| 23663 | 19  | 
val hyps_of : computer -> term list  | 
20  | 
val shyps_of : computer -> sort list  | 
|
| 25217 | 21  | 
(* ! *) val update : computer -> thm list -> unit  | 
| 25520 | 22  | 
(* ! *) val update_with_cache : computer -> term list -> thm list -> unit  | 
| 25217 | 23  | 
|
24  | 
(* ! *) val set_naming : computer -> naming -> unit  | 
|
25  | 
val naming_of : computer -> naming  | 
|
26  | 
||
27  | 
exception Compute of string  | 
|
28  | 
val simplify : computer -> theorem -> thm  | 
|
29  | 
val rewrite : computer -> cterm -> thm  | 
|
| 23663 | 30  | 
|
| 25217 | 31  | 
val make_theorem : computer -> thm -> string list -> theorem  | 
32  | 
(* ! *) val instantiate : computer -> (string * cterm) list -> theorem -> theorem  | 
|
33  | 
(* ! *) val evaluate_prem : computer -> int -> theorem -> theorem  | 
|
34  | 
(* ! *) val modus_ponens : computer -> int -> thm -> theorem -> theorem  | 
|
| 23663 | 35  | 
|
| 23174 | 36  | 
end  | 
37  | 
||
| 23663 | 38  | 
structure Compute :> COMPUTE = struct  | 
39  | 
||
| 25217 | 40  | 
open Report;  | 
| 24654 | 41  | 
|
| 
26626
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
42  | 
datatype machine = BARRAS | BARRAS_COMPILED | HASKELL | SML  | 
| 23663 | 43  | 
|
44  | 
(* Terms are mapped to integer codes *)  | 
|
45  | 
structure Encode :>  | 
|
46  | 
sig  | 
|
47  | 
type encoding  | 
|
48  | 
val empty : encoding  | 
|
49  | 
val insert : term -> encoding -> int * encoding  | 
|
50  | 
val lookup_code : term -> encoding -> int option  | 
|
| 46536 | 51  | 
val lookup_term : int -> encoding -> term option  | 
| 23663 | 52  | 
val remove_code : int -> encoding -> encoding  | 
53  | 
val remove_term : term -> encoding -> encoding  | 
|
54  | 
end  | 
|
55  | 
=  | 
|
56  | 
struct  | 
|
57  | 
||
58  | 
type encoding = int * (int Termtab.table) * (term Inttab.table)  | 
|
59  | 
||
60  | 
val empty = (0, Termtab.empty, Inttab.empty)  | 
|
61  | 
||
62  | 
fun insert t (e as (count, term2int, int2term)) =  | 
|
63  | 
(case Termtab.lookup term2int t of  | 
|
| 
26626
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
64  | 
NONE => (count, (count+1, Termtab.update_new (t, count) term2int, Inttab.update_new (count, t) int2term))  | 
| 23663 | 65  | 
| SOME code => (code, e))  | 
66  | 
||
67  | 
fun lookup_code t (_, term2int, _) = Termtab.lookup term2int t  | 
|
68  | 
||
69  | 
fun lookup_term c (_, _, int2term) = Inttab.lookup int2term c  | 
|
70  | 
||
71  | 
fun remove_code c (e as (count, term2int, int2term)) =  | 
|
72  | 
(case lookup_term c e of NONE => e | SOME t => (count, Termtab.delete t term2int, Inttab.delete c int2term))  | 
|
73  | 
||
74  | 
fun remove_term t (e as (count, term2int, int2term)) =  | 
|
75  | 
(case lookup_code t e of NONE => e | SOME c => (count, Termtab.delete t term2int, Inttab.delete c int2term))  | 
|
76  | 
||
77  | 
end  | 
|
78  | 
||
| 23174 | 79  | 
exception Make of string;  | 
| 23663 | 80  | 
exception Compute of string;  | 
| 23174 | 81  | 
|
| 23663 | 82  | 
local  | 
| 46531 | 83  | 
fun make_constant t encoding =  | 
| 
26626
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
84  | 
let  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
85  | 
val (code, encoding) = Encode.insert t encoding  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
86  | 
in  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
87  | 
(encoding, AbstractMachine.Const code)  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
88  | 
end  | 
| 23663 | 89  | 
in  | 
| 23174 | 90  | 
|
| 23663 | 91  | 
fun remove_types encoding t =  | 
92  | 
case t of  | 
|
| 46531 | 93  | 
Var _ => make_constant t encoding  | 
94  | 
| Free _ => make_constant t encoding  | 
|
95  | 
| Const _ => make_constant t encoding  | 
|
96  | 
| Abs (_, _, t') =>  | 
|
| 
26626
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
97  | 
let val (encoding, t'') = remove_types encoding t' in  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
98  | 
(encoding, AbstractMachine.Abs t'')  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
99  | 
end  | 
| 23663 | 100  | 
| a $ b =>  | 
| 
26626
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
101  | 
let  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
102  | 
val (encoding, a) = remove_types encoding a  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
103  | 
val (encoding, b) = remove_types encoding b  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
104  | 
in  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
105  | 
(encoding, AbstractMachine.App (a,b))  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
106  | 
end  | 
| 23663 | 107  | 
| Bound b => (encoding, AbstractMachine.Var b)  | 
108  | 
end  | 
|
109  | 
||
110  | 
local  | 
|
111  | 
fun type_of (Free (_, ty)) = ty  | 
|
112  | 
| type_of (Const (_, ty)) = ty  | 
|
113  | 
| type_of (Var (_, ty)) = ty  | 
|
| 
40314
 
b5ec88d9ac03
replaced ancient sys_error by raise Fail, assuming that the latter is not handled specifically by the environment;
 
wenzelm 
parents: 
38808 
diff
changeset
 | 
114  | 
| type_of _ = raise Fail "infer_types: type_of error"  | 
| 23663 | 115  | 
in  | 
116  | 
fun infer_types naming encoding =  | 
|
| 23174 | 117  | 
let  | 
| 42364 | 118  | 
fun infer_types _ bounds _ (AbstractMachine.Var v) = (Bound v, nth bounds v)  | 
| 
26626
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
119  | 
| infer_types _ bounds _ (AbstractMachine.Const code) =  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
120  | 
let  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
121  | 
val c = the (Encode.lookup_term code encoding)  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
122  | 
in  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
123  | 
(c, type_of c)  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
124  | 
end  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
125  | 
| infer_types level bounds _ (AbstractMachine.App (a, b)) =  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
126  | 
let  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
127  | 
val (a, aty) = infer_types level bounds NONE a  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
128  | 
val (adom, arange) =  | 
| 23174 | 129  | 
case aty of  | 
130  | 
                        Type ("fun", [dom, range]) => (dom, range)
 | 
|
| 
40314
 
b5ec88d9ac03
replaced ancient sys_error by raise Fail, assuming that the latter is not handled specifically by the environment;
 
wenzelm 
parents: 
38808 
diff
changeset
 | 
131  | 
| _ => raise Fail "infer_types: function type expected"  | 
| 46531 | 132  | 
val (b, _) = infer_types level bounds (SOME adom) b  | 
| 
26626
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
133  | 
in  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
134  | 
(a $ b, arange)  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
135  | 
end  | 
| 23663 | 136  | 
          | infer_types level bounds (SOME (ty as Type ("fun", [dom, range]))) (AbstractMachine.Abs m) =
 | 
| 23174 | 137  | 
let  | 
| 23663 | 138  | 
val (m, _) = infer_types (level+1) (dom::bounds) (SOME range) m  | 
| 23174 | 139  | 
in  | 
| 23663 | 140  | 
(Abs (naming level, dom, m), ty)  | 
| 23174 | 141  | 
end  | 
| 46531 | 142  | 
| infer_types _ _ NONE (AbstractMachine.Abs _) =  | 
| 
40314
 
b5ec88d9ac03
replaced ancient sys_error by raise Fail, assuming that the latter is not handled specifically by the environment;
 
wenzelm 
parents: 
38808 
diff
changeset
 | 
143  | 
raise Fail "infer_types: cannot infer type of abstraction"  | 
| 23174 | 144  | 
|
| 23663 | 145  | 
fun infer ty term =  | 
| 23174 | 146  | 
let  | 
| 23663 | 147  | 
val (term', _) = infer_types 0 [] (SOME ty) term  | 
| 23174 | 148  | 
in  | 
149  | 
term'  | 
|
150  | 
end  | 
|
151  | 
in  | 
|
152  | 
infer  | 
|
153  | 
end  | 
|
| 23663 | 154  | 
end  | 
| 23174 | 155  | 
|
| 23663 | 156  | 
datatype prog =  | 
| 
26626
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
157  | 
ProgBarras of AM_Interpreter.program  | 
| 23663 | 158  | 
| ProgBarrasC of AM_Compiler.program  | 
159  | 
| ProgHaskell of AM_GHC.program  | 
|
160  | 
| ProgSML of AM_SML.program  | 
|
| 23174 | 161  | 
|
| 25217 | 162  | 
fun machine_of_prog (ProgBarras _) = BARRAS  | 
163  | 
| machine_of_prog (ProgBarrasC _) = BARRAS_COMPILED  | 
|
164  | 
| machine_of_prog (ProgHaskell _) = HASKELL  | 
|
165  | 
| machine_of_prog (ProgSML _) = SML  | 
|
166  | 
||
167  | 
type naming = int -> string  | 
|
168  | 
||
| 41491 | 169  | 
fun default_naming i = "v_" ^ string_of_int i  | 
| 25217 | 170  | 
|
| 32740 | 171  | 
datatype computer = Computer of  | 
172  | 
(theory_ref * Encode.encoding * term list * unit Sorttab.table * prog * unit Unsynchronized.ref * naming)  | 
|
173  | 
option Unsynchronized.ref  | 
|
| 25217 | 174  | 
|
| 32740 | 175  | 
fun theory_of (Computer (Unsynchronized.ref (SOME (rthy,_,_,_,_,_,_)))) = Theory.deref rthy  | 
176  | 
fun hyps_of (Computer (Unsynchronized.ref (SOME (_,_,hyps,_,_,_,_)))) = hyps  | 
|
177  | 
fun shyps_of (Computer (Unsynchronized.ref (SOME (_,_,_,shyptable,_,_,_)))) = Sorttab.keys (shyptable)  | 
|
178  | 
fun shyptab_of (Computer (Unsynchronized.ref (SOME (_,_,_,shyptable,_,_,_)))) = shyptable  | 
|
179  | 
fun stamp_of (Computer (Unsynchronized.ref (SOME (_,_,_,_,_,stamp,_)))) = stamp  | 
|
180  | 
fun prog_of (Computer (Unsynchronized.ref (SOME (_,_,_,_,prog,_,_)))) = prog  | 
|
181  | 
fun encoding_of (Computer (Unsynchronized.ref (SOME (_,encoding,_,_,_,_,_)))) = encoding  | 
|
| 46531 | 182  | 
fun set_encoding (Computer (r as Unsynchronized.ref (SOME (p1,_,p2,p3,p4,p5,p6)))) encoding' =  | 
| 25217 | 183  | 
(r := SOME (p1,encoding',p2,p3,p4,p5,p6))  | 
| 32740 | 184  | 
fun naming_of (Computer (Unsynchronized.ref (SOME (_,_,_,_,_,_,n)))) = n  | 
| 46531 | 185  | 
fun set_naming (Computer (r as Unsynchronized.ref (SOME (p1,p2,p3,p4,p5,p6,_)))) naming'=  | 
| 25217 | 186  | 
(r := SOME (p1,p2,p3,p4,p5,p6,naming'))  | 
187  | 
||
188  | 
fun ref_of (Computer r) = r  | 
|
| 23663 | 189  | 
|
190  | 
datatype cthm = ComputeThm of term list * sort list * term  | 
|
191  | 
||
192  | 
fun thm2cthm th =  | 
|
| 23174 | 193  | 
let  | 
| 
26626
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
194  | 
        val {hyps, prop, tpairs, shyps, ...} = Thm.rep_thm th
 | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
195  | 
val _ = if not (null tpairs) then raise Make "theorems may not contain tpairs" else ()  | 
| 23663 | 196  | 
in  | 
| 
26626
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
197  | 
ComputeThm (hyps, shyps, prop)  | 
| 23663 | 198  | 
end  | 
| 23174 | 199  | 
|
| 25520 | 200  | 
fun make_internal machine thy stamp encoding cache_pattern_terms raw_ths =  | 
| 23663 | 201  | 
let  | 
| 
26626
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
202  | 
fun transfer (x:thm) = Thm.transfer thy x  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
203  | 
val ths = map (thm2cthm o Thm.strip_shyps o transfer) raw_ths  | 
| 23174 | 204  | 
|
| 46531 | 205  | 
fun make_pattern encoding n vars (AbstractMachine.Abs _) =  | 
| 
26626
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
206  | 
raise (Make "no lambda abstractions allowed in pattern")  | 
| 46531 | 207  | 
| make_pattern encoding n vars (AbstractMachine.Var _) =  | 
| 
26626
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
208  | 
raise (Make "no bound variables allowed in pattern")  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
209  | 
| make_pattern encoding n vars (AbstractMachine.Const code) =  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
210  | 
(case the (Encode.lookup_term code encoding) of  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
211  | 
Var _ => ((n+1, Inttab.update_new (code, n) vars, AbstractMachine.PVar)  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
212  | 
handle Inttab.DUP _ => raise (Make "no duplicate variable in pattern allowed"))  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
213  | 
| _ => (n, vars, AbstractMachine.PConst (code, [])))  | 
| 25520 | 214  | 
| make_pattern encoding n vars (AbstractMachine.App (a, b)) =  | 
215  | 
let  | 
|
216  | 
val (n, vars, pa) = make_pattern encoding n vars a  | 
|
217  | 
val (n, vars, pb) = make_pattern encoding n vars b  | 
|
218  | 
in  | 
|
219  | 
case pa of  | 
|
220  | 
AbstractMachine.PVar =>  | 
|
221  | 
raise (Make "patterns may not start with a variable")  | 
|
222  | 
| AbstractMachine.PConst (c, args) =>  | 
|
223  | 
(n, vars, AbstractMachine.PConst (c, args@[pb]))  | 
|
224  | 
end  | 
|
225  | 
||
| 23663 | 226  | 
fun thm2rule (encoding, hyptable, shyptable) th =  | 
227  | 
let  | 
|
| 
26626
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
228  | 
val (ComputeThm (hyps, shyps, prop)) = th  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
229  | 
val hyptable = fold (fn h => Termtab.update (h, ())) hyps hyptable  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
230  | 
val shyptable = fold (fn sh => Sorttab.update (sh, ())) shyps shyptable  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
231  | 
val (prems, prop) = (Logic.strip_imp_prems prop, Logic.strip_imp_concl prop)  | 
| 23663 | 232  | 
val (a, b) = Logic.dest_equals prop  | 
233  | 
handle TERM _ => raise (Make "theorems must be meta-level equations (with optional guards)")  | 
|
| 
26626
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
234  | 
val a = Envir.eta_contract a  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
235  | 
val b = Envir.eta_contract b  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
236  | 
val prems = map Envir.eta_contract prems  | 
| 23174 | 237  | 
|
| 23663 | 238  | 
val (encoding, left) = remove_types encoding a  | 
| 
26626
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
239  | 
val (encoding, right) = remove_types encoding b  | 
| 23663 | 240  | 
fun remove_types_of_guard encoding g =  | 
| 
26626
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
241  | 
(let  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
242  | 
val (t1, t2) = Logic.dest_equals g  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
243  | 
val (encoding, t1) = remove_types encoding t1  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
244  | 
val (encoding, t2) = remove_types encoding t2  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
245  | 
in  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
246  | 
(encoding, AbstractMachine.Guard (t1, t2))  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
247  | 
end handle TERM _ => raise (Make "guards must be meta-level equations"))  | 
| 23663 | 248  | 
val (encoding, prems) = fold_rev (fn p => fn (encoding, ps) => let val (e, p) = remove_types_of_guard encoding p in (e, p::ps) end) prems (encoding, [])  | 
| 23174 | 249  | 
|
| 23663 | 250  | 
(* Principally, a check should be made here to see if the (meta-) hyps contain any of the variables of the rule.  | 
251  | 
As it is, all variables of the rule are schematic, and there are no schematic variables in meta-hyps, therefore  | 
|
252  | 
this check can be left out. *)  | 
|
253  | 
||
254  | 
val (vcount, vars, pattern) = make_pattern encoding 0 Inttab.empty left  | 
|
| 23174 | 255  | 
val _ = (case pattern of  | 
| 23663 | 256  | 
AbstractMachine.PVar =>  | 
| 23174 | 257  | 
raise (Make "patterns may not start with a variable")  | 
| 
26626
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
258  | 
| _ => ())  | 
| 23174 | 259  | 
|
260  | 
(* finally, provide a function for renaming the  | 
|
| 23663 | 261  | 
pattern bound variables on the right hand side *)  | 
| 23174 | 262  | 
|
| 23663 | 263  | 
fun rename level vars (var as AbstractMachine.Var _) = var  | 
| 
26626
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
264  | 
| rename level vars (c as AbstractMachine.Const code) =  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
265  | 
(case Inttab.lookup vars code of  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
266  | 
NONE => c  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
267  | 
| SOME n => AbstractMachine.Var (vcount-n-1+level))  | 
| 23663 | 268  | 
| rename level vars (AbstractMachine.App (a, b)) =  | 
269  | 
AbstractMachine.App (rename level vars a, rename level vars b)  | 
|
270  | 
| rename level vars (AbstractMachine.Abs m) =  | 
|
271  | 
AbstractMachine.Abs (rename (level+1) vars m)  | 
|
| 
26626
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
272  | 
|
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
273  | 
fun rename_guard (AbstractMachine.Guard (a,b)) =  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
274  | 
AbstractMachine.Guard (rename 0 vars a, rename 0 vars b)  | 
| 23174 | 275  | 
in  | 
| 23663 | 276  | 
((encoding, hyptable, shyptable), (map rename_guard prems, pattern, rename 0 vars right))  | 
| 23174 | 277  | 
end  | 
278  | 
||
| 23663 | 279  | 
val ((encoding, hyptable, shyptable), rules) =  | 
280  | 
fold_rev (fn th => fn (encoding_hyptable, rules) =>  | 
|
| 23174 | 281  | 
let  | 
| 23663 | 282  | 
val (encoding_hyptable, rule) = thm2rule encoding_hyptable th  | 
283  | 
in (encoding_hyptable, rule::rules) end)  | 
|
| 25217 | 284  | 
ths ((encoding, Termtab.empty, Sorttab.empty), [])  | 
| 23174 | 285  | 
|
| 25520 | 286  | 
fun make_cache_pattern t (encoding, cache_patterns) =  | 
| 
26626
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
287  | 
let  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
288  | 
val (encoding, a) = remove_types encoding t  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
289  | 
val (_,_,p) = make_pattern encoding 0 Inttab.empty a  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
290  | 
in  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
291  | 
(encoding, p::cache_patterns)  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
292  | 
end  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
293  | 
|
| 46534 | 294  | 
val (encoding, _) = fold_rev make_cache_pattern cache_pattern_terms (encoding, [])  | 
| 25520 | 295  | 
|
| 23663 | 296  | 
val prog =  | 
| 
26626
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
297  | 
case machine of  | 
| 46534 | 298  | 
BARRAS => ProgBarras (AM_Interpreter.compile rules)  | 
299  | 
| BARRAS_COMPILED => ProgBarrasC (AM_Compiler.compile rules)  | 
|
300  | 
| HASKELL => ProgHaskell (AM_GHC.compile rules)  | 
|
301  | 
| SML => ProgSML (AM_SML.compile rules)  | 
|
| 23174 | 302  | 
|
| 23663 | 303  | 
fun has_witness s = not (null (Sign.witness_sorts thy [] [s]))  | 
| 23174 | 304  | 
|
| 
26626
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
305  | 
val shyptable = fold Sorttab.delete (filter has_witness (Sorttab.keys (shyptable))) shyptable  | 
| 23663 | 306  | 
|
| 25217 | 307  | 
in (Theory.check_thy thy, encoding, Termtab.keys hyptable, shyptable, prog, stamp, default_naming) end  | 
308  | 
||
| 32740 | 309  | 
fun make_with_cache machine thy cache_patterns raw_thms =  | 
310  | 
Computer (Unsynchronized.ref (SOME (make_internal machine thy (Unsynchronized.ref ()) Encode.empty cache_patterns raw_thms)))  | 
|
| 23663 | 311  | 
|
| 25520 | 312  | 
fun make machine thy raw_thms = make_with_cache machine thy [] raw_thms  | 
313  | 
||
314  | 
fun update_with_cache computer cache_patterns raw_thms =  | 
|
| 25217 | 315  | 
let  | 
| 
26626
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
316  | 
val c = make_internal (machine_of_prog (prog_of computer)) (theory_of computer) (stamp_of computer)  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
317  | 
(encoding_of computer) cache_patterns raw_thms  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
318  | 
val _ = (ref_of computer) := SOME c  | 
| 25217 | 319  | 
in  | 
| 
26626
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
320  | 
()  | 
| 25217 | 321  | 
end  | 
322  | 
||
| 25520 | 323  | 
fun update computer raw_thms = update_with_cache computer [] raw_thms  | 
324  | 
||
| 25217 | 325  | 
fun runprog (ProgBarras p) = AM_Interpreter.run p  | 
326  | 
| runprog (ProgBarrasC p) = AM_Compiler.run p  | 
|
327  | 
| runprog (ProgHaskell p) = AM_GHC.run p  | 
|
| 
26626
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
328  | 
| runprog (ProgSML p) = AM_SML.run p  | 
| 25217 | 329  | 
|
330  | 
(* ------------------------------------------------------------------------------------- *)  | 
|
331  | 
(* An oracle for exporting theorems; must only be accessible from inside this structure! *)  | 
|
332  | 
(* ------------------------------------------------------------------------------------- *)  | 
|
333  | 
||
334  | 
fun merge_hyps hyps1 hyps2 =  | 
|
335  | 
let  | 
|
336  | 
fun add hyps tab = fold (fn h => fn tab => Termtab.update (h, ()) tab) hyps tab  | 
|
337  | 
in  | 
|
338  | 
Termtab.keys (add hyps2 (add hyps1 Termtab.empty))  | 
|
339  | 
end  | 
|
340  | 
||
341  | 
fun add_shyps shyps tab = fold (fn h => fn tab => Sorttab.update (h, ()) tab) shyps tab  | 
|
342  | 
||
343  | 
fun merge_shyps shyps1 shyps2 = Sorttab.keys (add_shyps shyps2 (add_shyps shyps1 Sorttab.empty))  | 
|
344  | 
||
| 28290 | 345  | 
val (_, export_oracle) = Context.>>> (Context.map_theory_result  | 
| 38808 | 346  | 
  (Thm.add_oracle (@{binding compute}, fn (thy, hyps, shyps, prop) =>
 | 
| 25217 | 347  | 
let  | 
| 
26626
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
348  | 
val shyptab = add_shyps shyps Sorttab.empty  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
349  | 
fun delete s shyptab = Sorttab.delete s shyptab handle Sorttab.UNDEF _ => shyptab  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
350  | 
fun delete_term t shyptab = fold delete (Sorts.insert_term t []) shyptab  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
351  | 
fun has_witness s = not (null (Sign.witness_sorts thy [] [s]))  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
352  | 
val shyptab = fold Sorttab.delete (filter has_witness (Sorttab.keys (shyptab))) shyptab  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
353  | 
val shyps = if Sorttab.is_empty shyptab then [] else Sorttab.keys (fold delete_term (prop::hyps) shyptab)  | 
| 
31322
 
526e149999cc
attempt to eliminate adhoc makestring at runtime (which is not well-defined);
 
wenzelm 
parents: 
30288 
diff
changeset
 | 
354  | 
val _ =  | 
| 
 
526e149999cc
attempt to eliminate adhoc makestring at runtime (which is not well-defined);
 
wenzelm 
parents: 
30288 
diff
changeset
 | 
355  | 
if not (null shyps) then  | 
| 
 
526e149999cc
attempt to eliminate adhoc makestring at runtime (which is not well-defined);
 
wenzelm 
parents: 
30288 
diff
changeset
 | 
356  | 
            raise Compute ("dangling sort hypotheses: " ^
 | 
| 
 
526e149999cc
attempt to eliminate adhoc makestring at runtime (which is not well-defined);
 
wenzelm 
parents: 
30288 
diff
changeset
 | 
357  | 
commas (map (Syntax.string_of_sort_global thy) shyps))  | 
| 
 
526e149999cc
attempt to eliminate adhoc makestring at runtime (which is not well-defined);
 
wenzelm 
parents: 
30288 
diff
changeset
 | 
358  | 
else ()  | 
| 25217 | 359  | 
in  | 
| 28290 | 360  | 
Thm.cterm_of thy (fold_rev (fn hyp => fn p => Logic.mk_implies (hyp, p)) hyps prop)  | 
361  | 
end)));  | 
|
| 25217 | 362  | 
|
363  | 
fun export_thm thy hyps shyps prop =  | 
|
364  | 
let  | 
|
| 28290 | 365  | 
val th = export_oracle (thy, hyps, shyps, prop)  | 
| 36945 | 366  | 
val hyps = map (fn h => Thm.assume (cterm_of thy h)) hyps  | 
| 25217 | 367  | 
in  | 
| 36945 | 368  | 
fold (fn h => fn p => Thm.implies_elim p h) hyps th  | 
| 25217 | 369  | 
end  | 
| 
26626
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
370  | 
|
| 25217 | 371  | 
(* --------- Rewrite ----------- *)  | 
372  | 
||
373  | 
fun rewrite computer ct =  | 
|
374  | 
let  | 
|
| 
26626
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
375  | 
val thy = Thm.theory_of_cterm ct  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
376  | 
        val {t=t',T=ty,...} = rep_cterm ct
 | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
377  | 
val _ = Theory.assert_super (theory_of computer) thy  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
378  | 
val naming = naming_of computer  | 
| 25217 | 379  | 
val (encoding, t) = remove_types (encoding_of computer) t'  | 
380  | 
val t = runprog (prog_of computer) t  | 
|
381  | 
val t = infer_types naming encoding ty t  | 
|
| 
26626
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
382  | 
val eq = Logic.mk_equals (t', t)  | 
| 25217 | 383  | 
in  | 
384  | 
export_thm thy (hyps_of computer) (Sorttab.keys (shyptab_of computer)) eq  | 
|
385  | 
end  | 
|
386  | 
||
387  | 
(* --------- Simplify ------------ *)  | 
|
| 23663 | 388  | 
|
| 25217 | 389  | 
datatype prem = EqPrem of AbstractMachine.term * AbstractMachine.term * Term.typ * int  | 
| 
26626
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
390  | 
| Prem of AbstractMachine.term  | 
| 32740 | 391  | 
datatype theorem = Theorem of theory_ref * unit Unsynchronized.ref * (int * typ) Symtab.table * (AbstractMachine.term option) Inttab.table  | 
| 25217 | 392  | 
* prem list * AbstractMachine.term * term list * sort list  | 
393  | 
||
394  | 
||
395  | 
exception ParamSimplify of computer * theorem  | 
|
396  | 
||
397  | 
fun make_theorem computer th vars =  | 
|
398  | 
let  | 
|
399  | 
val _ = Theory.assert_super (theory_of computer) (theory_of_thm th)  | 
|
400  | 
||
401  | 
val (ComputeThm (hyps, shyps, prop)) = thm2cthm th  | 
|
402  | 
||
403  | 
val encoding = encoding_of computer  | 
|
404  | 
||
405  | 
(* variables in the theorem are identified upfront *)  | 
|
406  | 
fun collect_vars (Abs (_, _, t)) tab = collect_vars t tab  | 
|
407  | 
| collect_vars (a $ b) tab = collect_vars b (collect_vars a tab)  | 
|
408  | 
| collect_vars (Const _) tab = tab  | 
|
409  | 
| collect_vars (Free _) tab = tab  | 
|
410  | 
| collect_vars (Var ((s, i), ty)) tab =  | 
|
| 
26626
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
411  | 
if List.find (fn x => x=s) vars = NONE then  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
412  | 
tab  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
413  | 
else  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
414  | 
(case Symtab.lookup tab s of  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
415  | 
SOME ((s',i'),ty') =>  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
416  | 
if s' <> s orelse i' <> i orelse ty <> ty' then  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
417  | 
                         raise Compute ("make_theorem: variable name '"^s^"' is not unique")
 | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
418  | 
else  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
419  | 
tab  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
420  | 
| NONE => Symtab.update (s, ((s, i), ty)) tab)  | 
| 25217 | 421  | 
val vartab = collect_vars prop Symtab.empty  | 
422  | 
fun encodevar (s, t as (_, ty)) (encoding, tab) =  | 
|
| 
26626
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
423  | 
let  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
424  | 
val (x, encoding) = Encode.insert (Var t) encoding  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
425  | 
in  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
426  | 
(encoding, Symtab.update (s, (x, ty)) tab)  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
427  | 
end  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
428  | 
val (encoding, vartab) = Symtab.fold encodevar vartab (encoding, Symtab.empty)  | 
| 46531 | 429  | 
val varsubst = Inttab.make (map (fn (_, (x, _)) => (x, NONE)) (Symtab.dest vartab))  | 
| 23174 | 430  | 
|
| 25217 | 431  | 
(* make the premises and the conclusion *)  | 
432  | 
fun mk_prem encoding t =  | 
|
| 
26626
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
433  | 
(let  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
434  | 
val (a, b) = Logic.dest_equals t  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
435  | 
val ty = type_of a  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
436  | 
val (encoding, a) = remove_types encoding a  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
437  | 
val (encoding, b) = remove_types encoding b  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
438  | 
             val (eq, encoding) = Encode.insert (Const ("==", ty --> ty --> @{typ "prop"})) encoding 
 | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
439  | 
in  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
440  | 
(encoding, EqPrem (a, b, ty, eq))  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
441  | 
end handle TERM _ => let val (encoding, t) = remove_types encoding t in (encoding, Prem t) end)  | 
| 25217 | 442  | 
val (encoding, prems) =  | 
| 
26626
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
443  | 
(fold_rev (fn t => fn (encoding, l) =>  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
444  | 
case mk_prem encoding t of  | 
| 25217 | 445  | 
(encoding, t) => (encoding, t::l)) (Logic.strip_imp_prems prop) (encoding, []))  | 
446  | 
val (encoding, concl) = remove_types encoding (Logic.strip_imp_concl prop)  | 
|
447  | 
val _ = set_encoding computer encoding  | 
|
448  | 
in  | 
|
449  | 
Theorem (Theory.check_thy (theory_of_thm th), stamp_of computer, vartab, varsubst,  | 
|
| 
26626
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
450  | 
prems, concl, hyps, shyps)  | 
| 25217 | 451  | 
end  | 
452  | 
||
453  | 
fun theory_of_theorem (Theorem (rthy,_,_,_,_,_,_,_)) = Theory.deref rthy  | 
|
454  | 
fun update_theory thy (Theorem (_,p0,p1,p2,p3,p4,p5,p6)) =  | 
|
455  | 
Theorem (Theory.check_thy thy,p0,p1,p2,p3,p4,p5,p6)  | 
|
456  | 
fun stamp_of_theorem (Theorem (_,s, _, _, _, _, _, _)) = s  | 
|
457  | 
fun vartab_of_theorem (Theorem (_,_,vt,_,_,_,_,_)) = vt  | 
|
458  | 
fun varsubst_of_theorem (Theorem (_,_,_,vs,_,_,_,_)) = vs  | 
|
459  | 
fun update_varsubst vs (Theorem (p0,p1,p2,_,p3,p4,p5,p6)) = Theorem (p0,p1,p2,vs,p3,p4,p5,p6)  | 
|
460  | 
fun prems_of_theorem (Theorem (_,_,_,_,prems,_,_,_)) = prems  | 
|
461  | 
fun update_prems prems (Theorem (p0,p1,p2,p3,_,p4,p5,p6)) = Theorem (p0,p1,p2,p3,prems,p4,p5,p6)  | 
|
462  | 
fun concl_of_theorem (Theorem (_,_,_,_,_,concl,_,_)) = concl  | 
|
463  | 
fun hyps_of_theorem (Theorem (_,_,_,_,_,_,hyps,_)) = hyps  | 
|
464  | 
fun update_hyps hyps (Theorem (p0,p1,p2,p3,p4,p5,_,p6)) = Theorem (p0,p1,p2,p3,p4,p5,hyps,p6)  | 
|
465  | 
fun shyps_of_theorem (Theorem (_,_,_,_,_,_,_,shyps)) = shyps  | 
|
466  | 
fun update_shyps shyps (Theorem (p0,p1,p2,p3,p4,p5,p6,_)) = Theorem (p0,p1,p2,p3,p4,p5,p6,shyps)  | 
|
467  | 
||
468  | 
fun check_compatible computer th s =  | 
|
469  | 
if stamp_of computer <> stamp_of_theorem th then  | 
|
| 
26626
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
470  | 
raise Compute (s^": computer and theorem are incompatible")  | 
| 25217 | 471  | 
else ()  | 
472  | 
||
473  | 
fun instantiate computer insts th =  | 
|
474  | 
let  | 
|
475  | 
val _ = check_compatible computer th  | 
|
476  | 
||
477  | 
val thy = theory_of computer  | 
|
478  | 
||
479  | 
val vartab = vartab_of_theorem th  | 
|
480  | 
||
481  | 
fun rewrite computer t =  | 
|
482  | 
let  | 
|
483  | 
val (encoding, t) = remove_types (encoding_of computer) t  | 
|
484  | 
val t = runprog (prog_of computer) t  | 
|
| 
26626
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
485  | 
val _ = set_encoding computer encoding  | 
| 23174 | 486  | 
in  | 
487  | 
t  | 
|
488  | 
end  | 
|
489  | 
||
| 25217 | 490  | 
fun assert_varfree vs t =  | 
| 
26626
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
491  | 
if AbstractMachine.forall_consts (fn x => Inttab.lookup vs x = NONE) t then  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
492  | 
()  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
493  | 
else  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
494  | 
raise Compute "instantiate: assert_varfree failed"  | 
| 25217 | 495  | 
|
496  | 
fun assert_closed t =  | 
|
| 
26626
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
497  | 
if AbstractMachine.closed t then  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
498  | 
()  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
499  | 
else  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
500  | 
raise Compute "instantiate: not a closed term"  | 
| 23663 | 501  | 
|
| 25217 | 502  | 
fun compute_inst (s, ct) vs =  | 
| 
26626
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
503  | 
let  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
504  | 
val _ = Theory.assert_super (theory_of_cterm ct) thy  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
505  | 
val ty = typ_of (ctyp_of_term ct)  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
506  | 
in  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
507  | 
(case Symtab.lookup vartab s of  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
508  | 
                 NONE => raise Compute ("instantiate: variable '"^s^"' not found in theorem")
 | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
509  | 
| SOME (x, ty') =>  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
510  | 
(case Inttab.lookup vs x of  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
511  | 
                      SOME (SOME _) => raise Compute ("instantiate: variable '"^s^"' has already been instantiated")
 | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
512  | 
| SOME NONE =>  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
513  | 
if ty <> ty' then  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
514  | 
                          raise Compute ("instantiate: wrong type for variable '"^s^"'")
 | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
515  | 
else  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
516  | 
let  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
517  | 
val t = rewrite computer (term_of ct)  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
518  | 
val _ = assert_varfree vs t  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
519  | 
val _ = assert_closed t  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
520  | 
in  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
521  | 
Inttab.update (x, SOME t) vs  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
522  | 
end  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
523  | 
| NONE => raise Compute "instantiate: internal error"))  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
524  | 
end  | 
| 23174 | 525  | 
|
| 25217 | 526  | 
val vs = fold compute_inst insts (varsubst_of_theorem th)  | 
527  | 
in  | 
|
528  | 
update_varsubst vs th  | 
|
529  | 
end  | 
|
| 23174 | 530  | 
|
| 25217 | 531  | 
fun match_aterms subst =  | 
| 
26626
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
532  | 
let  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
533  | 
exception no_match  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
534  | 
open AbstractMachine  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
535  | 
fun match subst (b as (Const c)) a =  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
536  | 
if a = b then subst  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
537  | 
else  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
538  | 
(case Inttab.lookup subst c of  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
539  | 
SOME (SOME a') => if a=a' then subst else raise no_match  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
540  | 
| SOME NONE => if AbstractMachine.closed a then  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
541  | 
Inttab.update (c, SOME a) subst  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
542  | 
else raise no_match  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
543  | 
| NONE => raise no_match)  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
544  | 
| match subst (b as (Var _)) a = if a=b then subst else raise no_match  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
545  | 
| match subst (App (u, v)) (App (u', v')) = match (match subst u u') v v'  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
546  | 
| match subst (Abs u) (Abs u') = match subst u u'  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
547  | 
| match subst _ _ = raise no_match  | 
| 23663 | 548  | 
in  | 
| 
26626
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
549  | 
fn b => fn a => (SOME (match subst b a) handle no_match => NONE)  | 
| 25217 | 550  | 
end  | 
551  | 
||
552  | 
fun apply_subst vars_allowed subst =  | 
|
553  | 
let  | 
|
| 
26626
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
554  | 
open AbstractMachine  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
555  | 
fun app (t as (Const c)) =  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
556  | 
(case Inttab.lookup subst c of  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
557  | 
NONE => t  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
558  | 
| SOME (SOME t) => Computed t  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
559  | 
| SOME NONE => if vars_allowed then t else raise Compute "apply_subst: no vars allowed")  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
560  | 
| app (t as (Var _)) = t  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
561  | 
| app (App (u, v)) = App (app u, app v)  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
562  | 
| app (Abs m) = Abs (app m)  | 
| 25217 | 563  | 
in  | 
| 
26626
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
564  | 
app  | 
| 23663 | 565  | 
end  | 
566  | 
||
| 25217 | 567  | 
fun splicein n l L = List.take (L, n) @ l @ List.drop (L, n+1)  | 
| 23663 | 568  | 
|
| 25217 | 569  | 
fun evaluate_prem computer prem_no th =  | 
570  | 
let  | 
|
571  | 
val _ = check_compatible computer th  | 
|
572  | 
val prems = prems_of_theorem th  | 
|
573  | 
val varsubst = varsubst_of_theorem th  | 
|
574  | 
fun run vars_allowed t =  | 
|
| 
26626
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
575  | 
runprog (prog_of computer) (apply_subst vars_allowed varsubst t)  | 
| 25217 | 576  | 
in  | 
| 42364 | 577  | 
case nth prems prem_no of  | 
| 
26626
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
578  | 
Prem _ => raise Compute "evaluate_prem: no equality premise"  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
579  | 
| EqPrem (a, b, ty, _) =>  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
580  | 
let  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
581  | 
val a' = run false a  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
582  | 
val b' = run true b  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
583  | 
in  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
584  | 
case match_aterms varsubst b' a' of  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
585  | 
NONE =>  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
586  | 
let  | 
| 
31322
 
526e149999cc
attempt to eliminate adhoc makestring at runtime (which is not well-defined);
 
wenzelm 
parents: 
30288 
diff
changeset
 | 
587  | 
fun mk s = Syntax.string_of_term_global Pure.thy  | 
| 
 
526e149999cc
attempt to eliminate adhoc makestring at runtime (which is not well-defined);
 
wenzelm 
parents: 
30288 
diff
changeset
 | 
588  | 
(infer_types (naming_of computer) (encoding_of computer) ty s)  | 
| 
26626
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
589  | 
val left = "computed left side: "^(mk a')  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
590  | 
val right = "computed right side: "^(mk b')  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
591  | 
in  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
592  | 
                    raise Compute ("evaluate_prem: cannot assign computed left to right hand side\n"^left^"\n"^right^"\n")
 | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
593  | 
end  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
594  | 
| SOME varsubst =>  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
595  | 
update_prems (splicein prem_no [] prems) (update_varsubst varsubst th)  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
596  | 
end  | 
| 25217 | 597  | 
end  | 
| 23663 | 598  | 
|
| 25217 | 599  | 
fun prem2term (Prem t) = t  | 
600  | 
| prem2term (EqPrem (a,b,_,eq)) =  | 
|
601  | 
AbstractMachine.App (AbstractMachine.App (AbstractMachine.Const eq, a), b)  | 
|
| 23663 | 602  | 
|
| 25217 | 603  | 
fun modus_ponens computer prem_no th' th =  | 
604  | 
let  | 
|
605  | 
val _ = check_compatible computer th  | 
|
606  | 
val thy =  | 
|
| 
26626
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
607  | 
let  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
608  | 
val thy1 = theory_of_theorem th  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
609  | 
val thy2 = theory_of_thm th'  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
610  | 
in  | 
| 26674 | 611  | 
if Theory.subthy (thy1, thy2) then thy2  | 
612  | 
else if Theory.subthy (thy2, thy1) then thy1 else  | 
|
| 
26626
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
613  | 
raise Compute "modus_ponens: theorems are not compatible with each other"  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
614  | 
end  | 
| 25217 | 615  | 
val th' = make_theorem computer th' []  | 
616  | 
val varsubst = varsubst_of_theorem th  | 
|
617  | 
fun run vars_allowed t =  | 
|
| 
26626
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
618  | 
runprog (prog_of computer) (apply_subst vars_allowed varsubst t)  | 
| 25217 | 619  | 
val prems = prems_of_theorem th  | 
| 42364 | 620  | 
val prem = run true (prem2term (nth prems prem_no))  | 
| 25217 | 621  | 
val concl = run false (concl_of_theorem th')  | 
622  | 
in  | 
|
623  | 
case match_aterms varsubst prem concl of  | 
|
| 
26626
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
624  | 
NONE => raise Compute "modus_ponens: conclusion does not match premise"  | 
| 25217 | 625  | 
| SOME varsubst =>  | 
| 
26626
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
626  | 
let  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
627  | 
val th = update_varsubst varsubst th  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
628  | 
val th = update_prems (splicein prem_no (prems_of_theorem th') prems) th  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
629  | 
val th = update_hyps (merge_hyps (hyps_of_theorem th) (hyps_of_theorem th')) th  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
630  | 
val th = update_shyps (merge_shyps (shyps_of_theorem th) (shyps_of_theorem th')) th  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
631  | 
in  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
632  | 
update_theory thy th  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
25520 
diff
changeset
 | 
633  | 
end  | 
| 25217 | 634  | 
end  | 
635  | 
||
636  | 
fun simplify computer th =  | 
|
637  | 
let  | 
|
638  | 
val _ = check_compatible computer th  | 
|
639  | 
val varsubst = varsubst_of_theorem th  | 
|
640  | 
val encoding = encoding_of computer  | 
|
641  | 
val naming = naming_of computer  | 
|
642  | 
    fun infer t = infer_types naming encoding @{typ "prop"} t
 | 
|
643  | 
fun run t = infer (runprog (prog_of computer) (apply_subst true varsubst t))  | 
|
644  | 
fun runprem p = run (prem2term p)  | 
|
645  | 
val prop = Logic.list_implies (map runprem (prems_of_theorem th), run (concl_of_theorem th))  | 
|
646  | 
val hyps = merge_hyps (hyps_of computer) (hyps_of_theorem th)  | 
|
647  | 
val shyps = merge_shyps (shyps_of_theorem th) (Sorttab.keys (shyptab_of computer))  | 
|
648  | 
in  | 
|
649  | 
export_thm (theory_of_theorem th) hyps shyps prop  | 
|
650  | 
end  | 
|
| 23174 | 651  | 
|
652  | 
end  | 
|
| 23663 | 653  |