| author | haftmann | 
| Thu, 18 Apr 2019 06:06:54 +0000 | |
| changeset 70183 | 3ea80c950023 | 
| parent 61343 | 5b5656a63bd6 | 
| permissions | -rw-r--r-- | 
| 30021 | 1  | 
(* Author: Florian Haftmann, TU Muenchen *)  | 
| 23268 | 2  | 
|
| 61343 | 3  | 
section \<open>Small examples for evaluation mechanisms\<close>  | 
| 23268 | 4  | 
|
5  | 
theory Eval_Examples  | 
|
| 32067 | 6  | 
imports Complex_Main  | 
| 23268 | 7  | 
begin  | 
8  | 
||
| 61343 | 9  | 
text \<open>evaluation oracle\<close>  | 
| 25099 | 10  | 
|
11  | 
lemma "True \<or> False" by eval  | 
|
| 40760 | 12  | 
lemma "Suc 0 \<noteq> Suc 1" by eval  | 
13  | 
lemma "[] = ([] :: int list)" by eval  | 
|
| 25099 | 14  | 
lemma "[()] = [()]" by eval  | 
| 40760 | 15  | 
lemma "fst ([] :: nat list, Suc 0) = []" by eval  | 
| 25099 | 16  | 
|
| 61343 | 17  | 
text \<open>normalization\<close>  | 
| 23268 | 18  | 
|
| 25099 | 19  | 
lemma "True \<or> False" by normalization  | 
| 40760 | 20  | 
lemma "Suc 0 \<noteq> Suc 1" by normalization  | 
21  | 
lemma "[] = ([] :: int list)" by normalization  | 
|
| 25099 | 22  | 
lemma "[()] = [()]" by normalization  | 
| 40760 | 23  | 
lemma "fst ([] :: nat list, Suc 0) = []" by normalization  | 
| 23268 | 24  | 
|
| 61343 | 25  | 
text \<open>term evaluation\<close>  | 
| 23268 | 26  | 
|
| 24587 | 27  | 
value "(Suc 2 + 1) * 4"  | 
| 
24835
 
8c26128f8997
clarified relationship of code generator conversions and evaluations
 
haftmann 
parents: 
24659 
diff
changeset
 | 
28  | 
|
| 24587 | 29  | 
value "(Suc 2 + Suc 0) * Suc 3"  | 
| 
24835
 
8c26128f8997
clarified relationship of code generator conversions and evaluations
 
haftmann 
parents: 
24659 
diff
changeset
 | 
30  | 
|
| 24587 | 31  | 
value "nat 100"  | 
| 
24835
 
8c26128f8997
clarified relationship of code generator conversions and evaluations
 
haftmann 
parents: 
24659 
diff
changeset
 | 
32  | 
|
| 40760 | 33  | 
value "(10::int) \<le> 12"  | 
| 
24835
 
8c26128f8997
clarified relationship of code generator conversions and evaluations
 
haftmann 
parents: 
24659 
diff
changeset
 | 
34  | 
|
| 
 
8c26128f8997
clarified relationship of code generator conversions and evaluations
 
haftmann 
parents: 
24659 
diff
changeset
 | 
35  | 
value "max (2::int) 4"  | 
| 
 
8c26128f8997
clarified relationship of code generator conversions and evaluations
 
haftmann 
parents: 
24659 
diff
changeset
 | 
36  | 
|
| 
 
8c26128f8997
clarified relationship of code generator conversions and evaluations
 
haftmann 
parents: 
24659 
diff
changeset
 | 
37  | 
value "of_int 2 / of_int 4 * (1::rat)"  | 
| 
 
8c26128f8997
clarified relationship of code generator conversions and evaluations
 
haftmann 
parents: 
24659 
diff
changeset
 | 
38  | 
|
| 40760 | 39  | 
value "[] :: nat list"  | 
| 
24835
 
8c26128f8997
clarified relationship of code generator conversions and evaluations
 
haftmann 
parents: 
24659 
diff
changeset
 | 
40  | 
|
| 24587 | 41  | 
value "[(nat 100, ())]"  | 
| 24659 | 42  | 
|
| 61343 | 43  | 
text \<open>a fancy datatype\<close>  | 
| 23268 | 44  | 
|
| 58310 | 45  | 
datatype ('a, 'b) foo =
 | 
| 61076 | 46  | 
Foo "'a::order" 'b  | 
| 40757 | 47  | 
  | Bla "('a, 'b) bar"
 | 
48  | 
| Dummy nat  | 
|
49  | 
and ('a, 'b) bar =
 | 
|
50  | 
Bar 'a 'b  | 
|
51  | 
  | Blubb "('a, 'b) foo"
 | 
|
| 23268 | 52  | 
|
| 40757 | 53  | 
value "Bla (Bar (4::nat) [Suc 0])"  | 
| 23268 | 54  | 
|
55  | 
end  |