src/Tools/Code_Generator.thy
author blanchet
Tue, 07 Dec 2010 11:56:01 +0100
changeset 41051 2ed1b971fc20
parent 40116 9ed3711366c8
child 43018 121aa59b4d17
permissions -rw-r--r--
give the inner timeout mechanism a chance, since it gives more precise information to the user

(*  Title:   Tools/Code_Generator.thy
    Author:  Florian Haftmann, TU Muenchen
*)

header {* Loading the code generator and related modules *}

theory Code_Generator
imports Pure
uses
  "~~/src/Tools/cache_io.ML"
  "~~/src/Tools/auto_tools.ML"
  "~~/src/Tools/solve_direct.ML"
  "~~/src/Tools/quickcheck.ML"
  "~~/src/Tools/value.ML"
  "~~/src/Tools/Code/code_preproc.ML" 
  "~~/src/Tools/Code/code_thingol.ML"
  "~~/src/Tools/Code/code_simp.ML"
  "~~/src/Tools/Code/code_printer.ML"
  "~~/src/Tools/Code/code_target.ML"
  "~~/src/Tools/Code/code_namespace.ML"
  "~~/src/Tools/Code/code_ml.ML"
  "~~/src/Tools/Code/code_haskell.ML"
  "~~/src/Tools/Code/code_scala.ML"
  ("~~/src/Tools/Code/code_runtime.ML")
  ("~~/src/Tools/nbe.ML")
begin

setup {*
  Solve_Direct.setup
  #> Code_Preproc.setup
  #> Code_Simp.setup
  #> Code_ML.setup
  #> Code_Haskell.setup
  #> Code_Scala.setup
  #> Quickcheck.setup
*}

code_datatype "TYPE('a\<Colon>{})"

definition holds :: "prop" where
  "holds \<equiv> ((\<lambda>x::prop. x) \<equiv> (\<lambda>x. x))"

lemma holds: "PROP holds"
  by (unfold holds_def) (rule reflexive)

code_datatype holds

lemma implies_code [code]:
  "(PROP holds \<Longrightarrow> PROP P) \<equiv> PROP P"
  "(PROP P \<Longrightarrow> PROP holds) \<equiv> PROP holds"
proof -
  show "(PROP holds \<Longrightarrow> PROP P) \<equiv> PROP P"
  proof
    assume "PROP holds \<Longrightarrow> PROP P"
    then show "PROP P" using holds .
  next
    assume "PROP P"
    then show "PROP P" .
  qed
next
  show "(PROP P \<Longrightarrow> PROP holds) \<equiv> PROP holds"
    by rule (rule holds)+
qed  

use "~~/src/Tools/Code/code_runtime.ML"
use "~~/src/Tools/nbe.ML"

setup Nbe.setup

hide_const (open) holds

end