Nominal/Ex/Multi_Recs.thy
author Christian Urban <urbanc@in.tum.de>
Sat, 17 Dec 2011 17:08:47 +0000
branchNominal2-Isabelle2011-1
changeset 3071 11f6a561eb4b
parent 2950 0911cb7bf696
permissions -rw-r--r--
cleaned examples for stable branch
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
2481
3a5ebb2fcdbf made supp proofs more robust by not using the standard induction; renamed some example files
Christian Urban <urbanc@in.tum.de>
parents: 2480
diff changeset
     1
theory Multi_Recs
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
2481
3a5ebb2fcdbf made supp proofs more robust by not using the standard induction; renamed some example files
Christian Urban <urbanc@in.tum.de>
parents: 2480
diff changeset
     5
(* 
3a5ebb2fcdbf made supp proofs more robust by not using the standard induction; renamed some example files
Christian Urban <urbanc@in.tum.de>
parents: 2480
diff changeset
     6
  multiple recursive binders
3a5ebb2fcdbf made supp proofs more robust by not using the standard induction; renamed some example files
Christian Urban <urbanc@in.tum.de>
parents: 2480
diff changeset
     7
3a5ebb2fcdbf made supp proofs more robust by not using the standard induction; renamed some example files
Christian Urban <urbanc@in.tum.de>
parents: 2480
diff changeset
     8
  example 7 from Peter Sewell's bestiary 
3a5ebb2fcdbf made supp proofs more robust by not using the standard induction; renamed some example files
Christian Urban <urbanc@in.tum.de>
parents: 2480
diff changeset
     9
3a5ebb2fcdbf made supp proofs more robust by not using the standard induction; renamed some example files
Christian Urban <urbanc@in.tum.de>
parents: 2480
diff changeset
    10
*)
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
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
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
    13
2481
3a5ebb2fcdbf made supp proofs more robust by not using the standard induction; renamed some example files
Christian Urban <urbanc@in.tum.de>
parents: 2480
diff changeset
    14
nominal_datatype multi_recs: exp =
2082
0854af516f14 cleaned up a bit the examples; added equivariance to all examples
Christian Urban <urbanc@in.tum.de>
parents: 2043
diff changeset
    15
  Var name
0854af516f14 cleaned up a bit the examples; added equivariance to all examples
Christian Urban <urbanc@in.tum.de>
parents: 2043
diff changeset
    16
| Unit
0854af516f14 cleaned up a bit the examples; added equivariance to all examples
Christian Urban <urbanc@in.tum.de>
parents: 2043
diff changeset
    17
| Pair exp exp
2950
0911cb7bf696 changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
Christian Urban <urbanc@in.tum.de>
parents: 2616
diff changeset
    18
| LetRec l::"lrbs" e::"exp"  binds (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
    19
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
    20
  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
    21
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
    22
  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
    23
| 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
    24
binder
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 :: "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
    26
  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
    27
where
2082
0854af516f14 cleaned up a bit the examples; added equivariance to all examples
Christian Urban <urbanc@in.tum.de>
parents: 2043
diff changeset
    28
  "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
    29
| "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
    30
| "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
    31
2481
3a5ebb2fcdbf made supp proofs more robust by not using the standard induction; renamed some example files
Christian Urban <urbanc@in.tum.de>
parents: 2480
diff changeset
    32
thm multi_recs.distinct
3a5ebb2fcdbf made supp proofs more robust by not using the standard induction; renamed some example files
Christian Urban <urbanc@in.tum.de>
parents: 2480
diff changeset
    33
thm multi_recs.induct
3a5ebb2fcdbf made supp proofs more robust by not using the standard induction; renamed some example files
Christian Urban <urbanc@in.tum.de>
parents: 2480
diff changeset
    34
thm multi_recs.inducts
3a5ebb2fcdbf made supp proofs more robust by not using the standard induction; renamed some example files
Christian Urban <urbanc@in.tum.de>
parents: 2480
diff changeset
    35
thm multi_recs.exhaust
3a5ebb2fcdbf made supp proofs more robust by not using the standard induction; renamed some example files
Christian Urban <urbanc@in.tum.de>
parents: 2480
diff changeset
    36
thm multi_recs.fv_defs
3a5ebb2fcdbf made supp proofs more robust by not using the standard induction; renamed some example files
Christian Urban <urbanc@in.tum.de>
parents: 2480
diff changeset
    37
thm multi_recs.bn_defs
3a5ebb2fcdbf made supp proofs more robust by not using the standard induction; renamed some example files
Christian Urban <urbanc@in.tum.de>
parents: 2480
diff changeset
    38
thm multi_recs.perm_simps
3a5ebb2fcdbf made supp proofs more robust by not using the standard induction; renamed some example files
Christian Urban <urbanc@in.tum.de>
parents: 2480
diff changeset
    39
thm multi_recs.eq_iff
3a5ebb2fcdbf made supp proofs more robust by not using the standard induction; renamed some example files
Christian Urban <urbanc@in.tum.de>
parents: 2480
diff changeset
    40
thm multi_recs.fv_bn_eqvt
3a5ebb2fcdbf made supp proofs more robust by not using the standard induction; renamed some example files
Christian Urban <urbanc@in.tum.de>
parents: 2480
diff changeset
    41
thm multi_recs.size_eqvt
3a5ebb2fcdbf made supp proofs more robust by not using the standard induction; renamed some example files
Christian Urban <urbanc@in.tum.de>
parents: 2480
diff changeset
    42
thm multi_recs.supports
3a5ebb2fcdbf made supp proofs more robust by not using the standard induction; renamed some example files
Christian Urban <urbanc@in.tum.de>
parents: 2480
diff changeset
    43
thm multi_recs.fsupp
3a5ebb2fcdbf made supp proofs more robust by not using the standard induction; renamed some example files
Christian Urban <urbanc@in.tum.de>
parents: 2480
diff changeset
    44
thm multi_recs.supp
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
    45
2615
d5713db7e146 one interesting case done
Christian Urban <urbanc@in.tum.de>
parents: 2481
diff changeset
    46
thm multi_recs.bn_defs
d5713db7e146 one interesting case done
Christian Urban <urbanc@in.tum.de>
parents: 2481
diff changeset
    47
thm multi_recs.permute_bn
d5713db7e146 one interesting case done
Christian Urban <urbanc@in.tum.de>
parents: 2481
diff changeset
    48
thm multi_recs.perm_bn_alpha
d5713db7e146 one interesting case done
Christian Urban <urbanc@in.tum.de>
parents: 2481
diff changeset
    49
thm multi_recs.perm_bn_simps
d5713db7e146 one interesting case done
Christian Urban <urbanc@in.tum.de>
parents: 2481
diff changeset
    50
thm multi_recs.bn_finite
d5713db7e146 one interesting case done
Christian Urban <urbanc@in.tum.de>
parents: 2481
diff changeset
    51
d5713db7e146 one interesting case done
Christian Urban <urbanc@in.tum.de>
parents: 2481
diff changeset
    52
d5713db7e146 one interesting case done
Christian Urban <urbanc@in.tum.de>
parents: 2481
diff changeset
    53
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
    54
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
    55
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
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