src/HOL/ex/Computations.thy
author wenzelm
Sat, 10 Nov 2018 19:01:20 +0100
changeset 69281 599b6d0d199b
parent 68660 4ce18f389f53
child 69597 ff784d5a5bfb
permissions -rw-r--r--
tuned signature;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
64959
9ca021bd718d ML antiquotation for generated computations
haftmann
parents:
diff changeset
     1
(*  Title:      HOL/ex/Computations.thy
9ca021bd718d ML antiquotation for generated computations
haftmann
parents:
diff changeset
     2
    Author:     Florian Haftmann, TU Muenchen
9ca021bd718d ML antiquotation for generated computations
haftmann
parents:
diff changeset
     3
*)
9ca021bd718d ML antiquotation for generated computations
haftmann
parents:
diff changeset
     4
9ca021bd718d ML antiquotation for generated computations
haftmann
parents:
diff changeset
     5
section \<open>Simple example for computations generated by the code generator\<close>
9ca021bd718d ML antiquotation for generated computations
haftmann
parents:
diff changeset
     6
9ca021bd718d ML antiquotation for generated computations
haftmann
parents:
diff changeset
     7
theory Computations
68660
4ce18f389f53 slightly more canonical imports
haftmann
parents: 66453
diff changeset
     8
  imports Main
64959
9ca021bd718d ML antiquotation for generated computations
haftmann
parents:
diff changeset
     9
begin
9ca021bd718d ML antiquotation for generated computations
haftmann
parents:
diff changeset
    10
9ca021bd718d ML antiquotation for generated computations
haftmann
parents:
diff changeset
    11
fun even :: "nat \<Rightarrow> bool"
9ca021bd718d ML antiquotation for generated computations
haftmann
parents:
diff changeset
    12
  where "even 0 \<longleftrightarrow> True"
9ca021bd718d ML antiquotation for generated computations
haftmann
parents:
diff changeset
    13
      | "even (Suc 0) \<longleftrightarrow> False"
9ca021bd718d ML antiquotation for generated computations
haftmann
parents:
diff changeset
    14
      | "even (Suc (Suc n)) \<longleftrightarrow> even n"
9ca021bd718d ML antiquotation for generated computations
haftmann
parents:
diff changeset
    15
  
9ca021bd718d ML antiquotation for generated computations
haftmann
parents:
diff changeset
    16
fun fib :: "nat \<Rightarrow> nat"
9ca021bd718d ML antiquotation for generated computations
haftmann
parents:
diff changeset
    17
  where "fib 0 = 0"
9ca021bd718d ML antiquotation for generated computations
haftmann
parents:
diff changeset
    18
      | "fib (Suc 0) = Suc 0"
9ca021bd718d ML antiquotation for generated computations
haftmann
parents:
diff changeset
    19
      | "fib (Suc (Suc n)) = fib (Suc n) + fib n"
9ca021bd718d ML antiquotation for generated computations
haftmann
parents:
diff changeset
    20
9ca021bd718d ML antiquotation for generated computations
haftmann
parents:
diff changeset
    21
declare [[ML_source_trace]]
9ca021bd718d ML antiquotation for generated computations
haftmann
parents:
diff changeset
    22
9ca021bd718d ML antiquotation for generated computations
haftmann
parents:
diff changeset
    23
ML \<open>
9ca021bd718d ML antiquotation for generated computations
haftmann
parents:
diff changeset
    24
local 
9ca021bd718d ML antiquotation for generated computations
haftmann
parents:
diff changeset
    25
9ca021bd718d ML antiquotation for generated computations
haftmann
parents:
diff changeset
    26
fun int_of_nat @{code "0 :: nat"} = 0
9ca021bd718d ML antiquotation for generated computations
haftmann
parents:
diff changeset
    27
  | int_of_nat (@{code Suc} n) = int_of_nat n + 1;
9ca021bd718d ML antiquotation for generated computations
haftmann
parents:
diff changeset
    28
9ca021bd718d ML antiquotation for generated computations
haftmann
parents:
diff changeset
    29
in
9ca021bd718d ML antiquotation for generated computations
haftmann
parents:
diff changeset
    30
64992
41e2c3617582 extended syntax allows to include datatype constructors directly in computations
haftmann
parents: 64989
diff changeset
    31
val comp_nat = @{computation nat
41e2c3617582 extended syntax allows to include datatype constructors directly in computations
haftmann
parents: 64989
diff changeset
    32
  terms: "plus :: nat \<Rightarrow>_" "times :: nat \<Rightarrow> _" fib
41e2c3617582 extended syntax allows to include datatype constructors directly in computations
haftmann
parents: 64989
diff changeset
    33
  datatypes: nat}
41e2c3617582 extended syntax allows to include datatype constructors directly in computations
haftmann
parents: 64989
diff changeset
    34
  (fn post => post o HOLogic.mk_nat o int_of_nat o the);
41e2c3617582 extended syntax allows to include datatype constructors directly in computations
haftmann
parents: 64989
diff changeset
    35
41e2c3617582 extended syntax allows to include datatype constructors directly in computations
haftmann
parents: 64989
diff changeset
    36
val comp_numeral = @{computation nat
41e2c3617582 extended syntax allows to include datatype constructors directly in computations
haftmann
parents: 64989
diff changeset
    37
  terms: "0 :: nat" "1 :: nat" "2 :: nat" "3 :: nat"}
64988
93aaff2b0ae0 computations and partiality
haftmann
parents: 64987
diff changeset
    38
  (fn post => post o HOLogic.mk_nat o int_of_nat o the);
64959
9ca021bd718d ML antiquotation for generated computations
haftmann
parents:
diff changeset
    39
64992
41e2c3617582 extended syntax allows to include datatype constructors directly in computations
haftmann
parents: 64989
diff changeset
    40
val comp_bool = @{computation bool
41e2c3617582 extended syntax allows to include datatype constructors directly in computations
haftmann
parents: 64989
diff changeset
    41
  terms: HOL.conj HOL.disj HOL.implies
41e2c3617582 extended syntax allows to include datatype constructors directly in computations
haftmann
parents: 64989
diff changeset
    42
    HOL.iff even "less_eq :: nat \<Rightarrow> _" "less :: nat \<Rightarrow> _" "HOL.eq :: nat \<Rightarrow> _"
41e2c3617582 extended syntax allows to include datatype constructors directly in computations
haftmann
parents: 64989
diff changeset
    43
  datatypes: bool}
64988
93aaff2b0ae0 computations and partiality
haftmann
parents: 64987
diff changeset
    44
  (K the);
64959
9ca021bd718d ML antiquotation for generated computations
haftmann
parents:
diff changeset
    45
64992
41e2c3617582 extended syntax allows to include datatype constructors directly in computations
haftmann
parents: 64989
diff changeset
    46
val comp_check = @{computation_check terms: Trueprop};
41e2c3617582 extended syntax allows to include datatype constructors directly in computations
haftmann
parents: 64989
diff changeset
    47
41e2c3617582 extended syntax allows to include datatype constructors directly in computations
haftmann
parents: 64989
diff changeset
    48
val comp_dummy = @{computation "(nat \<times> unit) option"
41e2c3617582 extended syntax allows to include datatype constructors directly in computations
haftmann
parents: 64989
diff changeset
    49
  datatypes: "(nat \<times> unit) option"}
64989
40c36a4aee1f more computation antiquotations
haftmann
parents: 64988
diff changeset
    50
64959
9ca021bd718d ML antiquotation for generated computations
haftmann
parents:
diff changeset
    51
end
9ca021bd718d ML antiquotation for generated computations
haftmann
parents:
diff changeset
    52
\<close>
9ca021bd718d ML antiquotation for generated computations
haftmann
parents:
diff changeset
    53
9ca021bd718d ML antiquotation for generated computations
haftmann
parents:
diff changeset
    54
declare [[ML_source_trace = false]]
9ca021bd718d ML antiquotation for generated computations
haftmann
parents:
diff changeset
    55
  
9ca021bd718d ML antiquotation for generated computations
haftmann
parents:
diff changeset
    56
ML_val \<open>
9ca021bd718d ML antiquotation for generated computations
haftmann
parents:
diff changeset
    57
  comp_nat @{context} @{term "fib (Suc (Suc (Suc 0)) * Suc (Suc (Suc 0))) + Suc 0"}
9ca021bd718d ML antiquotation for generated computations
haftmann
parents:
diff changeset
    58
  |> Syntax.string_of_term @{context}
9ca021bd718d ML antiquotation for generated computations
haftmann
parents:
diff changeset
    59
  |> writeln
9ca021bd718d ML antiquotation for generated computations
haftmann
parents:
diff changeset
    60
\<close>
9ca021bd718d ML antiquotation for generated computations
haftmann
parents:
diff changeset
    61
  
9ca021bd718d ML antiquotation for generated computations
haftmann
parents:
diff changeset
    62
ML_val \<open>
9ca021bd718d ML antiquotation for generated computations
haftmann
parents:
diff changeset
    63
  comp_bool @{context} @{term "fib (Suc (Suc (Suc 0)) * Suc (Suc (Suc 0))) + Suc 0 < fib (Suc (Suc 0))"}
9ca021bd718d ML antiquotation for generated computations
haftmann
parents:
diff changeset
    64
\<close>
9ca021bd718d ML antiquotation for generated computations
haftmann
parents:
diff changeset
    65
9ca021bd718d ML antiquotation for generated computations
haftmann
parents:
diff changeset
    66
ML_val \<open>
64989
40c36a4aee1f more computation antiquotations
haftmann
parents: 64988
diff changeset
    67
  comp_check @{context} @{cprop "fib (Suc (Suc (Suc 0)) * Suc (Suc (Suc 0))) + Suc 0 > fib (Suc (Suc 0))"}
40c36a4aee1f more computation antiquotations
haftmann
parents: 64988
diff changeset
    68
\<close>
40c36a4aee1f more computation antiquotations
haftmann
parents: 64988
diff changeset
    69
  
40c36a4aee1f more computation antiquotations
haftmann
parents: 64988
diff changeset
    70
ML_val \<open>
64959
9ca021bd718d ML antiquotation for generated computations
haftmann
parents:
diff changeset
    71
  comp_numeral @{context} @{term "Suc 42 + 7"}
9ca021bd718d ML antiquotation for generated computations
haftmann
parents:
diff changeset
    72
  |> Syntax.string_of_term @{context}
9ca021bd718d ML antiquotation for generated computations
haftmann
parents:
diff changeset
    73
  |> writeln
9ca021bd718d ML antiquotation for generated computations
haftmann
parents:
diff changeset
    74
\<close>
9ca021bd718d ML antiquotation for generated computations
haftmann
parents:
diff changeset
    75
9ca021bd718d ML antiquotation for generated computations
haftmann
parents:
diff changeset
    76
end