src/Pure/Pure.thy
changeset 59939 7d46aa03696e
parent 59923 b21c82422d65
child 59990 a81dc82ecba3
equal deleted inserted replaced
59938:f84b93187ab6 59939:7d46aa03696e
     9     "!!" "!" "+" "--" ":" ";" "<" "<=" "=" "=>" "?" "[" "\<equiv>"
     9     "!!" "!" "+" "--" ":" ";" "<" "<=" "=" "=>" "?" "[" "\<equiv>"
    10     "\<leftharpoondown>" "\<rightharpoonup>" "\<rightleftharpoons>"
    10     "\<leftharpoondown>" "\<rightharpoonup>" "\<rightleftharpoons>"
    11     "\<subseteq>" "]" "assumes" "attach" "binder" "constrains"
    11     "\<subseteq>" "]" "assumes" "attach" "binder" "constrains"
    12     "defines" "fixes" "for" "identifier" "if" "in" "includes" "infix"
    12     "defines" "fixes" "for" "identifier" "if" "in" "includes" "infix"
    13     "infixl" "infixr" "is" "notes" "obtains" "open" "output"
    13     "infixl" "infixr" "is" "notes" "obtains" "open" "output"
    14     "overloaded" "pervasive" "private" "shows" "structure" "unchecked"
    14     "overloaded" "pervasive" "private" "restricted" "shows"
    15     "where" "|"
    15     "structure" "unchecked" "where" "|"
    16   and "text" "txt" :: document_body
    16   and "text" "txt" :: document_body
    17   and "text_raw" :: document_raw
    17   and "text_raw" :: document_raw
    18   and "default_sort" :: thy_decl == ""
    18   and "default_sort" :: thy_decl == ""
    19   and "typedecl" "type_synonym" "nonterminal" "judgment"
    19   and "typedecl" "type_synonym" "nonterminal" "judgment"
    20     "consts" "syntax" "no_syntax" "translations" "no_translations" "defs"
    20     "consts" "syntax" "no_syntax" "translations" "no_translations" "defs"