| author | wenzelm | 
| Thu, 22 Mar 2012 15:41:49 +0100 | |
| changeset 47081 | 5e70b457b704 | 
| parent 44064 | 5bce8ff0d9ae | 
| child 51940 | 958d439b3013 | 
| permissions | -rw-r--r-- | 
| 24584 | 1 | (* Title: HOL/ex/svc_funcs.ML | 
| 12869 | 2 | Author: Lawrence C Paulson | 
| 3 | Copyright 1999 University of Cambridge | |
| 4 | ||
| 29276 | 5 | Translation functions for the interface to SVC. | 
| 12869 | 6 | |
| 19336 
fb5e19d26d5e
removed some illegal characters: they were crashing SML/NJ
 paulson parents: 
19277diff
changeset | 7 | Based upon the work of Soren T. Heilmann | 
| 12869 | 8 | |
| 9 | Integers and naturals are translated as follows: | |
| 10 | In a positive context, replace x<y by x+1<=y | |
| 11 | In a negative context, replace x<=y by x<y+1 | |
| 12 | In a negative context, replace x=y by x<y+1 & y<x+1 | |
| 13 | Biconditionals (if-and-only-iff) are expanded if they require such translations | |
| 14 | in either operand. | |
| 15 | ||
| 16 | For each variable of type nat, an assumption is added that it is non-negative. | |
| 17 | *) | |
| 18 | ||
| 19 | structure Svc = | |
| 20 | struct | |
| 32740 | 21 | val trace = Unsynchronized.ref false; | 
| 12869 | 22 | |
| 23 | datatype expr = | |
| 24 | Buildin of string * expr list | |
| 25 | | Interp of string * expr list | |
| 26 | | UnInterp of string * expr list | |
| 16836 | 27 | | FalseExpr | 
| 12869 | 28 | | TrueExpr | 
| 24630 
351a308ab58d
simplified type int (eliminated IntInf.int, integer);
 wenzelm parents: 
24584diff
changeset | 29 | | Int of int | 
| 
351a308ab58d
simplified type int (eliminated IntInf.int, integer);
 wenzelm parents: 
24584diff
changeset | 30 | | Rat of int * int; | 
| 16836 | 31 | |
| 12869 | 32 | fun is_intnat T = T = HOLogic.intT orelse T = HOLogic.natT; | 
| 16836 | 33 | |
| 12869 | 34 | fun is_numeric T = is_intnat T orelse T = HOLogic.realT; | 
| 16836 | 35 | |
| 12869 | 36 | fun is_numeric_op T = is_numeric (domain_type T); | 
| 37 | ||
| 38 | fun toString t = | |
| 16836 | 39 | let fun ue (Buildin(s, l)) = | 
| 40 |              "(" ^ s ^ (Library.foldl (fn (a, b) => a ^ " " ^ (ue b)) ("", l)) ^ ") "
 | |
| 41 | | ue (Interp(s, l)) = | |
| 42 |              "{" ^ s ^ (Library.foldl (fn (a, b) => a ^ " " ^ (ue b)) ("", l)) ^ "} "
 | |
| 43 | | ue (UnInterp(s, l)) = | |
| 44 |              "(" ^ s ^ (Library.foldl (fn (a, b) => a ^ " " ^ (ue b)) ("", l)) ^ ") "
 | |
| 45 | | ue (FalseExpr) = "FALSE " | |
| 46 | | ue (TrueExpr) = "TRUE " | |
| 24630 
351a308ab58d
simplified type int (eliminated IntInf.int, integer);
 wenzelm parents: 
24584diff
changeset | 47 | | ue (Int i) = signed_string_of_int i ^ " " | 
| 
351a308ab58d
simplified type int (eliminated IntInf.int, integer);
 wenzelm parents: 
24584diff
changeset | 48 | | ue (Rat(i, j)) = signed_string_of_int i ^ "|" ^ signed_string_of_int j ^ " " | 
| 12869 | 49 | in | 
| 16836 | 50 | ue t | 
| 12869 | 51 | end; | 
| 52 | ||
| 16836 | 53 | fun valid e = | 
| 54 | let val svc_home = getenv "SVC_HOME" | |
| 12869 | 55 | val svc_machine = getenv "SVC_MACHINE" | 
| 56 | val check_valid = if svc_home = "" | |
| 16836 | 57 | then error "Environment variable SVC_HOME not set" | 
| 58 | else if svc_machine = "" | |
| 59 | then error "Environment variable SVC_MACHINE not set" | |
| 60 | else svc_home ^ "/" ^ svc_machine ^ "/bin/check_valid" | |
| 12869 | 61 | val svc_input = toString e | 
| 62 |       val _ = if !trace then tracing ("Calling SVC:\n" ^ svc_input) else ()
 | |
| 63 | val svc_input_file = File.tmp_path (Path.basic "SVM_in"); | |
| 64 | val svc_output_file = File.tmp_path (Path.basic "SVM_out"); | |
| 35010 
d6e492cea6e4
renamed system/system_out to bash/bash_output -- to emphasized that this is really GNU bash, not some undefined POSIX sh;
 wenzelm parents: 
34974diff
changeset | 65 | val _ = File.write svc_input_file svc_input; | 
| 
d6e492cea6e4
renamed system/system_out to bash/bash_output -- to emphasized that this is really GNU bash, not some undefined POSIX sh;
 wenzelm parents: 
34974diff
changeset | 66 | val _ = | 
| 43850 
7f2cbc713344
moved bash operations to Isabelle_System (cf. Scala version);
 wenzelm parents: 
42364diff
changeset | 67 | Isabelle_System.bash_output (check_valid ^ " -dump-result " ^ | 
| 35010 
d6e492cea6e4
renamed system/system_out to bash/bash_output -- to emphasized that this is really GNU bash, not some undefined POSIX sh;
 wenzelm parents: 
34974diff
changeset | 68 | File.shell_path svc_output_file ^ " " ^ File.shell_path svc_input_file ^ | 
| 
d6e492cea6e4
renamed system/system_out to bash/bash_output -- to emphasized that this is really GNU bash, not some undefined POSIX sh;
 wenzelm parents: 
34974diff
changeset | 69 | ">/dev/null 2>&1") | 
| 12869 | 70 | val svc_output = | 
| 21395 
f34ac19659ae
moved some fundamental concepts to General/basics.ML;
 wenzelm parents: 
21078diff
changeset | 71 | (case try File.read svc_output_file of | 
| 15531 | 72 | SOME out => out | 
| 73 | | NONE => error "SVC returned no output"); | |
| 12869 | 74 | in | 
| 75 |       if ! trace then tracing ("SVC Returns:\n" ^ svc_output)
 | |
| 76 | else (File.rm svc_input_file; File.rm svc_output_file); | |
| 77 | String.isPrefix "VALID" svc_output | |
| 78 | end | |
| 79 | ||
| 16836 | 80 |  fun fail t = raise TERM ("SVC oracle", [t]);
 | 
| 12869 | 81 | |
| 82 | fun apply c args = | |
| 83 | let val (ts, bs) = ListPair.unzip args | |
| 84 | in (list_comb(c,ts), exists I bs) end; | |
| 85 | ||
| 86 | (*Determining whether the biconditionals must be unfolded: if there are | |
| 87 | int or nat comparisons below*) | |
| 88 | val iff_tag = | |
| 89 | let fun tag t = | |
| 16836 | 90 | let val (c,ts) = strip_comb t | 
| 91 | in case c of | |
| 38795 
848be46708dc
formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
 haftmann parents: 
38786diff
changeset | 92 |              Const(@{const_name HOL.conj}, _)   => apply c (map tag ts)
 | 
| 
848be46708dc
formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
 haftmann parents: 
38786diff
changeset | 93 |            | Const(@{const_name HOL.disj}, _)   => apply c (map tag ts)
 | 
| 38786 
e46e7a9cb622
formerly unnamed infix impliciation now named HOL.implies
 haftmann parents: 
38558diff
changeset | 94 |            | Const(@{const_name HOL.implies}, _) => apply c (map tag ts)
 | 
| 38558 | 95 |            | Const(@{const_name Not}, _)    => apply c (map tag ts)
 | 
| 96 |            | Const(@{const_name True}, _)   => (c, false)
 | |
| 97 |            | Const(@{const_name False}, _)  => (c, false)
 | |
| 38864 
4abe644fcea5
formerly unnamed infix equality now named HOL.eq
 haftmann parents: 
38795diff
changeset | 98 |            | Const(@{const_name HOL.eq}, Type ("fun", [T,_])) =>
 | 
| 16836 | 99 | if T = HOLogic.boolT then | 
| 100 | (*biconditional: with int/nat comparisons below?*) | |
| 101 | let val [t1,t2] = ts | |
| 102 | val (u1,b1) = tag t1 | |
| 103 | and (u2,b2) = tag t2 | |
| 104 | val cname = if b1 orelse b2 then "unfold" else "keep" | |
| 105 | in | |
| 106 |                         (Const ("SVC_Oracle.iff_" ^ cname, dummyT) $ u1 $ u2,
 | |
| 107 | b1 orelse b2) | |
| 108 | end | |
| 109 | else (*might be numeric equality*) (t, is_intnat T) | |
| 35092 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 haftmann parents: 
35084diff
changeset | 110 |            | Const(@{const_name Orderings.less}, Type ("fun", [T,_]))  => (t, is_intnat T)
 | 
| 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 haftmann parents: 
35084diff
changeset | 111 |            | Const(@{const_name Orderings.less_eq}, Type ("fun", [T,_])) => (t, is_intnat T)
 | 
| 16836 | 112 | | _ => (t, false) | 
| 113 | end | |
| 12869 | 114 | in #1 o tag end; | 
| 115 | ||
| 116 | (*Map expression e to 0<=a --> e, where "a" is the name of a nat variable*) | |
| 21078 | 117 | fun add_nat_var a e = | 
| 12869 | 118 |      Buildin("=>", [Buildin("<=", [Int 0, UnInterp (a, [])]),
 | 
| 16836 | 119 | e]); | 
| 12869 | 120 | |
| 121 | fun param_string [] = "" | |
| 122 | | param_string is = "_" ^ space_implode "_" (map string_of_int is) | |
| 123 | ||
| 124 | (*Translate an Isabelle formula into an SVC expression | |
| 125 | pos ["positive"]: true if an assumption, false if a goal*) | |
| 126 | fun expr_of pos t = | |
| 127 | let | |
| 29276 | 128 | val params = rev (Term.rename_wrt_term t (Term.strip_all_vars t)) | 
| 12869 | 129 | and body = Term.strip_all_body t | 
| 32740 | 130 | val nat_vars = Unsynchronized.ref ([] : string list) | 
| 12869 | 131 | (*translation of a variable: record all natural numbers*) | 
| 132 | fun trans_var (a,T,is) = | |
| 20853 | 133 | (if T = HOLogic.natT then nat_vars := (insert (op =) a (!nat_vars)) | 
| 16836 | 134 | else (); | 
| 12869 | 135 | UnInterp (a ^ param_string is, [])) | 
| 136 | (*A variable, perhaps applied to a series of parameters*) | |
| 137 |     fun var (Free(a,T), is)      = trans_var ("F_" ^ a, T, is)
 | |
| 138 | | var (Var((a, 0), T), is) = trans_var (a, T, is) | |
| 16836 | 139 | | var (Bound i, is) = | 
| 42364 | 140 | let val (a,T) = nth params i | 
| 16836 | 141 |           in  trans_var ("B_" ^ a, T, is)  end
 | 
| 12869 | 142 | | var (t $ Bound i, is) = var(t,i::is) | 
| 143 | (*removing a parameter from a Var: the bound var index will | |
| 144 | become part of the Var's name*) | |
| 16836 | 145 | | var (t,_) = fail t; | 
| 12869 | 146 | (*translation of a literal*) | 
| 21820 
2f2b6a965ccc
introduced mk/dest_numeral/number for mk/dest_binum etc.
 haftmann parents: 
21395diff
changeset | 147 | val lit = snd o HOLogic.dest_number; | 
| 12869 | 148 | (*translation of a literal expression [no variables]*) | 
| 35267 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35092diff
changeset | 149 |     fun litExp (Const(@{const_name Groups.plus}, T) $ x $ y) =
 | 
| 16836 | 150 | if is_numeric_op T then (litExp x) + (litExp y) | 
| 151 | else fail t | |
| 35267 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35092diff
changeset | 152 |       | litExp (Const(@{const_name Groups.minus}, T) $ x $ y) =
 | 
| 16836 | 153 | if is_numeric_op T then (litExp x) - (litExp y) | 
| 154 | else fail t | |
| 35267 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35092diff
changeset | 155 |       | litExp (Const(@{const_name Groups.times}, T) $ x $ y) =
 | 
| 16836 | 156 | if is_numeric_op T then (litExp x) * (litExp y) | 
| 157 | else fail t | |
| 35267 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35092diff
changeset | 158 |       | litExp (Const(@{const_name Groups.uminus}, T) $ x)   =
 | 
| 16836 | 159 | if is_numeric_op T then ~(litExp x) | 
| 160 | else fail t | |
| 161 | | litExp t = lit t | |
| 162 | handle Match => fail t | |
| 12869 | 163 | (*translation of a real/rational expression*) | 
| 164 |     fun suc t = Interp("+", [Int 1, t])
 | |
| 22997 | 165 |     fun tm (Const(@{const_name Suc}, T) $ x) = suc (tm x)
 | 
| 35267 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35092diff
changeset | 166 |       | tm (Const(@{const_name Groups.plus}, T) $ x $ y) =
 | 
| 16836 | 167 |           if is_numeric_op T then Interp("+", [tm x, tm y])
 | 
| 168 | else fail t | |
| 35267 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35092diff
changeset | 169 |       | tm (Const(@{const_name Groups.minus}, T) $ x $ y) =
 | 
| 16836 | 170 | if is_numeric_op T then | 
| 171 |               Interp("+", [tm x, Interp("*", [Int ~1, tm y])])
 | |
| 172 | else fail t | |
| 35267 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35092diff
changeset | 173 |       | tm (Const(@{const_name Groups.times}, T) $ x $ y) =
 | 
| 16836 | 174 |           if is_numeric_op T then Interp("*", [tm x, tm y])
 | 
| 175 | else fail t | |
| 44064 
5bce8ff0d9ae
moved division ring stuff from Rings.thy to Fields.thy
 huffman parents: 
43850diff
changeset | 176 |       | tm (Const(@{const_name Fields.inverse}, T) $ x) =
 | 
| 16836 | 177 | if domain_type T = HOLogic.realT then | 
| 178 | Rat(1, litExp x) | |
| 179 | else fail t | |
| 35267 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35092diff
changeset | 180 |       | tm (Const(@{const_name Groups.uminus}, T) $ x) =
 | 
| 16836 | 181 |           if is_numeric_op T then Interp("*", [Int ~1, tm x])
 | 
| 182 | else fail t | |
| 183 | | tm t = Int (lit t) | |
| 184 | handle Match => var (t,[]) | |
| 12869 | 185 | (*translation of a formula*) | 
| 38795 
848be46708dc
formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
 haftmann parents: 
38786diff
changeset | 186 |     and fm pos (Const(@{const_name HOL.conj}, _) $ p $ q) =
 | 
| 16836 | 187 |             Buildin("AND", [fm pos p, fm pos q])
 | 
| 38795 
848be46708dc
formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
 haftmann parents: 
38786diff
changeset | 188 |       | fm pos (Const(@{const_name HOL.disj}, _) $ p $ q) =
 | 
| 16836 | 189 |             Buildin("OR", [fm pos p, fm pos q])
 | 
| 38786 
e46e7a9cb622
formerly unnamed infix impliciation now named HOL.implies
 haftmann parents: 
38558diff
changeset | 190 |       | fm pos (Const(@{const_name HOL.implies}, _) $ p $ q) =
 | 
| 16836 | 191 |             Buildin("=>", [fm (not pos) p, fm pos q])
 | 
| 38558 | 192 |       | fm pos (Const(@{const_name Not}, _) $ p) =
 | 
| 16836 | 193 |             Buildin("NOT", [fm (not pos) p])
 | 
| 38558 | 194 |       | fm pos (Const(@{const_name True}, _)) = TrueExpr
 | 
| 195 |       | fm pos (Const(@{const_name False}, _)) = FalseExpr
 | |
| 16836 | 196 |       | fm pos (Const("SVC_Oracle.iff_keep", _) $ p $ q) =
 | 
| 197 | (*polarity doesn't matter*) | |
| 198 |             Buildin("=", [fm pos p, fm pos q])
 | |
| 199 |       | fm pos (Const("SVC_Oracle.iff_unfold", _) $ p $ q) =
 | |
| 200 |             Buildin("AND",   (*unfolding uses both polarities*)
 | |
| 201 |                          [Buildin("=>", [fm (not pos) p, fm pos q]),
 | |
| 202 |                           Buildin("=>", [fm (not pos) q, fm pos p])])
 | |
| 38864 
4abe644fcea5
formerly unnamed infix equality now named HOL.eq
 haftmann parents: 
38795diff
changeset | 203 |       | fm pos (t as Const(@{const_name HOL.eq}, Type ("fun", [T,_])) $ x $ y) =
 | 
| 16836 | 204 | let val tx = tm x and ty = tm y | 
| 205 | in if pos orelse T = HOLogic.realT then | |
| 206 |                        Buildin("=", [tx, ty])
 | |
| 207 | else if is_intnat T then | |
| 208 |                        Buildin("AND",
 | |
| 209 |                                     [Buildin("<", [tx, suc ty]),
 | |
| 210 |                                      Buildin("<", [ty, suc tx])])
 | |
| 211 | else fail t | |
| 212 | end | |
| 213 | (*inequalities: possible types are nat, int, real*) | |
| 35092 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 haftmann parents: 
35084diff
changeset | 214 |       | fm pos (t as Const(@{const_name Orderings.less},  Type ("fun", [T,_])) $ x $ y) =
 | 
| 16836 | 215 | if not pos orelse T = HOLogic.realT then | 
| 216 |                 Buildin("<", [tm x, tm y])
 | |
| 217 | else if is_intnat T then | |
| 218 |                 Buildin("<=", [suc (tm x), tm y])
 | |
| 219 | else fail t | |
| 35092 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 haftmann parents: 
35084diff
changeset | 220 |       | fm pos (t as Const(@{const_name Orderings.less_eq},  Type ("fun", [T,_])) $ x $ y) =
 | 
| 16836 | 221 | if pos orelse T = HOLogic.realT then | 
| 222 |                 Buildin("<=", [tm x, tm y])
 | |
| 223 | else if is_intnat T then | |
| 224 |                 Buildin("<", [tm x, suc (tm y)])
 | |
| 225 | else fail t | |
| 12869 | 226 | | fm pos t = var(t,[]); | 
| 227 | (*entry point, and translation of a meta-formula*) | |
| 38558 | 228 |       fun mt pos ((c as Const(@{const_name Trueprop}, _)) $ p) = fm pos (iff_tag p)
 | 
| 16836 | 229 |         | mt pos ((c as Const("==>", _)) $ p $ q) =
 | 
| 230 |             Buildin("=>", [mt (not pos) p, mt pos q])
 | |
| 231 | | mt pos t = fm pos (iff_tag t) (*it might be a formula*) | |
| 12869 | 232 | |
| 233 | val body_e = mt pos body (*evaluate now to assign into !nat_vars*) | |
| 16836 | 234 | in | 
| 21078 | 235 | fold_rev add_nat_var (!nat_vars) body_e | 
| 12869 | 236 | end; | 
| 237 | ||
| 238 | ||
| 28290 | 239 | (*The oracle proves the given formula, if possible*) | 
| 240 | fun oracle ct = | |
| 241 | let | |
| 242 | val thy = Thm.theory_of_cterm ct; | |
| 243 | val t = Thm.term_of ct; | |
| 244 | val _ = | |
| 245 |         if ! trace then tracing ("SVC oracle: problem is\n" ^ Syntax.string_of_term_global thy t)
 | |
| 246 | else (); | |
| 247 | in if valid (expr_of false t) then ct else fail t end; | |
| 12869 | 248 | |
| 249 | end; |