doc-src/TutorialI/Misc/types.thy
author nipkow
Sun, 06 Aug 2000 15:26:53 +0200
changeset 9541 d17c0b34d5c8
parent 8771 026f37a86ea7
child 9792 bbefb6ce5cb2
permissions -rw-r--r--
*** empty log message ***
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
     1
(*<*)
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
     2
theory "types" = Main:;
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
     3
(*>*)types number       = nat
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
     4
      gate         = "bool \\<Rightarrow> bool \\<Rightarrow> bool"
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
     5
      ('a,'b)alist = "('a * 'b)list";
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
     6
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
     7
text{*\noindent\indexbold{*types}%
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
     8
Internally all synonyms are fully expanded.  As a consequence Isabelle's
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
     9
output never contains synonyms.  Their main purpose is to improve the
8771
026f37a86ea7 *** empty log message ***
nipkow
parents: 8745
diff changeset
    10
readability of theories.  Synonyms can be used just like any other
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    11
type:
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    12
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    13
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    14
consts nand :: gate
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    15
       exor :: gate;
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    16
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    17
text{*
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    18
\subsection{Constant definitions}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    19
\label{sec:ConstDefinitions}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    20
\indexbold{definition}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    21
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    22
The above constants \isa{nand} and \isa{exor} are non-recursive and can
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    23
therefore be defined directly by
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    24
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    25
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    26
defs nand_def: "nand A B \\<equiv> \\<not>(A \\<and> B)"
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    27
     exor_def: "exor A B \\<equiv> A \\<and> \\<not>B \\<or> \\<not>A \\<and> B";
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    28
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    29
text{*\noindent%
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    30
where \isacommand{defs}\indexbold{*defs} is a keyword and \isa{nand_def} and
8771
026f37a86ea7 *** empty log message ***
nipkow
parents: 8745
diff changeset
    31
\isa{exor_def} are user-supplied names.
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    32
The symbol \indexboldpos{\isasymequiv}{$IsaEq} is a special form of equality
9541
d17c0b34d5c8 *** empty log message ***
nipkow
parents: 8771
diff changeset
    33
that must be used in constant definitions.
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    34
Declarations and definitions can also be merged
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    35
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    36
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    37
constdefs nor :: gate
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    38
         "nor A B \\<equiv> \\<not>(A \\<or> B)"
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    39
          exor2 :: gate
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    40
         "exor2 A B \\<equiv> (A \\<or> B) \\<and> (\\<not>A \\<or> \\<not>B)";
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    41
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    42
text{*\noindent\indexbold{*constdefs}%
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    43
in which case the default name of each definition is \isa{$f$_def}, where
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    44
$f$ is the name of the defined constant.*}(*<*)
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    45
end
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    46
(*>*)