Nominal/Ex/ExPS8.thy
author Cezary Kaliszyk <kaliszyk@in.tum.de>
Mon, 10 May 2010 18:29:45 +0200
changeset 2095 ae94bae5bb93
parent 2082 0854af516f14
child 2104 2205b572bc9b
permissions -rw-r--r--
Restore set bindings in CoreHaskell
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
2052
ca512f9c0b0a Move ExPS8 to new parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1941
diff changeset
     9
ML {* val _ = cheat_fv_rsp := true *}
ca512f9c0b0a Move ExPS8 to new parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1941
diff changeset
    10
ML {* val _ = cheat_alpha_bn_rsp := true *}
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
    11
1941
Christian Urban <urbanc@in.tum.de>
parents: 1773
diff changeset
    12
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
    13
  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
    14
| EUnit
1941
Christian Urban <urbanc@in.tum.de>
parents: 1773
diff changeset
    15
| EPair exp exp
2052
ca512f9c0b0a Move ExPS8 to new parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1941
diff changeset
    16
| 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
    17
and fnclause =
2052
ca512f9c0b0a Move ExPS8 to new parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1941
diff changeset
    18
  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
    19
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
    20
  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
    21
| ORs fnclause fnclauses
1941
Christian Urban <urbanc@in.tum.de>
parents: 1773
diff changeset
    22
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
    23
  Clause fnclauses
1941
Christian Urban <urbanc@in.tum.de>
parents: 1773
diff changeset
    24
and lrbs =
Christian Urban <urbanc@in.tum.de>
parents: 1773
diff changeset
    25
  Single lrb
Christian Urban <urbanc@in.tum.de>
parents: 1773
diff changeset
    26
| More lrb lrbs
Christian Urban <urbanc@in.tum.de>
parents: 1773
diff changeset
    27
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
    28
  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
    29
| PUnit
1941
Christian Urban <urbanc@in.tum.de>
parents: 1773
diff changeset
    30
| 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
    31
binder
1941
Christian Urban <urbanc@in.tum.de>
parents: 1773
diff changeset
    32
  b_lrbs :: "lrbs \<Rightarrow> atom set" and
Christian Urban <urbanc@in.tum.de>
parents: 1773
diff changeset
    33
  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
    34
  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
    35
  b_fnclause :: "fnclause \<Rightarrow> atom set" and
1941
Christian Urban <urbanc@in.tum.de>
parents: 1773
diff changeset
    36
  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
    37
where
1941
Christian Urban <urbanc@in.tum.de>
parents: 1773
diff changeset
    38
  "b_lrbs (Single l) = b_lrb l"
Christian Urban <urbanc@in.tum.de>
parents: 1773
diff changeset
    39
| "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
    40
| "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
    41
| "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
    42
| "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
    43
| "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
    44
| "b_fnclauses (ORs fc fcs) = (b_fnclause fc) \<union> (b_fnclauses fcs)"
1941
Christian Urban <urbanc@in.tum.de>
parents: 1773
diff changeset
    45
| "b_lrb (Clause fcs) = (b_fnclauses fcs)"
Christian Urban <urbanc@in.tum.de>
parents: 1773
diff changeset
    46
| "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
    47
2052
ca512f9c0b0a Move ExPS8 to new parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1941
diff changeset
    48
thm exp_fnclause_fnclauses_lrb_lrbs_pat.fv
ca512f9c0b0a Move ExPS8 to new parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1941
diff changeset
    49
thm exp_fnclause_fnclauses_lrb_lrbs_pat.eq_iff
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
    50
2082
0854af516f14 cleaned up a bit the examples; added equivariance to all examples
Christian Urban <urbanc@in.tum.de>
parents: 2052
diff changeset
    51
0854af516f14 cleaned up a bit the examples; added equivariance to all examples
Christian Urban <urbanc@in.tum.de>
parents: 2052
diff changeset
    52
declare permute_exp_raw_permute_fnclause_raw_permute_fnclauses_raw_permute_lrb_raw_permute_lrbs_raw_permute_pat_raw.simps[eqvt]
0854af516f14 cleaned up a bit the examples; added equivariance to all examples
Christian Urban <urbanc@in.tum.de>
parents: 2052
diff changeset
    53
declare alpha_gen_eqvt[eqvt]
0854af516f14 cleaned up a bit the examples; added equivariance to all examples
Christian Urban <urbanc@in.tum.de>
parents: 2052
diff changeset
    54
0854af516f14 cleaned up a bit the examples; added equivariance to all examples
Christian Urban <urbanc@in.tum.de>
parents: 2052
diff changeset
    55
equivariance alpha_exp_raw
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
    56
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
    57
9cec4269b7f9 Extract PS7 and PS8 from Test. PS7 needs the same fix as Core Haskell.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    58
9cec4269b7f9 Extract PS7 and PS8 from Test. PS7 needs the same fix as Core Haskell.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    59