|
18425
|
1 |
|
|
18661
|
2 |
theory Class
|
|
18425
|
3 |
imports "../nominal"
|
|
|
4 |
begin
|
|
18395
|
5 |
|
|
18881
|
6 |
section {* Term-Calculus from Urban's PhD *}
|
|
|
7 |
|
|
18395
|
8 |
atom_decl name coname
|
|
|
9 |
|
|
18881
|
10 |
nominal_datatype trm =
|
|
|
11 |
Ax "name" "coname"
|
|
|
12 |
| Cut "\<guillemotleft>coname\<guillemotright>trm" "\<guillemotleft>name\<guillemotright>trm" ("Cut <_>._ (_)._" [100,100,100,100] 100)
|
|
|
13 |
| NotR "\<guillemotleft>name\<guillemotright>trm" "coname" ("NotR (_)._ _" [100,100,100] 100)
|
|
|
14 |
| NotL "\<guillemotleft>coname\<guillemotright>trm" "name" ("NotL <_>._ _" [100,100,100] 100)
|
|
|
15 |
| AndR "\<guillemotleft>coname\<guillemotright>trm" "\<guillemotleft>coname\<guillemotright>trm" "coname" ("AndR <_>._ <_>._ _" [100,100,100,100,100] 100)
|
|
|
16 |
| AndL1 "\<guillemotleft>name\<guillemotright>trm" "name" ("AndL1 (_)._ _" [100,100,100] 100)
|
|
|
17 |
| AndL2 "\<guillemotleft>name\<guillemotright>trm" "name" ("AndL2 (_)._ _" [100,100,100] 100)
|
|
|
18 |
| OrR1 "\<guillemotleft>coname\<guillemotright>trm" "coname" ("OrR1 <_>._ _" [100,100,100] 100)
|
|
|
19 |
| OrR2 "\<guillemotleft>coname\<guillemotright>trm" "coname" ("OrR2 <_>._ _" [100,100,100] 100)
|
|
|
20 |
| OrL "\<guillemotleft>name\<guillemotright>trm" "\<guillemotleft>name\<guillemotright>trm" "name" ("OrL (_)._ (_)._ _" [100,100,100,100,100] 100)
|
|
|
21 |
| ImpR "\<guillemotleft>name\<guillemotright>(\<guillemotleft>coname\<guillemotright>trm)" "coname" ("ImpR (_).<_>._ _" [100,100,100,100] 100)
|
|
|
22 |
| ImpL "\<guillemotleft>coname\<guillemotright>trm" "\<guillemotleft>name\<guillemotright>trm" "name" ("ImpL <_>._ (_)._ _" [100,100,100,100,100] 100)
|
|
|
23 |
|
|
|
24 |
|
|
|
25 |
text {* Induction principles *}
|
|
18395
|
26 |
|
|
18881
|
27 |
thm trm.induct_weak --"weak"
|
|
|
28 |
thm trm.induct --"strong"
|
|
|
29 |
thm trm.induct' --"strong with explicit context (rarely needed)"
|
|
|
30 |
|
|
|
31 |
text {* named terms *}
|
|
|
32 |
|
|
|
33 |
nominal_datatype ntrm = N "\<guillemotleft>name\<guillemotright>trm" ("N (_)._" [100,100] 100)
|
|
|
34 |
|
|
|
35 |
text {* conamed terms *}
|
|
18395
|
36 |
|
|
18881
|
37 |
nominal_datatype ctrm = C "\<guillemotleft>coname\<guillemotright>trm" ("C <_>._" [100,100] 100)
|
|
|
38 |
|
|
|
39 |
text {* We should now define the two forms of substitition :o( *}
|
|
18395
|
40 |
|
|
18881
|
41 |
consts
|
|
|
42 |
substn :: "trm \<Rightarrow> name \<Rightarrow> ctrm \<Rightarrow> trm" ("_[_::=_]" [100,100,100] 100)
|
|
|
43 |
substc :: "trm \<Rightarrow> coname \<Rightarrow> ntrm \<Rightarrow> trm" ("_[_::=_]" [100,100,100] 100)
|
|
|
44 |
|
|
|
45 |
text {* does not work yet *} |