Nominal/Ex/ExPS8.thy
author Christian Urban <urbanc@in.tum.de>
Sun, 29 Aug 2010 12:14:40 +0800
changeset 2452 39f8d405d7a2
parent 2440 0a36825b16c1
child 2454 9ffee4eb1ae1
permissions -rw-r--r--
updated todos
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 ExPS8
2052
ca512f9c0b0a Move ExPS8 to new parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1941
diff changeset
     2
imports "../NewParser"
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 8 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
2440
0a36825b16c1 "isabelle make test" makes all major examples....they work up to supp theorems (excluding)
Christian Urban <urbanc@in.tum.de>
parents: 2438
diff changeset
     9
declare [[STEPS = 100]]
2438
abafea9b39bb corrected bug with fv-function generation (that was the problem with recursive binders)
Christian Urban <urbanc@in.tum.de>
parents: 2436
diff changeset
    10
1941
Christian Urban <urbanc@in.tum.de>
parents: 1773
diff changeset
    11
nominal_datatype 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
    12
  EVar 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
    13
| EUnit
1941
Christian Urban <urbanc@in.tum.de>
parents: 1773
diff changeset
    14
| EPair exp exp
2436
3885dc2669f9 cleaned up (almost completely) the examples
Christian Urban <urbanc@in.tum.de>
parents: 2120
diff changeset
    15
| ELetRec l::lrbs e::exp bind (set) "b_lrbs 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 fnclause =
2436
3885dc2669f9 cleaned up (almost completely) the examples
Christian Urban <urbanc@in.tum.de>
parents: 2120
diff changeset
    17
  K x::name p::pat f::exp bind (set) "b_pat p" in f
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 fnclauses =
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
  S fnclause
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
| ORs fnclause fnclauses
1941
Christian Urban <urbanc@in.tum.de>
parents: 1773
diff changeset
    21
and lrb =
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
    22
  Clause fnclauses
1941
Christian Urban <urbanc@in.tum.de>
parents: 1773
diff changeset
    23
and lrbs =
Christian Urban <urbanc@in.tum.de>
parents: 1773
diff changeset
    24
  Single lrb
Christian Urban <urbanc@in.tum.de>
parents: 1773
diff changeset
    25
| More lrb lrbs
Christian Urban <urbanc@in.tum.de>
parents: 1773
diff changeset
    26
and pat =
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
    27
  PVar 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
    28
| PUnit
1941
Christian Urban <urbanc@in.tum.de>
parents: 1773
diff changeset
    29
| PPair pat pat
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
binder
1941
Christian Urban <urbanc@in.tum.de>
parents: 1773
diff changeset
    31
  b_lrbs :: "lrbs \<Rightarrow> atom set" and
Christian Urban <urbanc@in.tum.de>
parents: 1773
diff changeset
    32
  b_pat :: "pat \<Rightarrow> atom set" and
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
    33
  b_fnclauses :: "fnclauses \<Rightarrow> atom set" and
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
  b_fnclause :: "fnclause \<Rightarrow> atom set" and
1941
Christian Urban <urbanc@in.tum.de>
parents: 1773
diff changeset
    35
  b_lrb :: "lrb \<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
    36
where
1941
Christian Urban <urbanc@in.tum.de>
parents: 1773
diff changeset
    37
  "b_lrbs (Single l) = b_lrb l"
Christian Urban <urbanc@in.tum.de>
parents: 1773
diff changeset
    38
| "b_lrbs (More l ls) = b_lrb l \<union> b_lrbs ls"
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
    39
| "b_pat (PVar x) = {atom x}"
9cec4269b7f9 Extract PS7 and PS8 from Test. PS7 needs the same fix as Core Haskell.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    40
| "b_pat (PUnit) = {}"
9cec4269b7f9 Extract PS7 and PS8 from Test. PS7 needs the same fix as Core Haskell.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    41
| "b_pat (PPair p1 p2) = b_pat p1 \<union> b_pat p2"
9cec4269b7f9 Extract PS7 and PS8 from Test. PS7 needs the same fix as Core Haskell.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    42
| "b_fnclauses (S fc) = (b_fnclause fc)"
9cec4269b7f9 Extract PS7 and PS8 from Test. PS7 needs the same fix as Core Haskell.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    43
| "b_fnclauses (ORs fc fcs) = (b_fnclause fc) \<union> (b_fnclauses fcs)"
1941
Christian Urban <urbanc@in.tum.de>
parents: 1773
diff changeset
    44
| "b_lrb (Clause fcs) = (b_fnclauses fcs)"
Christian Urban <urbanc@in.tum.de>
parents: 1773
diff changeset
    45
| "b_fnclause (K x pat exp) = {atom x}"
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
    46
9cec4269b7f9 Extract PS7 and PS8 from Test. PS7 needs the same fix as Core Haskell.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    47
9cec4269b7f9 Extract PS7 and PS8 from Test. PS7 needs the same fix as Core Haskell.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    48
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
    49
9cec4269b7f9 Extract PS7 and PS8 from Test. PS7 needs the same fix as Core Haskell.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    50
9cec4269b7f9 Extract PS7 and PS8 from Test. PS7 needs the same fix as Core Haskell.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    51