src/HOL/Library/OptionalSugar.thy
author haftmann
Thu, 19 Jul 2007 21:47:39 +0200
changeset 23854 688a8a7bcd4e
parent 22835 37d3a984d730
child 29494 a189c6274c7a
permissions -rw-r--r--
uniform naming conventions for CG theories

(*  Title:      HOL/Library/OptionalSugar.thy
    ID:         $Id$
    Author:     Gerwin Klain, Tobias Nipkow
    Copyright   2005 NICTA and TUM
*)
(*<*)
theory OptionalSugar
imports LaTeXsugar
begin

(* set *)
translations
  "xs" <= "set xs"

(* append *)
syntax (latex output)
  "appendL" :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" (infixl "\<^raw:\isacharat>" 65)
translations
  "appendL xs ys" <= "xs @ ys" 
  "appendL (appendL xs ys) zs" <= "appendL xs (appendL ys zs)"


(* aligning equations *)
notation (tab output)
  "op ="  ("(_) \<^raw:}\putisatab\isa{\ >=\<^raw:}\putisatab\isa{> (_)" [50,49] 50) and
  "=="  ("(_) \<^raw:}\putisatab\isa{\ >\<equiv>\<^raw:}\putisatab\isa{> (_)")

(* Let *)
translations 
  "_bind (p,DUMMY) e" <= "_bind p (fst e)"
  "_bind (DUMMY,p) e" <= "_bind p (snd e)"

  "_tuple_args x (_tuple_args y z)" ==
    "_tuple_args x (_tuple_arg (_tuple y z))"

  "_bind (Some p) e" <= "_bind p (the e)"
  "_bind (p#DUMMY) e" <= "_bind p (hd e)"
  "_bind (DUMMY#p) e" <= "_bind p (tl e)"


end
(*>*)