Nominal/Ex/Modules.thy
author Christian Urban <urbanc@in.tum.de>
Tue, 18 Jan 2011 21:26:58 +0100
changeset 2676 028d5511c15f
parent 2593 25dcb2b1329e
child 2950 0911cb7bf696
permissions -rw-r--r--
some tryes about substitution over type-schemes
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
2083
9568f9f31822 tuned file names for examples
Christian Urban <urbanc@in.tum.de>
parents: 2082
diff changeset
     1
theory Modules
2454
9ffee4eb1ae1 renamed NewParser to Nominal2
Christian Urban <urbanc@in.tum.de>
parents: 2436
diff changeset
     2
imports "../Nominal2"
1599
8b5a1ad60487 Move Leroy out of Test, rename accordingly.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     3
begin
8b5a1ad60487 Move Leroy out of Test, rename accordingly.
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
   example from Leroy'96 about modules
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
   see OTT example by Owens 
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
*)
1599
8b5a1ad60487 Move Leroy out of Test, rename accordingly.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    10
8b5a1ad60487 Move Leroy out of Test, rename accordingly.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    11
atom_decl name
8b5a1ad60487 Move Leroy out of Test, rename accordingly.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    12
2436
3885dc2669f9 cleaned up (almost completely) the examples
Christian Urban <urbanc@in.tum.de>
parents: 2120
diff changeset
    13
nominal_datatype modules: 
3885dc2669f9 cleaned up (almost completely) the examples
Christian Urban <urbanc@in.tum.de>
parents: 2120
diff changeset
    14
 mexp =
1599
8b5a1ad60487 Move Leroy out of Test, rename accordingly.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    15
  Acc "path"
8b5a1ad60487 Move Leroy out of Test, rename accordingly.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    16
| Stru "body"
2436
3885dc2669f9 cleaned up (almost completely) the examples
Christian Urban <urbanc@in.tum.de>
parents: 2120
diff changeset
    17
| Funct x::"name" "sexp" m::"mexp"    bind (set) x in m
1599
8b5a1ad60487 Move Leroy out of Test, rename accordingly.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    18
| FApp "mexp" "path"
8b5a1ad60487 Move Leroy out of Test, rename accordingly.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    19
| Ascr "mexp" "sexp"
8b5a1ad60487 Move Leroy out of Test, rename accordingly.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    20
and body =
8b5a1ad60487 Move Leroy out of Test, rename accordingly.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    21
  Empty
2436
3885dc2669f9 cleaned up (almost completely) the examples
Christian Urban <urbanc@in.tum.de>
parents: 2120
diff changeset
    22
| Seq c::"defn" d::"body"     bind (set) "cbinders c" in d
1599
8b5a1ad60487 Move Leroy out of Test, rename accordingly.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    23
and defn =
2082
0854af516f14 cleaned up a bit the examples; added equivariance to all examples
Christian Urban <urbanc@in.tum.de>
parents: 2042
diff changeset
    24
  Type "name" "ty"
1599
8b5a1ad60487 Move Leroy out of Test, rename accordingly.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    25
| Dty "name"
8b5a1ad60487 Move Leroy out of Test, rename accordingly.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    26
| DStru "name" "mexp"
2082
0854af516f14 cleaned up a bit the examples; added equivariance to all examples
Christian Urban <urbanc@in.tum.de>
parents: 2042
diff changeset
    27
| Val "name" "trm"
1599
8b5a1ad60487 Move Leroy out of Test, rename accordingly.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    28
and sexp =
8b5a1ad60487 Move Leroy out of Test, rename accordingly.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    29
  Sig sbody
8b5a1ad60487 Move Leroy out of Test, rename accordingly.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    30
| SFunc "name" "sexp" "sexp"
8b5a1ad60487 Move Leroy out of Test, rename accordingly.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    31
and sbody =
8b5a1ad60487 Move Leroy out of Test, rename accordingly.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    32
  SEmpty
2436
3885dc2669f9 cleaned up (almost completely) the examples
Christian Urban <urbanc@in.tum.de>
parents: 2120
diff changeset
    33
| SSeq C::"spec" D::"sbody"    bind (set) "Cbinders C" in D
1599
8b5a1ad60487 Move Leroy out of Test, rename accordingly.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    34
and spec =
8b5a1ad60487 Move Leroy out of Test, rename accordingly.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    35
  Type1 "name"
2082
0854af516f14 cleaned up a bit the examples; added equivariance to all examples
Christian Urban <urbanc@in.tum.de>
parents: 2042
diff changeset
    36
| Type2 "name" "ty"
1599
8b5a1ad60487 Move Leroy out of Test, rename accordingly.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    37
| SStru "name" "sexp"
2082
0854af516f14 cleaned up a bit the examples; added equivariance to all examples
Christian Urban <urbanc@in.tum.de>
parents: 2042
diff changeset
    38
| SVal "name" "ty"
0854af516f14 cleaned up a bit the examples; added equivariance to all examples
Christian Urban <urbanc@in.tum.de>
parents: 2042
diff changeset
    39
and ty =
1599
8b5a1ad60487 Move Leroy out of Test, rename accordingly.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    40
  Tyref1 "name"
2082
0854af516f14 cleaned up a bit the examples; added equivariance to all examples
Christian Urban <urbanc@in.tum.de>
parents: 2042
diff changeset
    41
| Tyref2 "path" "ty"
0854af516f14 cleaned up a bit the examples; added equivariance to all examples
Christian Urban <urbanc@in.tum.de>
parents: 2042
diff changeset
    42
| Fun "ty" "ty"
1599
8b5a1ad60487 Move Leroy out of Test, rename accordingly.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    43
and path =
8b5a1ad60487 Move Leroy out of Test, rename accordingly.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    44
  Sref1 "name"
8b5a1ad60487 Move Leroy out of Test, rename accordingly.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    45
| Sref2 "path" "name"
2082
0854af516f14 cleaned up a bit the examples; added equivariance to all examples
Christian Urban <urbanc@in.tum.de>
parents: 2042
diff changeset
    46
and trm =
1599
8b5a1ad60487 Move Leroy out of Test, rename accordingly.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    47
  Tref1 "name"
8b5a1ad60487 Move Leroy out of Test, rename accordingly.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    48
| Tref2 "path" "name"
2436
3885dc2669f9 cleaned up (almost completely) the examples
Christian Urban <urbanc@in.tum.de>
parents: 2120
diff changeset
    49
| Lam' v::"name" "ty" M::"trm"  bind (set) v in M
2082
0854af516f14 cleaned up a bit the examples; added equivariance to all examples
Christian Urban <urbanc@in.tum.de>
parents: 2042
diff changeset
    50
| App' "trm" "trm"
0854af516f14 cleaned up a bit the examples; added equivariance to all examples
Christian Urban <urbanc@in.tum.de>
parents: 2042
diff changeset
    51
| Let' "body" "trm"
1599
8b5a1ad60487 Move Leroy out of Test, rename accordingly.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    52
binder
8b5a1ad60487 Move Leroy out of Test, rename accordingly.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    53
    cbinders :: "defn \<Rightarrow> atom set"
8b5a1ad60487 Move Leroy out of Test, rename accordingly.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    54
and Cbinders :: "spec \<Rightarrow> atom set"
8b5a1ad60487 Move Leroy out of Test, rename accordingly.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    55
where
8b5a1ad60487 Move Leroy out of Test, rename accordingly.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    56
  "cbinders (Type t T) = {atom t}"
8b5a1ad60487 Move Leroy out of Test, rename accordingly.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    57
| "cbinders (Dty t) = {atom t}"
8b5a1ad60487 Move Leroy out of Test, rename accordingly.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    58
| "cbinders (DStru x s) = {atom x}"
8b5a1ad60487 Move Leroy out of Test, rename accordingly.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    59
| "cbinders (Val v M) = {atom v}"
8b5a1ad60487 Move Leroy out of Test, rename accordingly.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    60
| "Cbinders (Type1 t) = {atom t}"
8b5a1ad60487 Move Leroy out of Test, rename accordingly.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    61
| "Cbinders (Type2 t T) = {atom t}"
8b5a1ad60487 Move Leroy out of Test, rename accordingly.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    62
| "Cbinders (SStru x S) = {atom x}"
8b5a1ad60487 Move Leroy out of Test, rename accordingly.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    63
| "Cbinders (SVal v T) = {atom v}"
8b5a1ad60487 Move Leroy out of Test, rename accordingly.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    64
2436
3885dc2669f9 cleaned up (almost completely) the examples
Christian Urban <urbanc@in.tum.de>
parents: 2120
diff changeset
    65
2593
25dcb2b1329e ordered raw_bn_info to agree with the order of the raw_bn_functions; started alpha_bn proof
Christian Urban <urbanc@in.tum.de>
parents: 2481
diff changeset
    66
thm modules.perm_bn_alpha
25dcb2b1329e ordered raw_bn_info to agree with the order of the raw_bn_functions; started alpha_bn proof
Christian Urban <urbanc@in.tum.de>
parents: 2481
diff changeset
    67
thm modules.perm_bn_simps
25dcb2b1329e ordered raw_bn_info to agree with the order of the raw_bn_functions; started alpha_bn proof
Christian Urban <urbanc@in.tum.de>
parents: 2481
diff changeset
    68
thm modules.bn_finite
25dcb2b1329e ordered raw_bn_info to agree with the order of the raw_bn_functions; started alpha_bn proof
Christian Urban <urbanc@in.tum.de>
parents: 2481
diff changeset
    69
2436
3885dc2669f9 cleaned up (almost completely) the examples
Christian Urban <urbanc@in.tum.de>
parents: 2120
diff changeset
    70
thm modules.distinct
3885dc2669f9 cleaned up (almost completely) the examples
Christian Urban <urbanc@in.tum.de>
parents: 2120
diff changeset
    71
thm modules.induct
3885dc2669f9 cleaned up (almost completely) the examples
Christian Urban <urbanc@in.tum.de>
parents: 2120
diff changeset
    72
thm modules.exhaust
3885dc2669f9 cleaned up (almost completely) the examples
Christian Urban <urbanc@in.tum.de>
parents: 2120
diff changeset
    73
thm modules.fv_defs
3885dc2669f9 cleaned up (almost completely) the examples
Christian Urban <urbanc@in.tum.de>
parents: 2120
diff changeset
    74
thm modules.bn_defs
3885dc2669f9 cleaned up (almost completely) the examples
Christian Urban <urbanc@in.tum.de>
parents: 2120
diff changeset
    75
thm modules.perm_simps
3885dc2669f9 cleaned up (almost completely) the examples
Christian Urban <urbanc@in.tum.de>
parents: 2120
diff changeset
    76
thm modules.eq_iff
3885dc2669f9 cleaned up (almost completely) the examples
Christian Urban <urbanc@in.tum.de>
parents: 2120
diff changeset
    77
thm modules.fv_bn_eqvt
3885dc2669f9 cleaned up (almost completely) the examples
Christian Urban <urbanc@in.tum.de>
parents: 2120
diff changeset
    78
thm modules.size_eqvt
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
    79
thm modules.supp
2082
0854af516f14 cleaned up a bit the examples; added equivariance to all examples
Christian Urban <urbanc@in.tum.de>
parents: 2042
diff changeset
    80
1599
8b5a1ad60487 Move Leroy out of Test, rename accordingly.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    81
8b5a1ad60487 Move Leroy out of Test, rename accordingly.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    82
end
8b5a1ad60487 Move Leroy out of Test, rename accordingly.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    83
8b5a1ad60487 Move Leroy out of Test, rename accordingly.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    84
8b5a1ad60487 Move Leroy out of Test, rename accordingly.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    85