| author | urbanc |
| Tue, 24 Jan 2006 13:28:06 +0100 | |
| changeset 18773 | 0eabf66582d0 |
| parent 18709 | f174ebc26073 |
| child 18835 | 8e080d0252c5 |
| permissions | -rw-r--r-- |
|
15804
3a55e6e26c8a
The CPure theory -- Pure with alternative application syntax.
wenzelm
parents:
diff
changeset
|
1 |
(* Title: Pure/CPure.thy |
|
3a55e6e26c8a
The CPure theory -- Pure with alternative application syntax.
wenzelm
parents:
diff
changeset
|
2 |
ID: $Id$ |
|
3a55e6e26c8a
The CPure theory -- Pure with alternative application syntax.
wenzelm
parents:
diff
changeset
|
3 |
|
|
3a55e6e26c8a
The CPure theory -- Pure with alternative application syntax.
wenzelm
parents:
diff
changeset
|
4 |
The CPure theory -- Pure with alternative application syntax. |
|
3a55e6e26c8a
The CPure theory -- Pure with alternative application syntax.
wenzelm
parents:
diff
changeset
|
5 |
*) |
|
3a55e6e26c8a
The CPure theory -- Pure with alternative application syntax.
wenzelm
parents:
diff
changeset
|
6 |
|
|
3a55e6e26c8a
The CPure theory -- Pure with alternative application syntax.
wenzelm
parents:
diff
changeset
|
7 |
theory CPure |
|
3a55e6e26c8a
The CPure theory -- Pure with alternative application syntax.
wenzelm
parents:
diff
changeset
|
8 |
imports Pure |
|
3a55e6e26c8a
The CPure theory -- Pure with alternative application syntax.
wenzelm
parents:
diff
changeset
|
9 |
begin |
|
3a55e6e26c8a
The CPure theory -- Pure with alternative application syntax.
wenzelm
parents:
diff
changeset
|
10 |
|
| 18709 | 11 |
setup {*
|
12 |
Theory.del_modesyntax Syntax.default_mode Syntax.appl_syntax #> |
|
13 |
Theory.add_syntax Syntax.applC_syntax |
|
14 |
*} |
|
|
15804
3a55e6e26c8a
The CPure theory -- Pure with alternative application syntax.
wenzelm
parents:
diff
changeset
|
15 |
|
|
3a55e6e26c8a
The CPure theory -- Pure with alternative application syntax.
wenzelm
parents:
diff
changeset
|
16 |
end |