12197
|
1 |
(* Title: ZF/UNITY/GenPrefix.thy
|
|
2 |
ID: $Id$
|
|
3 |
Author: Sidi O Ehmety, Cambridge University Computer Laboratory
|
|
4 |
Copyright 2001 University of Cambridge
|
|
5 |
|
|
6 |
Charpentier's Generalized Prefix Relation
|
|
7 |
<xs,ys>:gen_prefix(r)
|
|
8 |
if ys = xs' @ zs where length(xs) = length(xs')
|
|
9 |
and corresponding elements of xs, xs' are pairwise related by r
|
|
10 |
|
|
11 |
Based on Lex/Prefix
|
|
12 |
*)
|
|
13 |
|
|
14 |
GenPrefix = ListPlus +
|
|
15 |
|
|
16 |
consts
|
|
17 |
gen_prefix :: "[i, i] => i"
|
|
18 |
|
|
19 |
inductive
|
|
20 |
(* Parameter A is the domain of zs's elements *)
|
|
21 |
|
|
22 |
domains "gen_prefix(A, r)" <= "list(A)*list(A)"
|
|
23 |
|
|
24 |
intrs
|
|
25 |
Nil "<[],[]>:gen_prefix(A, r)"
|
|
26 |
|
|
27 |
prepend "[| <xs,ys>:gen_prefix(A, r); <x,y>:r; x:A; y:A |]
|
|
28 |
==> <Cons(x,xs), Cons(y,ys)>: gen_prefix(A, r)"
|
|
29 |
|
|
30 |
append "[| <xs,ys>:gen_prefix(A, r); zs:list(A) |]
|
|
31 |
==> <xs, ys@zs>:gen_prefix(A, r)"
|
|
32 |
type_intrs "list.intrs@[app_type]"
|
|
33 |
|
|
34 |
constdefs
|
|
35 |
prefix :: i=>i
|
|
36 |
"prefix(A) == gen_prefix(A, id(A))"
|
|
37 |
|
|
38 |
strict_prefix :: i=>i
|
|
39 |
"strict_prefix(A) == prefix(A) - id(list(A))"
|
|
40 |
|
|
41 |
(* Probably to be moved elsewhere *)
|
|
42 |
|
|
43 |
Le :: i
|
|
44 |
"Le == {<x,y>:nat*nat. x le y}"
|
|
45 |
|
|
46 |
Ge :: i
|
|
47 |
"Ge == {<x,y>:nat*nat. y le x}"
|
|
48 |
|
|
49 |
syntax
|
|
50 |
(* less or equal and greater or equal over prefixes *)
|
|
51 |
pfixLe :: [i, i] => o (infixl "pfixLe" 50)
|
|
52 |
pfixGe :: [i, i] => o (infixl "pfixGe" 50)
|
|
53 |
|
|
54 |
translations
|
|
55 |
"xs pfixLe ys" == "<xs, ys>:gen_prefix(nat, Le)"
|
|
56 |
"xs pfixGe ys" == "<xs, ys>:gen_prefix(nat, Ge)"
|
|
57 |
|
|
58 |
|
|
59 |
end
|