src/HOL/Library/Rewrite.thy
author nipkow
Mon, 06 Apr 2015 15:23:50 +0200
changeset 59928 b9b7f913a19a
parent 59739 4ed50ebf5d36
child 59975 da10875adf8e
permissions -rw-r--r--
new theory Library/Tree_Multiset.thy

theory Rewrite
imports Main
begin

consts rewrite_HOLE :: "'a :: {}"
notation rewrite_HOLE ("HOLE")
notation rewrite_HOLE ("\<box>")

lemma eta_expand:
  fixes f :: "'a :: {} \<Rightarrow> 'b :: {}" shows "f \<equiv> (\<lambda>x. f x)"
  by (rule reflexive)

ML_file "cconv.ML"
ML_file "rewrite.ML"
setup Rewrite.setup

end