| author | wenzelm | 
| Wed, 30 Aug 2000 13:54:53 +0200 | |
| changeset 9738 | 2e1dca5af2d4 | 
| parent 5184 | 9b8547a9496a | 
| child 10338 | 291ce4c4b50e | 
| permissions | -rw-r--r-- | 
| 1476 | 1 | (* Title: HOL/Lex/AutoChopper.thy | 
| 1344 | 2 | ID: $Id$ | 
| 1476 | 3 | Author: Tobias Nipkow | 
| 1344 | 4 | Copyright 1995 TUM | 
| 5 | ||
| 6 | auto_chopper turns an automaton into a chopper. Tricky, because primrec. | |
| 7 | ||
| 8 | is_auto_chopper requires its argument to produce longest_prefix_choppers | |
| 9 | wrt the language accepted by the automaton. | |
| 10 | ||
| 11 | Main result: auto_chopper satisfies the is_auto_chopper specification. | |
| 1362 | 12 | |
| 13 | WARNING: auto_chopper is exponential(?) | |
| 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: 
1900diff
changeset | 15 | |
| 
2ce2e659c2b1
Added an alternativ version of AutoChopper and a theory for the conversion of
 nipkow parents: 
1900diff
changeset | 16 | A more efficient version is defined in AutoChopper1. | 
| 4931 | 17 | |
| 18 | But both versions are far too specific. Better development in Scanner.thy and | |
| 19 | its antecedents. | |
| 1344 | 20 | *) | 
| 21 | ||
| 4832 | 22 | AutoChopper = Prefix + DA + Chopper + | 
| 23 | ||
| 24 | constdefs | |
| 25 |  is_auto_chopper :: (('a,'s)da => 'a chopper) => bool
 | |
| 26 | "is_auto_chopper(chopperf) == | |
| 27 | !A. is_longest_prefix_chopper(accepts A)(chopperf A)" | |
| 1344 | 28 | |
| 29 | consts | |
| 4931 | 30 |   acc :: "[('a,'s)da, 's, 'a list list*'a list, 'a list, 'a list, 'a list]
 | 
| 1476 | 31 | => 'a list list * 'a list" | 
| 5184 | 32 | |
| 33 | primrec | |
| 4955 | 34 | "acc A s r ps [] zs = (if ps=[] then r else (ps#fst(r),snd(r)))" | 
| 35 | "acc A s r ps (x#xs) zs = | |
| 4832 | 36 | (let t = next A x s | 
| 37 | in if fin A t | |
| 4955 | 38 | then acc A t (acc A (start A) ([],xs) [] xs []) | 
| 39 | (zs@[x]) xs (zs@[x]) | |
| 40 | else acc A t r ps xs (zs@[x]))" | |
| 4832 | 41 | |
| 42 | constdefs | |
| 43 |  auto_chopper :: ('a,'s)da => 'a chopper
 | |
| 4955 | 44 | "auto_chopper A xs == acc A (start A) ([],xs) [] xs []" | 
| 4832 | 45 | |
| 46 | (* acc_prefix is an auxiliary notion for the proof *) | |
| 47 | constdefs | |
| 48 |  acc_prefix :: ('a,'s)da => 'a list => 's => bool
 | |
| 49 | "acc_prefix A xs s == ? us. us <= xs & us~=[] & fin A (delta A us s)" | |
| 1344 | 50 | |
| 51 | end |