| 1476 |      1 | (*  Title:      HOL/Lex/Chopper.thy
 | 
| 1344 |      2 |     ID:         $Id$
 | 
| 1476 |      3 |     Author:     Tobias Nipkow
 | 
| 1344 |      4 |     Copyright   1995 TUM
 | 
|  |      5 | 
 | 
|  |      6 | A 'chopper' is a function which returns
 | 
|  |      7 |   1. a chopped up prefix of the input string
 | 
|  |      8 |   2. together with the remaining suffix.
 | 
|  |      9 | 
 | 
|  |     10 | It is a longest_prefix_chopper if it
 | 
|  |     11 |   1. chops up as much of the input as possible and
 | 
|  |     12 |   2. chops it up into the longest substrings possible.
 | 
|  |     13 | 
 | 
|  |     14 | A chopper is parametrized by a language ('a list => bool),
 | 
|  |     15 | i.e. a set of strings.
 | 
|  |     16 | *)
 | 
|  |     17 | 
 | 
| 14428 |     18 | theory Chopper = List_Prefix:
 | 
| 1344 |     19 | 
 | 
| 14428 |     20 | types 'a chopper = "'a list => 'a list list * 'a list"
 | 
| 1344 |     21 | 
 | 
| 1559 |     22 | constdefs
 | 
| 14428 |     23 |   is_longest_prefix_chopper :: "('a list => bool) => 'a chopper => bool"
 | 
|  |     24 |  "is_longest_prefix_chopper L chopper == !xs.
 | 
| 1559 |     25 |        (!zs. chopper(xs) = ([],zs) -->
 | 
|  |     26 |              zs=xs & (!ys. ys ~= [] & ys <= xs --> ~L(ys))) &
 | 
|  |     27 |        (!ys yss zs. chopper(xs) = (ys#yss,zs) -->
 | 
| 2609 |     28 |           ys ~= [] & L(ys) & xs=ys@concat(yss)@zs &
 | 
|  |     29 |           chopper(concat(yss)@zs) = (yss,zs) &
 | 
| 1559 |     30 |           (!as. as <= xs & ys <= as & ys ~= as --> ~L(as)))"
 | 
| 1344 |     31 | 
 | 
|  |     32 | end
 |