src/HOL/Lex/AutoChopper.thy
author nipkow
Thu, 14 May 1998 16:54:20 +0200
changeset 4931 2ec84dee7911
parent 4832 bc11b5b06f87
child 4955 a9fa93e1a86e
permissions -rw-r--r--
Reordred arguments in AutoChopper. Updated README with ref to paper.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1476
608483c2122a expanded tabs; incorporated Konrad's changes
clasohm
parents: 1374
diff changeset
     1
(*  Title:      HOL/Lex/AutoChopper.thy
1344
f172a7f14e49 Half a lexical analyzer generator.
nipkow
parents:
diff changeset
     2
    ID:         $Id$
1476
608483c2122a expanded tabs; incorporated Konrad's changes
clasohm
parents: 1374
diff changeset
     3
    Author:     Tobias Nipkow
1344
f172a7f14e49 Half a lexical analyzer generator.
nipkow
parents:
diff changeset
     4
    Copyright   1995 TUM
f172a7f14e49 Half a lexical analyzer generator.
nipkow
parents:
diff changeset
     5
f172a7f14e49 Half a lexical analyzer generator.
nipkow
parents:
diff changeset
     6
auto_chopper turns an automaton into a chopper. Tricky, because primrec.
f172a7f14e49 Half a lexical analyzer generator.
nipkow
parents:
diff changeset
     7
f172a7f14e49 Half a lexical analyzer generator.
nipkow
parents:
diff changeset
     8
is_auto_chopper requires its argument to produce longest_prefix_choppers
f172a7f14e49 Half a lexical analyzer generator.
nipkow
parents:
diff changeset
     9
wrt the language accepted by the automaton.
f172a7f14e49 Half a lexical analyzer generator.
nipkow
parents:
diff changeset
    10
f172a7f14e49 Half a lexical analyzer generator.
nipkow
parents:
diff changeset
    11
Main result: auto_chopper satisfies the is_auto_chopper specification.
1362
5fdd4da11d49 Added lots of comments
nipkow
parents: 1344
diff changeset
    12
5fdd4da11d49 Added lots of comments
nipkow
parents: 1344
diff changeset
    13
WARNING: auto_chopper is exponential(?)
5fdd4da11d49 Added lots of comments
nipkow
parents: 1344
diff changeset
    14
if the recursive calls in the penultimate argument are evaluated eagerly.
4137
2ce2e659c2b1 Added an alternativ version of AutoChopper and a theory for the conversion of
nipkow
parents: 1900
diff changeset
    15
2ce2e659c2b1 Added an alternativ version of AutoChopper and a theory for the conversion of
nipkow
parents: 1900
diff changeset
    16
A more efficient version is defined in AutoChopper1.
4931
2ec84dee7911 Reordred arguments in AutoChopper.
nipkow
parents: 4832
diff changeset
    17
2ec84dee7911 Reordred arguments in AutoChopper.
nipkow
parents: 4832
diff changeset
    18
But both versions are far too specific. Better development in Scanner.thy and
2ec84dee7911 Reordred arguments in AutoChopper.
nipkow
parents: 4832
diff changeset
    19
its antecedents.
1344
f172a7f14e49 Half a lexical analyzer generator.
nipkow
parents:
diff changeset
    20
*)
f172a7f14e49 Half a lexical analyzer generator.
nipkow
parents:
diff changeset
    21
4832
bc11b5b06f87 Added conversion of reg.expr. to automata.
nipkow
parents: 4137
diff changeset
    22
AutoChopper = Prefix + DA + Chopper +
bc11b5b06f87 Added conversion of reg.expr. to automata.
nipkow
parents: 4137
diff changeset
    23
bc11b5b06f87 Added conversion of reg.expr. to automata.
nipkow
parents: 4137
diff changeset
    24
constdefs
bc11b5b06f87 Added conversion of reg.expr. to automata.
nipkow
parents: 4137
diff changeset
    25
 is_auto_chopper :: (('a,'s)da => 'a chopper) => bool
bc11b5b06f87 Added conversion of reg.expr. to automata.
nipkow
parents: 4137
diff changeset
    26
"is_auto_chopper(chopperf) ==
bc11b5b06f87 Added conversion of reg.expr. to automata.
nipkow
parents: 4137
diff changeset
    27
    !A. is_longest_prefix_chopper(accepts A)(chopperf A)"
1344
f172a7f14e49 Half a lexical analyzer generator.
nipkow
parents:
diff changeset
    28
f172a7f14e49 Half a lexical analyzer generator.
nipkow
parents:
diff changeset
    29
consts
4931
2ec84dee7911 Reordred arguments in AutoChopper.
nipkow
parents: 4832
diff changeset
    30
  acc :: "[('a,'s)da, 's, 'a list list*'a list, 'a list, 'a list, 'a list]
1476
608483c2122a expanded tabs; incorporated Konrad's changes
clasohm
parents: 1374
diff changeset
    31
          => 'a list list * 'a list"
1344
f172a7f14e49 Half a lexical analyzer generator.
nipkow
parents:
diff changeset
    32
primrec acc List.list
4931
2ec84dee7911 Reordred arguments in AutoChopper.
nipkow
parents: 4832
diff changeset
    33
  "acc A s r []     ys zs = (if ys=[] then r else (ys#fst(r),snd(r)))" 
2ec84dee7911 Reordred arguments in AutoChopper.
nipkow
parents: 4832
diff changeset
    34
  "acc A s r (x#xs) ys zs =
4832
bc11b5b06f87 Added conversion of reg.expr. to automata.
nipkow
parents: 4137
diff changeset
    35
            (let t = next A x s
bc11b5b06f87 Added conversion of reg.expr. to automata.
nipkow
parents: 4137
diff changeset
    36
             in if fin A t
4931
2ec84dee7911 Reordred arguments in AutoChopper.
nipkow
parents: 4832
diff changeset
    37
                then acc A t (acc A (start A) ([],xs) xs [] [])
2ec84dee7911 Reordred arguments in AutoChopper.
nipkow
parents: 4832
diff changeset
    38
                         xs (zs@[x]) (zs@[x])
2ec84dee7911 Reordred arguments in AutoChopper.
nipkow
parents: 4832
diff changeset
    39
                else acc A t r xs ys (zs@[x]))"
4832
bc11b5b06f87 Added conversion of reg.expr. to automata.
nipkow
parents: 4137
diff changeset
    40
bc11b5b06f87 Added conversion of reg.expr. to automata.
nipkow
parents: 4137
diff changeset
    41
constdefs
bc11b5b06f87 Added conversion of reg.expr. to automata.
nipkow
parents: 4137
diff changeset
    42
 auto_chopper :: ('a,'s)da => 'a chopper
4931
2ec84dee7911 Reordred arguments in AutoChopper.
nipkow
parents: 4832
diff changeset
    43
"auto_chopper A xs == acc A (start A) ([],xs) xs [] []"
4832
bc11b5b06f87 Added conversion of reg.expr. to automata.
nipkow
parents: 4137
diff changeset
    44
bc11b5b06f87 Added conversion of reg.expr. to automata.
nipkow
parents: 4137
diff changeset
    45
(* acc_prefix is an auxiliary notion for the proof *)
bc11b5b06f87 Added conversion of reg.expr. to automata.
nipkow
parents: 4137
diff changeset
    46
constdefs
bc11b5b06f87 Added conversion of reg.expr. to automata.
nipkow
parents: 4137
diff changeset
    47
 acc_prefix :: ('a,'s)da => 'a list => 's => bool
bc11b5b06f87 Added conversion of reg.expr. to automata.
nipkow
parents: 4137
diff changeset
    48
"acc_prefix A xs s == ? us. us <= xs & us~=[] & fin A (delta A us s)"
1344
f172a7f14e49 Half a lexical analyzer generator.
nipkow
parents:
diff changeset
    49
f172a7f14e49 Half a lexical analyzer generator.
nipkow
parents:
diff changeset
    50
end