24281
|
1 |
(* Title: HOL/Library/Eval_Witness.thy
|
|
2 |
ID: $Id$
|
|
3 |
Author: Alexander Krauss, TU Muenchen
|
|
4 |
|
|
5 |
*)
|
|
6 |
|
|
7 |
header {* Evaluation Oracle with ML witnesses *}
|
|
8 |
|
|
9 |
theory Eval_Witness
|
|
10 |
imports Main
|
|
11 |
begin
|
|
12 |
|
|
13 |
text {*
|
|
14 |
We provide an oracle method similar to "eval", but with the
|
|
15 |
possibility to provide ML values as witnesses for existential
|
|
16 |
statements.
|
|
17 |
|
|
18 |
Our oracle can prove statements of the form @{term "EX x. P x"}
|
|
19 |
where @{term "P"} is an executable predicate that can be compiled to
|
|
20 |
ML. The oracle generates code for @{term "P"} and applies
|
|
21 |
it to a user-specified ML value. If the evaluation
|
|
22 |
returns true, this is effectively a proof of @{term "EX x. P x"}.
|
|
23 |
|
|
24 |
However, this is only sound if for every ML value of the given type
|
|
25 |
there exists a corresponding HOL value, which could be used in an
|
|
26 |
explicit proof. Unfortunately this is not true for function types,
|
|
27 |
since ML functions are not equivalent to the pure HOL
|
|
28 |
functions. Thus, the oracle can only be used on first-order
|
|
29 |
types.
|
|
30 |
|
|
31 |
We define a type class to mark types that can be safely used
|
|
32 |
with the oracle.
|
|
33 |
*}
|
|
34 |
|
|
35 |
class ml_equiv = type
|
|
36 |
|
|
37 |
text {*
|
|
38 |
Instances of @{text "ml_equiv"} should only be declared for those types,
|
|
39 |
where the universe of ML values coincides with the HOL values.
|
|
40 |
|
|
41 |
Since this is essentially a statement about ML, there is no
|
|
42 |
logical characterization.
|
|
43 |
*}
|
|
44 |
|
|
45 |
instance nat :: ml_equiv .. (* Attention: This conflicts with the "EfficientNat" theory *)
|
|
46 |
instance bool :: ml_equiv ..
|
|
47 |
instance list :: (ml_equiv) ml_equiv ..
|
|
48 |
|
|
49 |
oracle eval_witness_oracle ("term * string list") = {* fn thy => fn (goal, ws) =>
|
|
50 |
let
|
|
51 |
fun check_type T =
|
|
52 |
if Sorts.of_sort (Sign.classes_of thy) (T, ["Eval_Witness.ml_equiv"])
|
|
53 |
then T
|
|
54 |
else error ("Type " ^ quote (Sign.string_of_typ thy T) ^ " not allowed for ML witnesses")
|
|
55 |
|
|
56 |
fun dest_exs 0 t = t
|
|
57 |
| dest_exs n (Const ("Ex", _) $ Abs (v,T,b)) =
|
|
58 |
Abs (v, check_type T, dest_exs (n - 1) b)
|
|
59 |
| dest_exs _ _ = sys_error "dest_exs";
|
|
60 |
|
|
61 |
val ct = Thm.cterm_of thy (dest_exs (length ws) (HOLogic.dest_Trueprop goal))
|
|
62 |
in
|
|
63 |
if CodePackage.satisfies thy ct ws
|
|
64 |
then goal
|
|
65 |
else HOLogic.Trueprop $ (HOLogic.true_const) (*dummy*)
|
|
66 |
end
|
|
67 |
*}
|
|
68 |
|
|
69 |
|
|
70 |
method_setup eval_witness = {*
|
|
71 |
let
|
|
72 |
|
|
73 |
fun eval_tac ws thy =
|
|
74 |
SUBGOAL (fn (t, i) => rtac (eval_witness_oracle thy (t, ws)) i)
|
|
75 |
|
|
76 |
in
|
|
77 |
Method.simple_args (Scan.repeat Args.name) (fn ws => fn ctxt =>
|
|
78 |
Method.SIMPLE_METHOD' (eval_tac ws (ProofContext.theory_of ctxt)))
|
|
79 |
end
|
|
80 |
*} "Evaluation with ML witnesses"
|
|
81 |
|
|
82 |
|
|
83 |
subsection {* Toy Examples *}
|
|
84 |
|
|
85 |
text {*
|
|
86 |
Note that we must use the generated data structure for the
|
|
87 |
naturals, since ML integers are different.
|
|
88 |
*}
|
|
89 |
|
|
90 |
lemma "\<exists>n::nat. n = 1"
|
|
91 |
apply (eval_witness "Isabelle_Eval.Suc Isabelle_Eval.Zero_nat")
|
|
92 |
done
|
|
93 |
|
|
94 |
text {*
|
|
95 |
Since polymorphism is not allowed, we must specify the
|
|
96 |
type explicitly:
|
|
97 |
*}
|
|
98 |
|
|
99 |
lemma "\<exists>l. length (l::bool list) = 3"
|
|
100 |
apply (eval_witness "[true,true,true]")
|
|
101 |
done
|
|
102 |
|
|
103 |
text {* Multiple witnesses *}
|
|
104 |
|
|
105 |
lemma "\<exists>k l. length (k::bool list) = length (l::bool list)"
|
|
106 |
apply (eval_witness "[]" "[]")
|
|
107 |
done
|
|
108 |
|
|
109 |
|
|
110 |
subsection {* Discussion *}
|
|
111 |
|
|
112 |
subsubsection {* Conflicts *}
|
|
113 |
|
|
114 |
text {*
|
|
115 |
This theory conflicts with EfficientNat, since the @{text ml_equiv} instance
|
|
116 |
for natural numbers is not valid when they are mapped to ML
|
|
117 |
integers. With that theory loaded, we could use our oracle to prove
|
|
118 |
@{term "\<exists>n. n < 0"} by providing @{text "~1"} as a witness.
|
|
119 |
|
|
120 |
This shows that @{text ml_equiv} declarations have to be used with care,
|
|
121 |
taking the configuration of the code generator into account.
|
|
122 |
*}
|
|
123 |
|
|
124 |
subsubsection {* Haskell *}
|
|
125 |
|
|
126 |
text {*
|
|
127 |
If we were able to run generated Haskell code, the situation would
|
|
128 |
be much nicer, since Haskell functions are pure and could be used as
|
|
129 |
witnesses for HOL functions: Although Haskell functions are partial,
|
|
130 |
we know that if the evaluation terminates, they are ``sufficiently
|
|
131 |
defined'' and could be completed arbitrarily to a total (HOL) function.
|
|
132 |
|
|
133 |
This would allow us to provide access to very efficient data
|
|
134 |
structures via lookup functions coded in Haskell and provided to HOL
|
|
135 |
as witnesses.
|
|
136 |
*}
|
|
137 |
|
|
138 |
end |