| author | blanchet | 
| Sun, 01 May 2011 18:37:24 +0200 | |
| changeset 42557 | ae0deb39a254 | 
| parent 41460 | ea56b98aee83 | 
| child 44064 | 5bce8ff0d9ae | 
| permissions | -rw-r--r-- | 
| 12869 | 1 | (* Title: HOL/ex/SVC_Oracle.thy | 
| 2 | Author: Lawrence C Paulson | |
| 3 | Copyright 1999 University of Cambridge | |
| 4 | ||
| 40945 | 5 | Based upon the work of Søren T. Heilmann. | 
| 17388 | 6 | *) | 
| 12869 | 7 | |
| 17388 | 8 | header {* Installing an oracle for SVC (Stanford Validity Checker) *}
 | 
| 12869 | 9 | |
| 16836 | 10 | theory SVC_Oracle | 
| 11 | imports Main | |
| 24470 
41c81e23c08d
removed Hoare/hoare.ML, Hoare/hoareAbort.ML, ex/svc_oracle.ML (which can be mistaken as attached ML script on case-insensitive file-system);
 wenzelm parents: 
20813diff
changeset | 12 | uses "svc_funcs.ML" | 
| 16836 | 13 | begin | 
| 12869 | 14 | |
| 15 | consts | |
| 16 | iff_keep :: "[bool, bool] => bool" | |
| 17 | iff_unfold :: "[bool, bool] => bool" | |
| 18 | ||
| 36176 
3fe7e97ccca8
replaced generic 'hide' command by more conventional 'hide_class', 'hide_type', 'hide_const', 'hide_fact' -- frees some popular keywords;
 wenzelm parents: 
35267diff
changeset | 19 | hide_const iff_keep iff_unfold | 
| 16836 | 20 | |
| 28290 | 21 | oracle svc_oracle = Svc.oracle | 
| 12869 | 22 | |
| 24470 
41c81e23c08d
removed Hoare/hoare.ML, Hoare/hoareAbort.ML, ex/svc_oracle.ML (which can be mistaken as attached ML script on case-insensitive file-system);
 wenzelm parents: 
20813diff
changeset | 23 | ML {*
 | 
| 
41c81e23c08d
removed Hoare/hoare.ML, Hoare/hoareAbort.ML, ex/svc_oracle.ML (which can be mistaken as attached ML script on case-insensitive file-system);
 wenzelm parents: 
20813diff
changeset | 24 | (* | 
| 
41c81e23c08d
removed Hoare/hoare.ML, Hoare/hoareAbort.ML, ex/svc_oracle.ML (which can be mistaken as attached ML script on case-insensitive file-system);
 wenzelm parents: 
20813diff
changeset | 25 | Installing the oracle for SVC (Stanford Validity Checker) | 
| 
41c81e23c08d
removed Hoare/hoare.ML, Hoare/hoareAbort.ML, ex/svc_oracle.ML (which can be mistaken as attached ML script on case-insensitive file-system);
 wenzelm parents: 
20813diff
changeset | 26 | |
| 
41c81e23c08d
removed Hoare/hoare.ML, Hoare/hoareAbort.ML, ex/svc_oracle.ML (which can be mistaken as attached ML script on case-insensitive file-system);
 wenzelm parents: 
20813diff
changeset | 27 | The following code merely CALLS the oracle; | 
| 
41c81e23c08d
removed Hoare/hoare.ML, Hoare/hoareAbort.ML, ex/svc_oracle.ML (which can be mistaken as attached ML script on case-insensitive file-system);
 wenzelm parents: 
20813diff
changeset | 28 | the soundness-critical functions are at svc_funcs.ML | 
| 
41c81e23c08d
removed Hoare/hoare.ML, Hoare/hoareAbort.ML, ex/svc_oracle.ML (which can be mistaken as attached ML script on case-insensitive file-system);
 wenzelm parents: 
20813diff
changeset | 29 | |
| 40945 | 30 | Based upon the work of Søren T. Heilmann | 
| 24470 
41c81e23c08d
removed Hoare/hoare.ML, Hoare/hoareAbort.ML, ex/svc_oracle.ML (which can be mistaken as attached ML script on case-insensitive file-system);
 wenzelm parents: 
20813diff
changeset | 31 | *) | 
| 
41c81e23c08d
removed Hoare/hoare.ML, Hoare/hoareAbort.ML, ex/svc_oracle.ML (which can be mistaken as attached ML script on case-insensitive file-system);
 wenzelm parents: 
20813diff
changeset | 32 | |
| 
41c81e23c08d
removed Hoare/hoare.ML, Hoare/hoareAbort.ML, ex/svc_oracle.ML (which can be mistaken as attached ML script on case-insensitive file-system);
 wenzelm parents: 
20813diff
changeset | 33 | |
| 
41c81e23c08d
removed Hoare/hoare.ML, Hoare/hoareAbort.ML, ex/svc_oracle.ML (which can be mistaken as attached ML script on case-insensitive file-system);
 wenzelm parents: 
20813diff
changeset | 34 | (*Generalize an Isabelle formula, replacing by Vars | 
| 
41c81e23c08d
removed Hoare/hoare.ML, Hoare/hoareAbort.ML, ex/svc_oracle.ML (which can be mistaken as attached ML script on case-insensitive file-system);
 wenzelm parents: 
20813diff
changeset | 35 | all subterms not intelligible to SVC.*) | 
| 
41c81e23c08d
removed Hoare/hoare.ML, Hoare/hoareAbort.ML, ex/svc_oracle.ML (which can be mistaken as attached ML script on case-insensitive file-system);
 wenzelm parents: 
20813diff
changeset | 36 | fun svc_abstract t = | 
| 
41c81e23c08d
removed Hoare/hoare.ML, Hoare/hoareAbort.ML, ex/svc_oracle.ML (which can be mistaken as attached ML script on case-insensitive file-system);
 wenzelm parents: 
20813diff
changeset | 37 | let | 
| 
41c81e23c08d
removed Hoare/hoare.ML, Hoare/hoareAbort.ML, ex/svc_oracle.ML (which can be mistaken as attached ML script on case-insensitive file-system);
 wenzelm parents: 
20813diff
changeset | 38 | (*The oracle's result is given to the subgoal using compose_tac because | 
| 
41c81e23c08d
removed Hoare/hoare.ML, Hoare/hoareAbort.ML, ex/svc_oracle.ML (which can be mistaken as attached ML script on case-insensitive file-system);
 wenzelm parents: 
20813diff
changeset | 39 | its premises are matched against the assumptions rather than used | 
| 
41c81e23c08d
removed Hoare/hoare.ML, Hoare/hoareAbort.ML, ex/svc_oracle.ML (which can be mistaken as attached ML script on case-insensitive file-system);
 wenzelm parents: 
20813diff
changeset | 40 | to make subgoals. Therefore , abstraction must copy the parameters | 
| 
41c81e23c08d
removed Hoare/hoare.ML, Hoare/hoareAbort.ML, ex/svc_oracle.ML (which can be mistaken as attached ML script on case-insensitive file-system);
 wenzelm parents: 
20813diff
changeset | 41 | precisely and make them available to all generated Vars.*) | 
| 
41c81e23c08d
removed Hoare/hoare.ML, Hoare/hoareAbort.ML, ex/svc_oracle.ML (which can be mistaken as attached ML script on case-insensitive file-system);
 wenzelm parents: 
20813diff
changeset | 42 | val params = Term.strip_all_vars t | 
| 
41c81e23c08d
removed Hoare/hoare.ML, Hoare/hoareAbort.ML, ex/svc_oracle.ML (which can be mistaken as attached ML script on case-insensitive file-system);
 wenzelm parents: 
20813diff
changeset | 43 | and body = Term.strip_all_body t | 
| 
41c81e23c08d
removed Hoare/hoare.ML, Hoare/hoareAbort.ML, ex/svc_oracle.ML (which can be mistaken as attached ML script on case-insensitive file-system);
 wenzelm parents: 
20813diff
changeset | 44 | val Us = map #2 params | 
| 
41c81e23c08d
removed Hoare/hoare.ML, Hoare/hoareAbort.ML, ex/svc_oracle.ML (which can be mistaken as attached ML script on case-insensitive file-system);
 wenzelm parents: 
20813diff
changeset | 45 | val nPar = length params | 
| 32740 | 46 | val vname = Unsynchronized.ref "V_a" | 
| 47 | val pairs = Unsynchronized.ref ([] : (term*term) list) | |
| 24470 
41c81e23c08d
removed Hoare/hoare.ML, Hoare/hoareAbort.ML, ex/svc_oracle.ML (which can be mistaken as attached ML script on case-insensitive file-system);
 wenzelm parents: 
20813diff
changeset | 48 | fun insert t = | 
| 
41c81e23c08d
removed Hoare/hoare.ML, Hoare/hoareAbort.ML, ex/svc_oracle.ML (which can be mistaken as attached ML script on case-insensitive file-system);
 wenzelm parents: 
20813diff
changeset | 49 | let val T = fastype_of t | 
| 
41c81e23c08d
removed Hoare/hoare.ML, Hoare/hoareAbort.ML, ex/svc_oracle.ML (which can be mistaken as attached ML script on case-insensitive file-system);
 wenzelm parents: 
20813diff
changeset | 50 | val v = Logic.combound (Var ((!vname,0), Us--->T), 0, nPar) | 
| 
41c81e23c08d
removed Hoare/hoare.ML, Hoare/hoareAbort.ML, ex/svc_oracle.ML (which can be mistaken as attached ML script on case-insensitive file-system);
 wenzelm parents: 
20813diff
changeset | 51 | in vname := Symbol.bump_string (!vname); | 
| 
41c81e23c08d
removed Hoare/hoare.ML, Hoare/hoareAbort.ML, ex/svc_oracle.ML (which can be mistaken as attached ML script on case-insensitive file-system);
 wenzelm parents: 
20813diff
changeset | 52 | pairs := (t, v) :: !pairs; | 
| 
41c81e23c08d
removed Hoare/hoare.ML, Hoare/hoareAbort.ML, ex/svc_oracle.ML (which can be mistaken as attached ML script on case-insensitive file-system);
 wenzelm parents: 
20813diff
changeset | 53 | v | 
| 
41c81e23c08d
removed Hoare/hoare.ML, Hoare/hoareAbort.ML, ex/svc_oracle.ML (which can be mistaken as attached ML script on case-insensitive file-system);
 wenzelm parents: 
20813diff
changeset | 54 | end; | 
| 
41c81e23c08d
removed Hoare/hoare.ML, Hoare/hoareAbort.ML, ex/svc_oracle.ML (which can be mistaken as attached ML script on case-insensitive file-system);
 wenzelm parents: 
20813diff
changeset | 55 | fun replace t = | 
| 
41c81e23c08d
removed Hoare/hoare.ML, Hoare/hoareAbort.ML, ex/svc_oracle.ML (which can be mistaken as attached ML script on case-insensitive file-system);
 wenzelm parents: 
20813diff
changeset | 56 | case t of | 
| 
41c81e23c08d
removed Hoare/hoare.ML, Hoare/hoareAbort.ML, ex/svc_oracle.ML (which can be mistaken as attached ML script on case-insensitive file-system);
 wenzelm parents: 
20813diff
changeset | 57 | Free _ => t (*but not existing Vars, lest the names clash*) | 
| 
41c81e23c08d
removed Hoare/hoare.ML, Hoare/hoareAbort.ML, ex/svc_oracle.ML (which can be mistaken as attached ML script on case-insensitive file-system);
 wenzelm parents: 
20813diff
changeset | 58 | | Bound _ => t | 
| 
41c81e23c08d
removed Hoare/hoare.ML, Hoare/hoareAbort.ML, ex/svc_oracle.ML (which can be mistaken as attached ML script on case-insensitive file-system);
 wenzelm parents: 
20813diff
changeset | 59 | | _ => (case AList.lookup Pattern.aeconv (!pairs) t of | 
| 
41c81e23c08d
removed Hoare/hoare.ML, Hoare/hoareAbort.ML, ex/svc_oracle.ML (which can be mistaken as attached ML script on case-insensitive file-system);
 wenzelm parents: 
20813diff
changeset | 60 | SOME v => v | 
| 
41c81e23c08d
removed Hoare/hoare.ML, Hoare/hoareAbort.ML, ex/svc_oracle.ML (which can be mistaken as attached ML script on case-insensitive file-system);
 wenzelm parents: 
20813diff
changeset | 61 | | NONE => insert t) | 
| 
41c81e23c08d
removed Hoare/hoare.ML, Hoare/hoareAbort.ML, ex/svc_oracle.ML (which can be mistaken as attached ML script on case-insensitive file-system);
 wenzelm parents: 
20813diff
changeset | 62 | (*abstraction of a numeric literal*) | 
| 25929 
6bfef23e26be
avoiding direct references to numeral presentation
 haftmann parents: 
25919diff
changeset | 63 | fun lit t = if can HOLogic.dest_number t then t else replace t; | 
| 24470 
41c81e23c08d
removed Hoare/hoare.ML, Hoare/hoareAbort.ML, ex/svc_oracle.ML (which can be mistaken as attached ML script on case-insensitive file-system);
 wenzelm parents: 
20813diff
changeset | 64 | (*abstraction of a real/rational expression*) | 
| 35267 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35092diff
changeset | 65 |     fun rat ((c as Const(@{const_name Groups.plus}, _)) $ x $ y) = c $ (rat x) $ (rat y)
 | 
| 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35092diff
changeset | 66 |       | rat ((c as Const(@{const_name Groups.minus}, _)) $ x $ y) = c $ (rat x) $ (rat y)
 | 
| 35084 | 67 |       | rat ((c as Const(@{const_name Rings.divide}, _)) $ x $ y) = c $ (rat x) $ (rat y)
 | 
| 35267 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35092diff
changeset | 68 |       | rat ((c as Const(@{const_name Groups.times}, _)) $ x $ y) = c $ (rat x) $ (rat y)
 | 
| 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35092diff
changeset | 69 |       | rat ((c as Const(@{const_name Groups.uminus}, _)) $ x) = c $ (rat x)
 | 
| 24470 
41c81e23c08d
removed Hoare/hoare.ML, Hoare/hoareAbort.ML, ex/svc_oracle.ML (which can be mistaken as attached ML script on case-insensitive file-system);
 wenzelm parents: 
20813diff
changeset | 70 | | rat t = lit t | 
| 
41c81e23c08d
removed Hoare/hoare.ML, Hoare/hoareAbort.ML, ex/svc_oracle.ML (which can be mistaken as attached ML script on case-insensitive file-system);
 wenzelm parents: 
20813diff
changeset | 71 | (*abstraction of an integer expression: no div, mod*) | 
| 35267 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35092diff
changeset | 72 |     fun int ((c as Const(@{const_name Groups.plus}, _)) $ x $ y) = c $ (int x) $ (int y)
 | 
| 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35092diff
changeset | 73 |       | int ((c as Const(@{const_name Groups.minus}, _)) $ x $ y) = c $ (int x) $ (int y)
 | 
| 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35092diff
changeset | 74 |       | int ((c as Const(@{const_name Groups.times}, _)) $ x $ y) = c $ (int x) $ (int y)
 | 
| 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35092diff
changeset | 75 |       | int ((c as Const(@{const_name Groups.uminus}, _)) $ x) = c $ (int x)
 | 
| 24470 
41c81e23c08d
removed Hoare/hoare.ML, Hoare/hoareAbort.ML, ex/svc_oracle.ML (which can be mistaken as attached ML script on case-insensitive file-system);
 wenzelm parents: 
20813diff
changeset | 76 | | int t = lit t | 
| 
41c81e23c08d
removed Hoare/hoare.ML, Hoare/hoareAbort.ML, ex/svc_oracle.ML (which can be mistaken as attached ML script on case-insensitive file-system);
 wenzelm parents: 
20813diff
changeset | 77 | (*abstraction of a natural number expression: no minus*) | 
| 35267 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35092diff
changeset | 78 |     fun nat ((c as Const(@{const_name Groups.plus}, _)) $ x $ y) = c $ (nat x) $ (nat y)
 | 
| 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35092diff
changeset | 79 |       | nat ((c as Const(@{const_name Groups.times}, _)) $ x $ y) = c $ (nat x) $ (nat y)
 | 
| 24470 
41c81e23c08d
removed Hoare/hoare.ML, Hoare/hoareAbort.ML, ex/svc_oracle.ML (which can be mistaken as attached ML script on case-insensitive file-system);
 wenzelm parents: 
20813diff
changeset | 80 |       | nat ((c as Const(@{const_name Suc}, _)) $ x) = c $ (nat x)
 | 
| 
41c81e23c08d
removed Hoare/hoare.ML, Hoare/hoareAbort.ML, ex/svc_oracle.ML (which can be mistaken as attached ML script on case-insensitive file-system);
 wenzelm parents: 
20813diff
changeset | 81 | | nat t = lit t | 
| 
41c81e23c08d
removed Hoare/hoare.ML, Hoare/hoareAbort.ML, ex/svc_oracle.ML (which can be mistaken as attached ML script on case-insensitive file-system);
 wenzelm parents: 
20813diff
changeset | 82 | (*abstraction of a relation: =, <, <=*) | 
| 
41c81e23c08d
removed Hoare/hoare.ML, Hoare/hoareAbort.ML, ex/svc_oracle.ML (which can be mistaken as attached ML script on case-insensitive file-system);
 wenzelm parents: 
20813diff
changeset | 83 | fun rel (T, c $ x $ y) = | 
| 
41c81e23c08d
removed Hoare/hoare.ML, Hoare/hoareAbort.ML, ex/svc_oracle.ML (which can be mistaken as attached ML script on case-insensitive file-system);
 wenzelm parents: 
20813diff
changeset | 84 | if T = HOLogic.realT then c $ (rat x) $ (rat y) | 
| 
41c81e23c08d
removed Hoare/hoare.ML, Hoare/hoareAbort.ML, ex/svc_oracle.ML (which can be mistaken as attached ML script on case-insensitive file-system);
 wenzelm parents: 
20813diff
changeset | 85 | else if T = HOLogic.intT then c $ (int x) $ (int y) | 
| 
41c81e23c08d
removed Hoare/hoare.ML, Hoare/hoareAbort.ML, ex/svc_oracle.ML (which can be mistaken as attached ML script on case-insensitive file-system);
 wenzelm parents: 
20813diff
changeset | 86 | else if T = HOLogic.natT then c $ (nat x) $ (nat y) | 
| 
41c81e23c08d
removed Hoare/hoare.ML, Hoare/hoareAbort.ML, ex/svc_oracle.ML (which can be mistaken as attached ML script on case-insensitive file-system);
 wenzelm parents: 
20813diff
changeset | 87 | else if T = HOLogic.boolT then c $ (fm x) $ (fm y) | 
| 
41c81e23c08d
removed Hoare/hoare.ML, Hoare/hoareAbort.ML, ex/svc_oracle.ML (which can be mistaken as attached ML script on case-insensitive file-system);
 wenzelm parents: 
20813diff
changeset | 88 | else replace (c $ x $ y) (*non-numeric comparison*) | 
| 
41c81e23c08d
removed Hoare/hoare.ML, Hoare/hoareAbort.ML, ex/svc_oracle.ML (which can be mistaken as attached ML script on case-insensitive file-system);
 wenzelm parents: 
20813diff
changeset | 89 | (*abstraction of a formula*) | 
| 38795 
848be46708dc
formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
 haftmann parents: 
38786diff
changeset | 90 |     and fm ((c as Const(@{const_name HOL.conj}, _)) $ p $ q) = c $ (fm p) $ (fm q)
 | 
| 
848be46708dc
formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
 haftmann parents: 
38786diff
changeset | 91 |       | fm ((c as Const(@{const_name HOL.disj}, _)) $ p $ q) = c $ (fm p) $ (fm q)
 | 
| 38786 
e46e7a9cb622
formerly unnamed infix impliciation now named HOL.implies
 haftmann parents: 
38558diff
changeset | 92 |       | fm ((c as Const(@{const_name HOL.implies}, _)) $ p $ q) = c $ (fm p) $ (fm q)
 | 
| 38558 | 93 |       | fm ((c as Const(@{const_name Not}, _)) $ p) = c $ (fm p)
 | 
| 94 |       | fm ((c as Const(@{const_name True}, _))) = c
 | |
| 95 |       | fm ((c as Const(@{const_name False}, _))) = c
 | |
| 38864 
4abe644fcea5
formerly unnamed infix equality now named HOL.eq
 haftmann parents: 
38795diff
changeset | 96 |       | fm (t as Const(@{const_name HOL.eq},  Type ("fun", [T,_])) $ _ $ _) = rel (T, t)
 | 
| 35092 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 haftmann parents: 
35084diff
changeset | 97 |       | fm (t as Const(@{const_name Orderings.less},  Type ("fun", [T,_])) $ _ $ _) = rel (T, t)
 | 
| 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 haftmann parents: 
35084diff
changeset | 98 |       | fm (t as Const(@{const_name Orderings.less_eq}, Type ("fun", [T,_])) $ _ $ _) = rel (T, t)
 | 
| 24470 
41c81e23c08d
removed Hoare/hoare.ML, Hoare/hoareAbort.ML, ex/svc_oracle.ML (which can be mistaken as attached ML script on case-insensitive file-system);
 wenzelm parents: 
20813diff
changeset | 99 | | fm t = replace t | 
| 
41c81e23c08d
removed Hoare/hoare.ML, Hoare/hoareAbort.ML, ex/svc_oracle.ML (which can be mistaken as attached ML script on case-insensitive file-system);
 wenzelm parents: 
20813diff
changeset | 100 | (*entry point, and abstraction of a meta-formula*) | 
| 38558 | 101 |     fun mt ((c as Const(@{const_name Trueprop}, _)) $ p) = c $ (fm p)
 | 
| 24470 
41c81e23c08d
removed Hoare/hoare.ML, Hoare/hoareAbort.ML, ex/svc_oracle.ML (which can be mistaken as attached ML script on case-insensitive file-system);
 wenzelm parents: 
20813diff
changeset | 102 |       | mt ((c as Const("==>", _)) $ p $ q)  = c $ (mt p) $ (mt q)
 | 
| 
41c81e23c08d
removed Hoare/hoare.ML, Hoare/hoareAbort.ML, ex/svc_oracle.ML (which can be mistaken as attached ML script on case-insensitive file-system);
 wenzelm parents: 
20813diff
changeset | 103 | | mt t = fm t (*it might be a formula*) | 
| 
41c81e23c08d
removed Hoare/hoare.ML, Hoare/hoareAbort.ML, ex/svc_oracle.ML (which can be mistaken as attached ML script on case-insensitive file-system);
 wenzelm parents: 
20813diff
changeset | 104 | in (list_all (params, mt body), !pairs) end; | 
| 
41c81e23c08d
removed Hoare/hoare.ML, Hoare/hoareAbort.ML, ex/svc_oracle.ML (which can be mistaken as attached ML script on case-insensitive file-system);
 wenzelm parents: 
20813diff
changeset | 105 | |
| 
41c81e23c08d
removed Hoare/hoare.ML, Hoare/hoareAbort.ML, ex/svc_oracle.ML (which can be mistaken as attached ML script on case-insensitive file-system);
 wenzelm parents: 
20813diff
changeset | 106 | |
| 
41c81e23c08d
removed Hoare/hoare.ML, Hoare/hoareAbort.ML, ex/svc_oracle.ML (which can be mistaken as attached ML script on case-insensitive file-system);
 wenzelm parents: 
20813diff
changeset | 107 | (*Present the entire subgoal to the oracle, assumptions and all, but possibly | 
| 
41c81e23c08d
removed Hoare/hoare.ML, Hoare/hoareAbort.ML, ex/svc_oracle.ML (which can be mistaken as attached ML script on case-insensitive file-system);
 wenzelm parents: 
20813diff
changeset | 108 | abstracted. Use via compose_tac, which performs no lifting but will | 
| 
41c81e23c08d
removed Hoare/hoare.ML, Hoare/hoareAbort.ML, ex/svc_oracle.ML (which can be mistaken as attached ML script on case-insensitive file-system);
 wenzelm parents: 
20813diff
changeset | 109 | instantiate variables.*) | 
| 
41c81e23c08d
removed Hoare/hoare.ML, Hoare/hoareAbort.ML, ex/svc_oracle.ML (which can be mistaken as attached ML script on case-insensitive file-system);
 wenzelm parents: 
20813diff
changeset | 110 | |
| 28290 | 111 | val svc_tac = CSUBGOAL (fn (ct, i) => | 
| 24470 
41c81e23c08d
removed Hoare/hoare.ML, Hoare/hoareAbort.ML, ex/svc_oracle.ML (which can be mistaken as attached ML script on case-insensitive file-system);
 wenzelm parents: 
20813diff
changeset | 112 | let | 
| 28290 | 113 | val thy = Thm.theory_of_cterm ct; | 
| 114 | val (abs_goal, _) = svc_abstract (Thm.term_of ct); | |
| 115 | val th = svc_oracle (Thm.cterm_of thy abs_goal); | |
| 116 | in compose_tac (false, th, 0) i end | |
| 117 | handle TERM _ => no_tac); | |
| 24470 
41c81e23c08d
removed Hoare/hoare.ML, Hoare/hoareAbort.ML, ex/svc_oracle.ML (which can be mistaken as attached ML script on case-insensitive file-system);
 wenzelm parents: 
20813diff
changeset | 118 | *} | 
| 20813 | 119 | |
| 12869 | 120 | end |