Nominal/Ex/Ex1rec.thy
author Cezary Kaliszyk <kaliszyk@in.tum.de>
Thu, 29 Apr 2010 10:59:08 +0200
changeset 1981 9f9c4965b608
parent 1946 3395e87d94b8
child 2056 a58c73e39479
permissions -rw-r--r--
Include support of unknown datatypes in new fv
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1596
c69d9fb16785 Move Ex1 and Ex2 out of Test
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     1
theory Ex1rec
1773
c0eac04ae3b4 added README and moved examples into separate directory
Christian Urban <urbanc@in.tum.de>
parents: 1706
diff changeset
     2
imports "../Parser" 
1596
c69d9fb16785 Move Ex1 and Ex2 out of Test
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     3
begin
c69d9fb16785 Move Ex1 and Ex2 out of Test
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     4
c69d9fb16785 Move Ex1 and Ex2 out of Test
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     5
atom_decl name
c69d9fb16785 Move Ex1 and Ex2 out of Test
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     6
c69d9fb16785 Move Ex1 and Ex2 out of Test
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     7
ML {* val _ = recursive := true *}
1706
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1596
diff changeset
     8
ML {* val _ = alpha_type := AlphaGen *}
1946
3395e87d94b8 tuned and made to compile
Christian Urban <urbanc@in.tum.de>
parents: 1773
diff changeset
     9
nominal_datatype lam =
3395e87d94b8 tuned and made to compile
Christian Urban <urbanc@in.tum.de>
parents: 1773
diff changeset
    10
  VAR "name"
3395e87d94b8 tuned and made to compile
Christian Urban <urbanc@in.tum.de>
parents: 1773
diff changeset
    11
| APP "lam" "lam"
3395e87d94b8 tuned and made to compile
Christian Urban <urbanc@in.tum.de>
parents: 1773
diff changeset
    12
| LAM x::"name" t::"lam"  bind x in t
3395e87d94b8 tuned and made to compile
Christian Urban <urbanc@in.tum.de>
parents: 1773
diff changeset
    13
| LET bp::"bp" t::"lam"   bind "bi bp" in t
3395e87d94b8 tuned and made to compile
Christian Urban <urbanc@in.tum.de>
parents: 1773
diff changeset
    14
and bp =
3395e87d94b8 tuned and made to compile
Christian Urban <urbanc@in.tum.de>
parents: 1773
diff changeset
    15
  BP "name" "lam"
1596
c69d9fb16785 Move Ex1 and Ex2 out of Test
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    16
binder
1946
3395e87d94b8 tuned and made to compile
Christian Urban <urbanc@in.tum.de>
parents: 1773
diff changeset
    17
  bi::"bp \<Rightarrow> atom set"
1596
c69d9fb16785 Move Ex1 and Ex2 out of Test
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    18
where
1946
3395e87d94b8 tuned and made to compile
Christian Urban <urbanc@in.tum.de>
parents: 1773
diff changeset
    19
  "bi (BP x t) = {atom x}"
1596
c69d9fb16785 Move Ex1 and Ex2 out of Test
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    20
1946
3395e87d94b8 tuned and made to compile
Christian Urban <urbanc@in.tum.de>
parents: 1773
diff changeset
    21
thm lam_bp.fv
3395e87d94b8 tuned and made to compile
Christian Urban <urbanc@in.tum.de>
parents: 1773
diff changeset
    22
thm lam_bp.eq_iff[no_vars]
3395e87d94b8 tuned and made to compile
Christian Urban <urbanc@in.tum.de>
parents: 1773
diff changeset
    23
thm lam_bp.bn
3395e87d94b8 tuned and made to compile
Christian Urban <urbanc@in.tum.de>
parents: 1773
diff changeset
    24
thm lam_bp.perm
3395e87d94b8 tuned and made to compile
Christian Urban <urbanc@in.tum.de>
parents: 1773
diff changeset
    25
thm lam_bp.induct
3395e87d94b8 tuned and made to compile
Christian Urban <urbanc@in.tum.de>
parents: 1773
diff changeset
    26
thm lam_bp.inducts
3395e87d94b8 tuned and made to compile
Christian Urban <urbanc@in.tum.de>
parents: 1773
diff changeset
    27
thm lam_bp.distinct
3395e87d94b8 tuned and made to compile
Christian Urban <urbanc@in.tum.de>
parents: 1773
diff changeset
    28
ML {* Sign.of_sort @{theory} (@{typ lam}, @{sort fs}) *}
3395e87d94b8 tuned and made to compile
Christian Urban <urbanc@in.tum.de>
parents: 1773
diff changeset
    29
thm lam_bp.fv[simplified lam_bp.supp]
1596
c69d9fb16785 Move Ex1 and Ex2 out of Test
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    30
c69d9fb16785 Move Ex1 and Ex2 out of Test
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    31
end
c69d9fb16785 Move Ex1 and Ex2 out of Test
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    32
c69d9fb16785 Move Ex1 and Ex2 out of Test
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    33
c69d9fb16785 Move Ex1 and Ex2 out of Test
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    34