Nominal/Ex/Foo1.thy
author Christian Urban <urbanc@in.tum.de>
Tue, 28 Sep 2010 05:56:11 -0400
changeset 2494 11133eb76f61
child 2500 3b6a70e73006
permissions -rw-r--r--
added Foo1 to explore a contrived example
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
2494
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     1
theory Foo1
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     2
imports "../Nominal2" 
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     3
begin
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     4
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     5
(* 
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     6
  Contrived example that has more than one
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     7
  binding function for a datatype
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     8
*)
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     9
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    10
atom_decl name
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    11
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    12
nominal_datatype foo: trm =
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    13
  Var "name"
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    14
| App "trm" "trm"
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    15
| Lam x::"name" t::"trm"  bind x in t
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    16
| Let1 a::"assg" t::"trm"  bind "bn1 a" in t
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    17
| Let2 a::"assg" t::"trm"  bind "bn2 a" in t
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    18
| Let3 a::"assg" t::"trm"  bind "bn3 a" in t
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    19
and assg =
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    20
  As "name" "name" "trm"
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    21
binder
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    22
  bn1::"assg \<Rightarrow> atom list" and
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    23
  bn2::"assg \<Rightarrow> atom list" and
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    24
  bn3::"assg \<Rightarrow> atom list"
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    25
where
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    26
  "bn1 (As x y t) = [atom x]"
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    27
| "bn2 (As x y t) = [atom y]"
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    28
| "bn3 (As x y t) = [atom x, atom y]"
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    29
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    30
thm foo.distinct
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    31
thm foo.induct
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    32
thm foo.inducts
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    33
thm foo.exhaust
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    34
thm foo.fv_defs
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    35
thm foo.bn_defs
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    36
thm foo.perm_simps
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    37
thm foo.eq_iff
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    38
thm foo.fv_bn_eqvt
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    39
thm foo.size_eqvt
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    40
thm foo.supports
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    41
thm foo.fsupp
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    42
thm foo.supp
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    43
thm foo.fresh
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    44
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    45
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    46
end
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    47
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    48
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    49