doc-src/TutorialI/Misc/Translations.thy
author wenzelm
Fri, 20 Jul 2001 17:49:21 +0200
changeset 11431 2328d48eeba8
parent 11428 332347b9b942
child 11456 7eb63f63e6c6
permissions -rw-r--r--
*** empty log message ***
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
11208
76bc8ea0c6f2 *** empty log message ***
nipkow
parents:
diff changeset
     1
(*<*)
76bc8ea0c6f2 *** empty log message ***
nipkow
parents:
diff changeset
     2
theory Translations = Main:
76bc8ea0c6f2 *** empty log message ***
nipkow
parents:
diff changeset
     3
(*>*)
76bc8ea0c6f2 *** empty log message ***
nipkow
parents:
diff changeset
     4
subsection{*Syntax Translations*}
76bc8ea0c6f2 *** empty log message ***
nipkow
parents:
diff changeset
     5
76bc8ea0c6f2 *** empty log message ***
nipkow
parents:
diff changeset
     6
text{*\label{sec:def-translations}
11428
332347b9b942 tidying the index
paulson
parents: 11310
diff changeset
     7
\indexbold{syntax translations|(}%
332347b9b942 tidying the index
paulson
parents: 11310
diff changeset
     8
\indexbold{translations@\isacommand {translations} (command)|(}
11208
76bc8ea0c6f2 *** empty log message ***
nipkow
parents:
diff changeset
     9
Isabelle offers an additional definition-like facility,
11428
332347b9b942 tidying the index
paulson
parents: 11310
diff changeset
    10
\textbf{syntax translations}.
11211
febb93798bb8 translations: a tweak
paulson
parents: 11208
diff changeset
    11
They resemble macros: upon parsing, the defined concept is immediately
11208
76bc8ea0c6f2 *** empty log message ***
nipkow
parents:
diff changeset
    12
replaced by its definition, and this is reversed upon printing. For example,
76bc8ea0c6f2 *** empty log message ***
nipkow
parents:
diff changeset
    13
the symbol @{text"\<noteq>"} is defined via a syntax translation:
76bc8ea0c6f2 *** empty log message ***
nipkow
parents:
diff changeset
    14
*}
76bc8ea0c6f2 *** empty log message ***
nipkow
parents:
diff changeset
    15
76bc8ea0c6f2 *** empty log message ***
nipkow
parents:
diff changeset
    16
translations "x \<noteq> y" \<rightleftharpoons> "\<not>(x = y)"
76bc8ea0c6f2 *** empty log message ***
nipkow
parents:
diff changeset
    17
11428
332347b9b942 tidying the index
paulson
parents: 11310
diff changeset
    18
text{*\indexbold{$IsaEqTrans@\isasymrightleftharpoons}
11208
76bc8ea0c6f2 *** empty log message ***
nipkow
parents:
diff changeset
    19
\noindent
76bc8ea0c6f2 *** empty log message ***
nipkow
parents:
diff changeset
    20
Internally, @{text"\<noteq>"} never appears.
76bc8ea0c6f2 *** empty log message ***
nipkow
parents:
diff changeset
    21
76bc8ea0c6f2 *** empty log message ***
nipkow
parents:
diff changeset
    22
In addition to @{text"\<rightleftharpoons>"} there are
76bc8ea0c6f2 *** empty log message ***
nipkow
parents:
diff changeset
    23
@{text"\<rightharpoonup>"}\indexbold{$IsaEqTrans1@\isasymrightharpoonup}
76bc8ea0c6f2 *** empty log message ***
nipkow
parents:
diff changeset
    24
and @{text"\<leftharpoondown>"}\indexbold{$IsaEqTrans2@\isasymleftharpoondown}
11211
febb93798bb8 translations: a tweak
paulson
parents: 11208
diff changeset
    25
for uni-directional translations, which only affect
febb93798bb8 translations: a tweak
paulson
parents: 11208
diff changeset
    26
parsing or printing.  We do not want to cover the details of
11310
51e70b7bc315 spelling check
paulson
parents: 11211
diff changeset
    27
translations at this point.  We have mentioned the concept merely because it
11211
febb93798bb8 translations: a tweak
paulson
parents: 11208
diff changeset
    28
crops up occasionally: a number of HOL's built-in constructs are defined
febb93798bb8 translations: a tweak
paulson
parents: 11208
diff changeset
    29
via translations.  Translations are preferable to definitions when the new 
febb93798bb8 translations: a tweak
paulson
parents: 11208
diff changeset
    30
concept is a trivial variation on an existing one.  For example, we
febb93798bb8 translations: a tweak
paulson
parents: 11208
diff changeset
    31
don't need to derive new theorems about @{text"\<noteq>"}, since existing theorems
11428
332347b9b942 tidying the index
paulson
parents: 11310
diff changeset
    32
about @{text"="} still apply.%
332347b9b942 tidying the index
paulson
parents: 11310
diff changeset
    33
\indexbold{syntax translations|)}%
332347b9b942 tidying the index
paulson
parents: 11310
diff changeset
    34
\indexbold{translations@\isacommand {translations} (command)|)}
11211
febb93798bb8 translations: a tweak
paulson
parents: 11208
diff changeset
    35
*}
11208
76bc8ea0c6f2 *** empty log message ***
nipkow
parents:
diff changeset
    36
76bc8ea0c6f2 *** empty log message ***
nipkow
parents:
diff changeset
    37
(*<*)
76bc8ea0c6f2 *** empty log message ***
nipkow
parents:
diff changeset
    38
end
76bc8ea0c6f2 *** empty log message ***
nipkow
parents:
diff changeset
    39
(*>*)