src/HOL/ex/Eval_Examples.thy
author haftmann
Fri, 25 Jan 2008 14:53:56 +0100
changeset 25963 07e08dad8a77
parent 25099 b2c19b9964db
child 26020 ffe1a032d24b
permissions -rw-r--r--
distinguished examples for Efficient_Nat.thy
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
23268
572a483de1b0 renamed ex/Eval_Examples.thy;
wenzelm
parents:
diff changeset
     1
(*  ID:         $Id$
572a483de1b0 renamed ex/Eval_Examples.thy;
wenzelm
parents:
diff changeset
     2
    Author:     Florian Haftmann, TU Muenchen
572a483de1b0 renamed ex/Eval_Examples.thy;
wenzelm
parents:
diff changeset
     3
*)
572a483de1b0 renamed ex/Eval_Examples.thy;
wenzelm
parents:
diff changeset
     4
572a483de1b0 renamed ex/Eval_Examples.thy;
wenzelm
parents:
diff changeset
     5
header {* Small examples for evaluation mechanisms *}
572a483de1b0 renamed ex/Eval_Examples.thy;
wenzelm
parents:
diff changeset
     6
572a483de1b0 renamed ex/Eval_Examples.thy;
wenzelm
parents:
diff changeset
     7
theory Eval_Examples
24659
6b7ac2a43df8 more permissive
haftmann
parents: 24587
diff changeset
     8
imports Eval "~~/src/HOL/Real/Rational"
23268
572a483de1b0 renamed ex/Eval_Examples.thy;
wenzelm
parents:
diff changeset
     9
begin
572a483de1b0 renamed ex/Eval_Examples.thy;
wenzelm
parents:
diff changeset
    10
25099
b2c19b9964db lemmas with normalization
haftmann
parents: 24916
diff changeset
    11
text {* evaluation oracle *}
b2c19b9964db lemmas with normalization
haftmann
parents: 24916
diff changeset
    12
b2c19b9964db lemmas with normalization
haftmann
parents: 24916
diff changeset
    13
lemma "True \<or> False" by eval
b2c19b9964db lemmas with normalization
haftmann
parents: 24916
diff changeset
    14
lemma "\<not> (Suc 0 = Suc 1)" by eval
b2c19b9964db lemmas with normalization
haftmann
parents: 24916
diff changeset
    15
lemma "[] = ([]\<Colon> int list)" by eval
b2c19b9964db lemmas with normalization
haftmann
parents: 24916
diff changeset
    16
lemma "[()] = [()]" by eval
b2c19b9964db lemmas with normalization
haftmann
parents: 24916
diff changeset
    17
lemma "fst ([]::nat list, Suc 0) = []" by eval
b2c19b9964db lemmas with normalization
haftmann
parents: 24916
diff changeset
    18
24292
26ac9fe0e80e added evaluation examples
haftmann
parents: 24280
diff changeset
    19
text {* SML evaluation oracle *}
26ac9fe0e80e added evaluation examples
haftmann
parents: 24280
diff changeset
    20
26ac9fe0e80e added evaluation examples
haftmann
parents: 24280
diff changeset
    21
lemma "True \<or> False" by evaluation
26ac9fe0e80e added evaluation examples
haftmann
parents: 24280
diff changeset
    22
lemma "\<not> (Suc 0 = Suc 1)" by evaluation
26ac9fe0e80e added evaluation examples
haftmann
parents: 24280
diff changeset
    23
lemma "[] = ([]\<Colon> int list)" by evaluation
26ac9fe0e80e added evaluation examples
haftmann
parents: 24280
diff changeset
    24
lemma "[()] = [()]" by evaluation
26ac9fe0e80e added evaluation examples
haftmann
parents: 24280
diff changeset
    25
lemma "fst ([]::nat list, Suc 0) = []" by evaluation
26ac9fe0e80e added evaluation examples
haftmann
parents: 24280
diff changeset
    26
25099
b2c19b9964db lemmas with normalization
haftmann
parents: 24916
diff changeset
    27
text {* normalization *}
23268
572a483de1b0 renamed ex/Eval_Examples.thy;
wenzelm
parents:
diff changeset
    28
25099
b2c19b9964db lemmas with normalization
haftmann
parents: 24916
diff changeset
    29
lemma "True \<or> False" by normalization
b2c19b9964db lemmas with normalization
haftmann
parents: 24916
diff changeset
    30
lemma "\<not> (Suc 0 = Suc 1)" by normalization
b2c19b9964db lemmas with normalization
haftmann
parents: 24916
diff changeset
    31
lemma "[] = ([]\<Colon> int list)" by normalization
b2c19b9964db lemmas with normalization
haftmann
parents: 24916
diff changeset
    32
lemma "[()] = [()]" by normalization
b2c19b9964db lemmas with normalization
haftmann
parents: 24916
diff changeset
    33
lemma "fst ([]::nat list, Suc 0) = []" by normalization
23268
572a483de1b0 renamed ex/Eval_Examples.thy;
wenzelm
parents:
diff changeset
    34
572a483de1b0 renamed ex/Eval_Examples.thy;
wenzelm
parents:
diff changeset
    35
text {* term evaluation *}
572a483de1b0 renamed ex/Eval_Examples.thy;
wenzelm
parents:
diff changeset
    36
24587
4f2cbf6e563f multi-functional value keyword
haftmann
parents: 24423
diff changeset
    37
value "(Suc 2 + 1) * 4"
24835
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
    38
value (code) "(Suc 2 + 1) * 4"
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
    39
value (SML) "(Suc 2 + 1) * 4"
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
    40
value ("normal_form") "(Suc 2 + 1) * 4"
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
    41
24587
4f2cbf6e563f multi-functional value keyword
haftmann
parents: 24423
diff changeset
    42
value "(Suc 2 + Suc 0) * Suc 3"
24835
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
    43
value (code) "(Suc 2 + Suc 0) * Suc 3"
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
    44
value (SML) "(Suc 2 + Suc 0) * Suc 3"
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
    45
value ("normal_form") "(Suc 2 + Suc 0) * Suc 3"
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
    46
24587
4f2cbf6e563f multi-functional value keyword
haftmann
parents: 24423
diff changeset
    47
value "nat 100"
24835
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
    48
value (code) "nat 100"
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
    49
value (SML) "nat 100"
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
    50
value ("normal_form") "nat 100"
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
    51
24587
4f2cbf6e563f multi-functional value keyword
haftmann
parents: 24423
diff changeset
    52
value "(10\<Colon>int) \<le> 12"
24835
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
    53
value (code) "(10\<Colon>int) \<le> 12"
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
    54
value (SML) "(10\<Colon>int) \<le> 12"
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
    55
value ("normal_form") "(10\<Colon>int) \<le> 12"
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
    56
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
    57
value "max (2::int) 4"
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
    58
value (code) "max (2::int) 4"
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
    59
value (SML) "max (2::int) 4"
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
    60
value ("normal_form") "max (2::int) 4"
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
    61
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
    62
value "of_int 2 / of_int 4 * (1::rat)"
24916
dc56dd1b3cda simplified evaluation
haftmann
parents: 24835
diff changeset
    63
(*value (code) "of_int 2 / of_int 4 * (1::rat)"*)
24835
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
    64
value (SML) "of_int 2 / of_int 4 * (1::rat)"
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
    65
value ("normal_form") "of_int 2 / of_int 4 * (1::rat)"
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
    66
24587
4f2cbf6e563f multi-functional value keyword
haftmann
parents: 24423
diff changeset
    67
value "[]::nat list"
24835
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
    68
value (code) "[]::nat list"
25099
b2c19b9964db lemmas with normalization
haftmann
parents: 24916
diff changeset
    69
value (SML) "[]::nat list"
24835
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
    70
value ("normal_form") "[]::nat list"
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
    71
24587
4f2cbf6e563f multi-functional value keyword
haftmann
parents: 24423
diff changeset
    72
value "[(nat 100, ())]"
24835
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
    73
value (code) "[(nat 100, ())]"
25099
b2c19b9964db lemmas with normalization
haftmann
parents: 24916
diff changeset
    74
value (SML) "[(nat 100, ())]"
24835
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
    75
value ("normal_form") "[(nat 100, ())]"
24659
6b7ac2a43df8 more permissive
haftmann
parents: 24587
diff changeset
    76
23268
572a483de1b0 renamed ex/Eval_Examples.thy;
wenzelm
parents:
diff changeset
    77
572a483de1b0 renamed ex/Eval_Examples.thy;
wenzelm
parents:
diff changeset
    78
text {* a fancy datatype *}
572a483de1b0 renamed ex/Eval_Examples.thy;
wenzelm
parents:
diff changeset
    79
572a483de1b0 renamed ex/Eval_Examples.thy;
wenzelm
parents:
diff changeset
    80
datatype ('a, 'b) bair =
572a483de1b0 renamed ex/Eval_Examples.thy;
wenzelm
parents:
diff changeset
    81
    Bair "'a\<Colon>order" 'b
572a483de1b0 renamed ex/Eval_Examples.thy;
wenzelm
parents:
diff changeset
    82
  | Shift "('a, 'b) cair"
572a483de1b0 renamed ex/Eval_Examples.thy;
wenzelm
parents:
diff changeset
    83
  | Dummy unit
572a483de1b0 renamed ex/Eval_Examples.thy;
wenzelm
parents:
diff changeset
    84
and ('a, 'b) cair =
572a483de1b0 renamed ex/Eval_Examples.thy;
wenzelm
parents:
diff changeset
    85
    Cair 'a 'b
572a483de1b0 renamed ex/Eval_Examples.thy;
wenzelm
parents:
diff changeset
    86
24587
4f2cbf6e563f multi-functional value keyword
haftmann
parents: 24423
diff changeset
    87
value "Shift (Cair (4::nat) [Suc 0])"
24835
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
    88
value (code) "Shift (Cair (4::nat) [Suc 0])"
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
    89
value (SML) "Shift (Cair (4::nat) [Suc 0])"
8c26128f8997 clarified relationship of code generator conversions and evaluations
haftmann
parents: 24659
diff changeset
    90
value ("normal_form") "Shift (Cair (4::nat) [Suc 0])"
23268
572a483de1b0 renamed ex/Eval_Examples.thy;
wenzelm
parents:
diff changeset
    91
572a483de1b0 renamed ex/Eval_Examples.thy;
wenzelm
parents:
diff changeset
    92
end