| author | wenzelm |
| Tue, 22 Feb 2000 21:49:34 +0100 | |
| changeset 8280 | 259073d16f84 |
| parent 7867 | 2efb66472812 |
| child 9101 | b643f4d7b9e9 |
| permissions | -rw-r--r-- |
|
3120
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
1 |
(* Title: HOL/ex/Acc.thy |
|
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
2 |
ID: $Id$ |
|
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
|
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
4 |
Copyright 1994 University of Cambridge |
|
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
5 |
|
|
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
6 |
Inductive definition of acc(r) |
|
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
7 |
|
|
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
8 |
See Ch. Paulin-Mohring, Inductive Definitions in the System Coq. |
|
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
9 |
Research Report 92-49, LIP, ENS Lyon. Dec 1992. |
|
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
10 |
*) |
|
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
11 |
|
| 7867 | 12 |
header {* The accessible part of a relation *};
|
| 7759 | 13 |
|
14 |
theory Acc = WF + Inductive:; |
|
|
3120
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
15 |
|
|
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
16 |
consts |
| 7759 | 17 |
acc :: "('a * 'a)set => 'a set" -- {* accessible part *};
|
|
3120
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
18 |
|
| 7721 | 19 |
inductive "acc r" |
|
3120
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
20 |
intrs |
| 7800 | 21 |
accI [rulify_prems]: |
22 |
"ALL y. (y, x) : r --> y : acc r ==> x : acc r" |
|
|
3120
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
23 |
|
| 7800 | 24 |
syntax |
25 |
termi :: "('a * 'a)set => 'a set"
|
|
26 |
translations |
|
27 |
"termi r" == "acc(r^-1)" |
|
| 5273 | 28 |
|
|
3120
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
29 |
end |