author | obua |
Sat, 27 Oct 2007 18:37:32 +0200 | |
changeset 25217 | 3224db6415ae |
parent 23663 | 84b5c89b8b49 |
child 26957 | e3f04fdd994d |
permissions | -rw-r--r-- |
23173
51179ca0c429
proper theory setup for compute oracle (based on CPure);
wenzelm
parents:
diff
changeset
|
1 |
(* Title: Tools/Compute_Oracle/Compute_Oracle.thy |
51179ca0c429
proper theory setup for compute oracle (based on CPure);
wenzelm
parents:
diff
changeset
|
2 |
ID: $Id$ |
51179ca0c429
proper theory setup for compute oracle (based on CPure);
wenzelm
parents:
diff
changeset
|
3 |
Author: Steven Obua, TU Munich |
51179ca0c429
proper theory setup for compute oracle (based on CPure);
wenzelm
parents:
diff
changeset
|
4 |
|
51179ca0c429
proper theory setup for compute oracle (based on CPure);
wenzelm
parents:
diff
changeset
|
5 |
Steven Obua's evaluator. |
51179ca0c429
proper theory setup for compute oracle (based on CPure);
wenzelm
parents:
diff
changeset
|
6 |
*) |
51179ca0c429
proper theory setup for compute oracle (based on CPure);
wenzelm
parents:
diff
changeset
|
7 |
|
23663 | 8 |
theory Compute_Oracle imports CPure |
25217 | 9 |
uses "am.ML" "am_compiler.ML" "am_interpreter.ML" "am_ghc.ML" "am_sml.ML" "report.ML" "compute.ML" "linker.ML" |
23173
51179ca0c429
proper theory setup for compute oracle (based on CPure);
wenzelm
parents:
diff
changeset
|
10 |
begin |
51179ca0c429
proper theory setup for compute oracle (based on CPure);
wenzelm
parents:
diff
changeset
|
11 |
|
25217 | 12 |
setup {* Compute.setup_compute; *} |
23173
51179ca0c429
proper theory setup for compute oracle (based on CPure);
wenzelm
parents:
diff
changeset
|
13 |
|
23663 | 14 |
end |