Theory HOL-Library.OptionalSugar
theory OptionalSugar
imports Complex_Main LaTeXsugar
begin
translations
  ("prop") "P ∧ Q ⟹ R" <= ("prop") "P ⟹ Q ⟹ R"
translations
  "xs" <= "CONST set xs"
translations
  "n" <= "CONST of_nat n"
  "n" <= "CONST int n"
  "n" <= "CONST real n"
  "n" <= "CONST real_of_nat n"
  "n" <= "CONST real_of_int n"
  "n" <= "CONST of_real n"
  "n" <= "CONST complex_of_real n"
syntax (latex output)
  "_appendL" :: "'a list ⇒ 'a list ⇒ 'a list" (infixl ‹\<^latex>‹\isacharat›› 65)
translations
  "_appendL xs ys" <= "xs @ ys" 
  "_appendL (_appendL xs ys) zs" <= "_appendL xs (_appendL ys zs)"
notation (tab output)
  "HOL.eq"  (‹(_) \<^latex>‹}\putisatab\isa{\ ›=\<^latex>‹}\putisatab\isa{› (_)› [50,49] 50) and
  "Pure.eq"  (‹(_) \<^latex>‹}\putisatab\isa{\ ›≡\<^latex>‹}\putisatab\isa{› (_)›)
translations 
  "_bind (p, CONST DUMMY) e" <= "_bind p (CONST fst e)"
  "_bind (CONST DUMMY, p) e" <= "_bind p (CONST snd e)"
  "_tuple_args x (_tuple_args y z)" ==
    "_tuple_args x (_tuple_arg (_tuple y z))"
  "_bind (CONST Some p) e" <= "_bind p (CONST the e)"
  "_bind (p # CONST DUMMY) e" <= "_bind p (CONST hd e)"
  "_bind (CONST DUMMY # p) e" <= "_bind p (CONST tl e)"
unbundle constrain_space_syntax
syntax (output)
  "_topsort" :: "sort"  (‹⊤› 1000)
  "_sort" :: "classes => sort"  (‹'(_')› 1000)
  "_classes" :: "id => classes => classes"  (‹_ ∩ _› 1000)
  "_classes" :: "longid => classes => classes"  (‹_ ∩ _› 1000)
end