| author | paulson | 
| Tue, 28 Jun 2005 15:27:45 +0200 | |
| changeset 16587 | b34c8aa657a5 | 
| parent 15965 | f422f8283491 | 
| child 16635 | bf7de5723c60 | 
| 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  | 
||
| 15131 | 8  | 
theory Main  | 
| 15872 | 9  | 
imports Refute Reconstruction  | 
| 15151 | 10  | 
|
| 15131 | 11  | 
begin  | 
| 
9650
 
6f0b89f2a1f9
Main now new-style theory; added Main.ML for compatibility;
 
wenzelm 
parents: 
9619 
diff
changeset
 | 
12  | 
|
| 12024 | 13  | 
text {*
 | 
14  | 
  Theory @{text Main} includes everything.  Note that theory @{text
 | 
|
15  | 
PreList} already includes most HOL theories.  | 
|
16  | 
*}  | 
|
17  | 
||
18  | 
subsection {* Configuration of the code generator *}
 | 
|
| 11533 | 19  | 
|
20  | 
types_code  | 
|
21  | 
  "bool"  ("bool")
 | 
|
22  | 
||
23  | 
consts_code  | 
|
24  | 
  "True"    ("true")
 | 
|
25  | 
  "False"   ("false")
 | 
|
26  | 
  "Not"     ("not")
 | 
|
27  | 
  "op |"    ("(_ orelse/ _)")
 | 
|
28  | 
  "op &"    ("(_ andalso/ _)")
 | 
|
| 16587 | 29  | 
  "HOL.If"      ("(if _/ then _/ else _)")
 | 
| 11533 | 30  | 
|
| 
13093
 
ab0335307905
code generator: wfrec combinator is now implemented by ML function wf_rec.
 
berghofe 
parents: 
12554 
diff
changeset
 | 
31  | 
  "wfrec"   ("wf'_rec?")
 | 
| 
 
ab0335307905
code generator: wfrec combinator is now implemented by ML function wf_rec.
 
berghofe 
parents: 
12554 
diff
changeset
 | 
32  | 
|
| 
14102
 
8af7334af4b3
- Installed specific code generator for equality enforcing that
 
berghofe 
parents: 
14049 
diff
changeset
 | 
33  | 
quickcheck_params [default_type = int]  | 
| 
 
8af7334af4b3
- Installed specific code generator for equality enforcing that
 
berghofe 
parents: 
14049 
diff
changeset
 | 
34  | 
|
| 
15965
 
f422f8283491
Use of IntInf.int instead of int in most numeric simprocs; avoids
 
paulson 
parents: 
15872 
diff
changeset
 | 
35  | 
(*FIXME: the IntInf.fromInt below hides a dependence on fixed-precision ints!*)  | 
| 13755 | 36  | 
ML {*
 | 
37  | 
fun wf_rec f x = f (wf_rec f) x;  | 
|
38  | 
||
| 
14102
 
8af7334af4b3
- Installed specific code generator for equality enforcing that
 
berghofe 
parents: 
14049 
diff
changeset
 | 
39  | 
fun term_of_bool b = if b then HOLogic.true_const else HOLogic.false_const;  | 
| 
15965
 
f422f8283491
Use of IntInf.int instead of int in most numeric simprocs; avoids
 
paulson 
parents: 
15872 
diff
changeset
 | 
40  | 
val term_of_int = HOLogic.mk_int o IntInf.fromInt;  | 
| 
14102
 
8af7334af4b3
- Installed specific code generator for equality enforcing that
 
berghofe 
parents: 
14049 
diff
changeset
 | 
41  | 
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
 | 
42  | 
|
| 
 
8af7334af4b3
- Installed specific code generator for equality enforcing that
 
berghofe 
parents: 
14049 
diff
changeset
 | 
43  | 
val eq_codegen_setup = [Codegen.add_codegen "eq_codegen"  | 
| 
 
8af7334af4b3
- Installed specific code generator for equality enforcing that
 
berghofe 
parents: 
14049 
diff
changeset
 | 
44  | 
(fn thy => fn gr => fn dep => fn b => fn t =>  | 
| 
 
8af7334af4b3
- Installed specific code generator for equality enforcing that
 
berghofe 
parents: 
14049 
diff
changeset
 | 
45  | 
(case strip_comb t of  | 
| 15531 | 46  | 
       (Const ("op =", Type (_, [Type ("fun", _), _])), _) => NONE
 | 
| 
14102
 
8af7334af4b3
- Installed specific code generator for equality enforcing that
 
berghofe 
parents: 
14049 
diff
changeset
 | 
47  | 
     | (Const ("op =", _), [t, u]) =>
 | 
| 
 
8af7334af4b3
- Installed specific code generator for equality enforcing that
 
berghofe 
parents: 
14049 
diff
changeset
 | 
48  | 
let  | 
| 
 
8af7334af4b3
- Installed specific code generator for equality enforcing that
 
berghofe 
parents: 
14049 
diff
changeset
 | 
49  | 
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
 | 
50  | 
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
 | 
51  | 
in  | 
| 15531 | 52  | 
SOME (gr'', Codegen.parens  | 
| 
14102
 
8af7334af4b3
- Installed specific code generator for equality enforcing that
 
berghofe 
parents: 
14049 
diff
changeset
 | 
53  | 
(Pretty.block [pt, Pretty.str " =", Pretty.brk 1, pu]))  | 
| 
 
8af7334af4b3
- Installed specific code generator for equality enforcing that
 
berghofe 
parents: 
14049 
diff
changeset
 | 
54  | 
end  | 
| 15531 | 55  | 
     | (t as Const ("op =", _), ts) => SOME (Codegen.invoke_codegen
 | 
| 
14102
 
8af7334af4b3
- Installed specific code generator for equality enforcing that
 
berghofe 
parents: 
14049 
diff
changeset
 | 
56  | 
thy dep b (gr, Codegen.eta_expand t ts 2))  | 
| 15531 | 57  | 
| _ => NONE))];  | 
| 
14102
 
8af7334af4b3
- Installed specific code generator for equality enforcing that
 
berghofe 
parents: 
14049 
diff
changeset
 | 
58  | 
|
| 
 
8af7334af4b3
- Installed specific code generator for equality enforcing that
 
berghofe 
parents: 
14049 
diff
changeset
 | 
59  | 
fun gen_bool i = one_of [false, true];  | 
| 
 
8af7334af4b3
- Installed specific code generator for equality enforcing that
 
berghofe 
parents: 
14049 
diff
changeset
 | 
60  | 
|
| 
 
8af7334af4b3
- Installed specific code generator for equality enforcing that
 
berghofe 
parents: 
14049 
diff
changeset
 | 
61  | 
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
 | 
62  | 
|
| 
 
8af7334af4b3
- Installed specific code generator for equality enforcing that
 
berghofe 
parents: 
14049 
diff
changeset
 | 
63  | 
fun gen_fun_type _ G i =  | 
| 
 
8af7334af4b3
- Installed specific code generator for equality enforcing that
 
berghofe 
parents: 
14049 
diff
changeset
 | 
64  | 
let  | 
| 
 
8af7334af4b3
- Installed specific code generator for equality enforcing that
 
berghofe 
parents: 
14049 
diff
changeset
 | 
65  | 
val f = ref (fn x => raise ERROR);  | 
| 
 
8af7334af4b3
- Installed specific code generator for equality enforcing that
 
berghofe 
parents: 
14049 
diff
changeset
 | 
66  | 
val _ = (f := (fn x =>  | 
| 
 
8af7334af4b3
- Installed specific code generator for equality enforcing that
 
berghofe 
parents: 
14049 
diff
changeset
 | 
67  | 
let  | 
| 
 
8af7334af4b3
- Installed specific code generator for equality enforcing that
 
berghofe 
parents: 
14049 
diff
changeset
 | 
68  | 
val y = G i;  | 
| 
 
8af7334af4b3
- Installed specific code generator for equality enforcing that
 
berghofe 
parents: 
14049 
diff
changeset
 | 
69  | 
val f' = !f  | 
| 
 
8af7334af4b3
- Installed specific code generator for equality enforcing that
 
berghofe 
parents: 
14049 
diff
changeset
 | 
70  | 
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
 | 
71  | 
in (fn x => !f x) end;  | 
| 13755 | 72  | 
*}  | 
| 
13093
 
ab0335307905
code generator: wfrec combinator is now implemented by ML function wf_rec.
 
berghofe 
parents: 
12554 
diff
changeset
 | 
73  | 
|
| 
14102
 
8af7334af4b3
- Installed specific code generator for equality enforcing that
 
berghofe 
parents: 
14049 
diff
changeset
 | 
74  | 
setup eq_codegen_setup  | 
| 
 
8af7334af4b3
- Installed specific code generator for equality enforcing that
 
berghofe 
parents: 
14049 
diff
changeset
 | 
75  | 
|
| 
12554
 
671b4d632c34
Declared characteristic equations for < on nat for code generation.
 
berghofe 
parents: 
12439 
diff
changeset
 | 
76  | 
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
 | 
77  | 
lemma [code]: "(0 < Suc n) = True" by simp  | 
| 
 
d6cb80cc1d20
Improved efficiency of code generated for < predicate on natural numbers.
 
berghofe 
parents: 
14102 
diff
changeset
 | 
78  | 
lemmas [code] = Suc_less_eq imp_conv_disj  | 
| 
12554
 
671b4d632c34
Declared characteristic equations for < on nat for code generation.
 
berghofe 
parents: 
12439 
diff
changeset
 | 
79  | 
|
| 14350 | 80  | 
subsection {* Configuration of the 'refute' command *}
 | 
81  | 
||
82  | 
text {*
 | 
|
| 14458 | 83  | 
The following are fairly reasonable default values. For an  | 
84  | 
explanation of these parameters, see 'HOL/Refute.thy'.  | 
|
| 14350 | 85  | 
*}  | 
86  | 
||
| 
15584
 
3478bb4f93ff
refute_params: default value itself=1 added (for type classes)
 
webertj 
parents: 
15531 
diff
changeset
 | 
87  | 
refute_params ["itself"=1,  | 
| 
 
3478bb4f93ff
refute_params: default value itself=1 added (for type classes)
 
webertj 
parents: 
15531 
diff
changeset
 | 
88  | 
minsize=1,  | 
| 14350 | 89  | 
maxsize=8,  | 
| 14806 | 90  | 
maxvars=10000,  | 
91  | 
maxtime=60,  | 
|
92  | 
satsolver="auto"]  | 
|
| 14350 | 93  | 
|
| 
9650
 
6f0b89f2a1f9
Main now new-style theory; added Main.ML for compatibility;
 
wenzelm 
parents: 
9619 
diff
changeset
 | 
94  | 
end  |