author | kleing |
Mon, 21 Jun 2004 10:25:57 +0200 | |
changeset 14981 | e73f8140af78 |
parent 14806 | b42ad431cbae |
child 15063 | a43d771c18ac |
permissions | -rw-r--r-- |
10519 | 1 |
(* Title: HOL/Main.thy |
2 |
ID: $Id$ |
|
14350 | 3 |
Author: Stefan Berghofer, Tobias Nipkow, Tjark Weber, Markus Wenzel (TU Muenchen) |
10519 | 4 |
*) |
9619 | 5 |
|
12024 | 6 |
header {* Main HOL *} |
7 |
||
14443
75910c7557c5
generic theorems about exponentials; general tidying up
paulson
parents:
14350
diff
changeset
|
8 |
theory Main = Map + Infinite_Set + Extraction + Refute: |
9650
6f0b89f2a1f9
Main now new-style theory; added Main.ML for compatibility;
wenzelm
parents:
9619
diff
changeset
|
9 |
|
12024 | 10 |
text {* |
11 |
Theory @{text Main} includes everything. Note that theory @{text |
|
12 |
PreList} already includes most HOL theories. |
|
13 |
*} |
|
14 |
||
15 |
subsection {* Configuration of the code generator *} |
|
11533 | 16 |
|
17 |
types_code |
|
18 |
"bool" ("bool") |
|
12439 | 19 |
"*" ("(_ */ _)") |
20 |
"list" ("_ list") |
|
11533 | 21 |
|
22 |
consts_code |
|
23 |
"True" ("true") |
|
24 |
"False" ("false") |
|
25 |
"Not" ("not") |
|
26 |
"op |" ("(_ orelse/ _)") |
|
27 |
"op &" ("(_ andalso/ _)") |
|
28 |
"If" ("(if _/ then _/ else _)") |
|
29 |
||
30 |
"Pair" ("(_,/ _)") |
|
31 |
"fst" ("fst") |
|
32 |
"snd" ("snd") |
|
33 |
||
34 |
"Nil" ("[]") |
|
35 |
"Cons" ("(_ ::/ _)") |
|
12439 | 36 |
|
13093
ab0335307905
code generator: wfrec combinator is now implemented by ML function wf_rec.
berghofe
parents:
12554
diff
changeset
|
37 |
"wfrec" ("wf'_rec?") |
ab0335307905
code generator: wfrec combinator is now implemented by ML function wf_rec.
berghofe
parents:
12554
diff
changeset
|
38 |
|
14102
8af7334af4b3
- Installed specific code generator for equality enforcing that
berghofe
parents:
14049
diff
changeset
|
39 |
quickcheck_params [default_type = int] |
8af7334af4b3
- Installed specific code generator for equality enforcing that
berghofe
parents:
14049
diff
changeset
|
40 |
|
13755 | 41 |
ML {* |
42 |
fun wf_rec f x = f (wf_rec f) x; |
|
43 |
||
14102
8af7334af4b3
- Installed specific code generator for equality enforcing that
berghofe
parents:
14049
diff
changeset
|
44 |
fun term_of_bool b = if b then HOLogic.true_const else HOLogic.false_const; |
13755 | 45 |
val term_of_list = HOLogic.mk_list; |
46 |
val term_of_int = HOLogic.mk_int; |
|
14049 | 47 |
fun term_of_id_42 f T g U (x, y) = HOLogic.pair_const T U $ f x $ g y; |
14102
8af7334af4b3
- Installed specific code generator for equality enforcing that
berghofe
parents:
14049
diff
changeset
|
48 |
fun term_of_fun_type _ T _ U _ = Free ("<function>", T --> U); |
8af7334af4b3
- Installed specific code generator for equality enforcing that
berghofe
parents:
14049
diff
changeset
|
49 |
|
8af7334af4b3
- Installed specific code generator for equality enforcing that
berghofe
parents:
14049
diff
changeset
|
50 |
val eq_codegen_setup = [Codegen.add_codegen "eq_codegen" |
8af7334af4b3
- Installed specific code generator for equality enforcing that
berghofe
parents:
14049
diff
changeset
|
51 |
(fn thy => fn gr => fn dep => fn b => fn t => |
8af7334af4b3
- Installed specific code generator for equality enforcing that
berghofe
parents:
14049
diff
changeset
|
52 |
(case strip_comb t of |
8af7334af4b3
- Installed specific code generator for equality enforcing that
berghofe
parents:
14049
diff
changeset
|
53 |
(Const ("op =", Type (_, [Type ("fun", _), _])), _) => None |
8af7334af4b3
- Installed specific code generator for equality enforcing that
berghofe
parents:
14049
diff
changeset
|
54 |
| (Const ("op =", _), [t, u]) => |
8af7334af4b3
- Installed specific code generator for equality enforcing that
berghofe
parents:
14049
diff
changeset
|
55 |
let |
8af7334af4b3
- Installed specific code generator for equality enforcing that
berghofe
parents:
14049
diff
changeset
|
56 |
val (gr', pt) = Codegen.invoke_codegen thy dep false (gr, t); |
8af7334af4b3
- Installed specific code generator for equality enforcing that
berghofe
parents:
14049
diff
changeset
|
57 |
val (gr'', pu) = Codegen.invoke_codegen thy dep false (gr', u) |
8af7334af4b3
- Installed specific code generator for equality enforcing that
berghofe
parents:
14049
diff
changeset
|
58 |
in |
8af7334af4b3
- Installed specific code generator for equality enforcing that
berghofe
parents:
14049
diff
changeset
|
59 |
Some (gr'', Codegen.parens |
8af7334af4b3
- Installed specific code generator for equality enforcing that
berghofe
parents:
14049
diff
changeset
|
60 |
(Pretty.block [pt, Pretty.str " =", Pretty.brk 1, pu])) |
8af7334af4b3
- Installed specific code generator for equality enforcing that
berghofe
parents:
14049
diff
changeset
|
61 |
end |
8af7334af4b3
- Installed specific code generator for equality enforcing that
berghofe
parents:
14049
diff
changeset
|
62 |
| (t as Const ("op =", _), ts) => Some (Codegen.invoke_codegen |
8af7334af4b3
- Installed specific code generator for equality enforcing that
berghofe
parents:
14049
diff
changeset
|
63 |
thy dep b (gr, Codegen.eta_expand t ts 2)) |
8af7334af4b3
- Installed specific code generator for equality enforcing that
berghofe
parents:
14049
diff
changeset
|
64 |
| _ => None))]; |
8af7334af4b3
- Installed specific code generator for equality enforcing that
berghofe
parents:
14049
diff
changeset
|
65 |
|
8af7334af4b3
- Installed specific code generator for equality enforcing that
berghofe
parents:
14049
diff
changeset
|
66 |
fun gen_bool i = one_of [false, true]; |
8af7334af4b3
- Installed specific code generator for equality enforcing that
berghofe
parents:
14049
diff
changeset
|
67 |
|
8af7334af4b3
- Installed specific code generator for equality enforcing that
berghofe
parents:
14049
diff
changeset
|
68 |
fun gen_list' aG i j = frequency |
8af7334af4b3
- Installed specific code generator for equality enforcing that
berghofe
parents:
14049
diff
changeset
|
69 |
[(i, fn () => aG j :: gen_list' aG (i-1) j), (1, fn () => [])] () |
8af7334af4b3
- Installed specific code generator for equality enforcing that
berghofe
parents:
14049
diff
changeset
|
70 |
and gen_list aG i = gen_list' aG i i; |
8af7334af4b3
- Installed specific code generator for equality enforcing that
berghofe
parents:
14049
diff
changeset
|
71 |
|
8af7334af4b3
- Installed specific code generator for equality enforcing that
berghofe
parents:
14049
diff
changeset
|
72 |
fun gen_int i = one_of [~1, 1] * random_range 0 i; |
8af7334af4b3
- Installed specific code generator for equality enforcing that
berghofe
parents:
14049
diff
changeset
|
73 |
|
8af7334af4b3
- Installed specific code generator for equality enforcing that
berghofe
parents:
14049
diff
changeset
|
74 |
fun gen_id_42 aG bG i = (aG i, bG i); |
8af7334af4b3
- Installed specific code generator for equality enforcing that
berghofe
parents:
14049
diff
changeset
|
75 |
|
8af7334af4b3
- Installed specific code generator for equality enforcing that
berghofe
parents:
14049
diff
changeset
|
76 |
fun gen_fun_type _ G i = |
8af7334af4b3
- Installed specific code generator for equality enforcing that
berghofe
parents:
14049
diff
changeset
|
77 |
let |
8af7334af4b3
- Installed specific code generator for equality enforcing that
berghofe
parents:
14049
diff
changeset
|
78 |
val f = ref (fn x => raise ERROR); |
8af7334af4b3
- Installed specific code generator for equality enforcing that
berghofe
parents:
14049
diff
changeset
|
79 |
val _ = (f := (fn x => |
8af7334af4b3
- Installed specific code generator for equality enforcing that
berghofe
parents:
14049
diff
changeset
|
80 |
let |
8af7334af4b3
- Installed specific code generator for equality enforcing that
berghofe
parents:
14049
diff
changeset
|
81 |
val y = G i; |
8af7334af4b3
- Installed specific code generator for equality enforcing that
berghofe
parents:
14049
diff
changeset
|
82 |
val f' = !f |
8af7334af4b3
- Installed specific code generator for equality enforcing that
berghofe
parents:
14049
diff
changeset
|
83 |
in (f := (fn x' => if x = x' then y else f' x'); y) end)) |
8af7334af4b3
- Installed specific code generator for equality enforcing that
berghofe
parents:
14049
diff
changeset
|
84 |
in (fn x => !f x) end; |
13755 | 85 |
*} |
13093
ab0335307905
code generator: wfrec combinator is now implemented by ML function wf_rec.
berghofe
parents:
12554
diff
changeset
|
86 |
|
14102
8af7334af4b3
- Installed specific code generator for equality enforcing that
berghofe
parents:
14049
diff
changeset
|
87 |
setup eq_codegen_setup |
8af7334af4b3
- Installed specific code generator for equality enforcing that
berghofe
parents:
14049
diff
changeset
|
88 |
|
12554
671b4d632c34
Declared characteristic equations for < on nat for code generation.
berghofe
parents:
12439
diff
changeset
|
89 |
lemma [code]: "((n::nat) < 0) = False" by simp |
14192
d6cb80cc1d20
Improved efficiency of code generated for < predicate on natural numbers.
berghofe
parents:
14102
diff
changeset
|
90 |
lemma [code]: "(0 < Suc n) = True" by simp |
d6cb80cc1d20
Improved efficiency of code generated for < predicate on natural numbers.
berghofe
parents:
14102
diff
changeset
|
91 |
lemmas [code] = Suc_less_eq imp_conv_disj |
12554
671b4d632c34
Declared characteristic equations for < on nat for code generation.
berghofe
parents:
12439
diff
changeset
|
92 |
|
14350 | 93 |
subsection {* Configuration of the 'refute' command *} |
94 |
||
95 |
text {* |
|
14458 | 96 |
The following are fairly reasonable default values. For an |
97 |
explanation of these parameters, see 'HOL/Refute.thy'. |
|
14350 | 98 |
*} |
99 |
||
100 |
refute_params [minsize=1, |
|
101 |
maxsize=8, |
|
14806 | 102 |
maxvars=10000, |
103 |
maxtime=60, |
|
104 |
satsolver="auto"] |
|
14350 | 105 |
|
9650
6f0b89f2a1f9
Main now new-style theory; added Main.ML for compatibility;
wenzelm
parents:
9619
diff
changeset
|
106 |
end |