Nominal/Ex/ExPS7.thy
author Christian Urban <urbanc@in.tum.de>
Mon, 20 Sep 2010 21:52:45 +0800
changeset 2480 ac7dff1194e8
parent 2454 9ffee4eb1ae1
permissions -rw-r--r--
introduced a general procedure for structural inductions; simplified reflexivity proof
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1655
9cec4269b7f9 Extract PS7 and PS8 from Test. PS7 needs the same fix as Core Haskell.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     1
theory ExPS7
2454
9ffee4eb1ae1 renamed NewParser to Nominal2
Christian Urban <urbanc@in.tum.de>
parents: 2436
diff changeset
     2
imports "../Nominal2"
1655
9cec4269b7f9 Extract PS7 and PS8 from Test. PS7 needs the same fix as Core Haskell.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     3
begin
9cec4269b7f9 Extract PS7 and PS8 from Test. PS7 needs the same fix as Core Haskell.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     4
9cec4269b7f9 Extract PS7 and PS8 from Test. PS7 needs the same fix as Core Haskell.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     5
(* example 7 from Peter Sewell's bestiary *)
9cec4269b7f9 Extract PS7 and PS8 from Test. PS7 needs the same fix as Core Haskell.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     6
9cec4269b7f9 Extract PS7 and PS8 from Test. PS7 needs the same fix as Core Haskell.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     7
atom_decl name
9cec4269b7f9 Extract PS7 and PS8 from Test. PS7 needs the same fix as Core Haskell.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     8
2480
ac7dff1194e8 introduced a general procedure for structural inductions; simplified reflexivity proof
Christian Urban <urbanc@in.tum.de>
parents: 2454
diff changeset
     9
declare [[STEPS = 31]]
ac7dff1194e8 introduced a general procedure for structural inductions; simplified reflexivity proof
Christian Urban <urbanc@in.tum.de>
parents: 2454
diff changeset
    10
2082
0854af516f14 cleaned up a bit the examples; added equivariance to all examples
Christian Urban <urbanc@in.tum.de>
parents: 2043
diff changeset
    11
nominal_datatype exp =
0854af516f14 cleaned up a bit the examples; added equivariance to all examples
Christian Urban <urbanc@in.tum.de>
parents: 2043
diff changeset
    12
  Var name
0854af516f14 cleaned up a bit the examples; added equivariance to all examples
Christian Urban <urbanc@in.tum.de>
parents: 2043
diff changeset
    13
| Unit
0854af516f14 cleaned up a bit the examples; added equivariance to all examples
Christian Urban <urbanc@in.tum.de>
parents: 2043
diff changeset
    14
| Pair exp exp
2436
3885dc2669f9 cleaned up (almost completely) the examples
Christian Urban <urbanc@in.tum.de>
parents: 2120
diff changeset
    15
| LetRec l::"lrbs" e::"exp"  bind (set) "bs l" in l e
1655
9cec4269b7f9 Extract PS7 and PS8 from Test. PS7 needs the same fix as Core Haskell.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    16
and lrb =
2082
0854af516f14 cleaned up a bit the examples; added equivariance to all examples
Christian Urban <urbanc@in.tum.de>
parents: 2043
diff changeset
    17
  Assign name exp
1655
9cec4269b7f9 Extract PS7 and PS8 from Test. PS7 needs the same fix as Core Haskell.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    18
and lrbs =
9cec4269b7f9 Extract PS7 and PS8 from Test. PS7 needs the same fix as Core Haskell.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    19
  Single lrb
9cec4269b7f9 Extract PS7 and PS8 from Test. PS7 needs the same fix as Core Haskell.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    20
| More lrb lrbs
9cec4269b7f9 Extract PS7 and PS8 from Test. PS7 needs the same fix as Core Haskell.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    21
binder
2082
0854af516f14 cleaned up a bit the examples; added equivariance to all examples
Christian Urban <urbanc@in.tum.de>
parents: 2043
diff changeset
    22
  b :: "lrb \<Rightarrow> atom set" and
0854af516f14 cleaned up a bit the examples; added equivariance to all examples
Christian Urban <urbanc@in.tum.de>
parents: 2043
diff changeset
    23
  bs :: "lrbs \<Rightarrow> atom set"
1655
9cec4269b7f9 Extract PS7 and PS8 from Test. PS7 needs the same fix as Core Haskell.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    24
where
2082
0854af516f14 cleaned up a bit the examples; added equivariance to all examples
Christian Urban <urbanc@in.tum.de>
parents: 2043
diff changeset
    25
  "b (Assign x e) = {atom x}"
0854af516f14 cleaned up a bit the examples; added equivariance to all examples
Christian Urban <urbanc@in.tum.de>
parents: 2043
diff changeset
    26
| "bs (Single a) = b a"
0854af516f14 cleaned up a bit the examples; added equivariance to all examples
Christian Urban <urbanc@in.tum.de>
parents: 2043
diff changeset
    27
| "bs (More a as) = (b a) \<union> (bs as)"
0854af516f14 cleaned up a bit the examples; added equivariance to all examples
Christian Urban <urbanc@in.tum.de>
parents: 2043
diff changeset
    28
0854af516f14 cleaned up a bit the examples; added equivariance to all examples
Christian Urban <urbanc@in.tum.de>
parents: 2043
diff changeset
    29
1655
9cec4269b7f9 Extract PS7 and PS8 from Test. PS7 needs the same fix as Core Haskell.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    30
9cec4269b7f9 Extract PS7 and PS8 from Test. PS7 needs the same fix as Core Haskell.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    31
end
9cec4269b7f9 Extract PS7 and PS8 from Test. PS7 needs the same fix as Core Haskell.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    32
9cec4269b7f9 Extract PS7 and PS8 from Test. PS7 needs the same fix as Core Haskell.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    33
9cec4269b7f9 Extract PS7 and PS8 from Test. PS7 needs the same fix as Core Haskell.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    34