| author | Cezary Kaliszyk <kaliszyk@in.tum.de> | 
| Tue, 23 Mar 2010 08:46:44 +0100 | |
| changeset 1598 | 2406350c358f | 
| parent 1597 | af34dd3418fe | 
| child 1599 | 8b5a1ad60487 | 
| permissions | -rw-r--r-- | 
| 954 
c009d2535896
very rough example file for how nominal2 specification can be parsed
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1 | theory Test | 
| 1596 
c69d9fb16785
Move Ex1 and Ex2 out of Test
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1595diff
changeset | 2 | imports "Parser" | 
| 954 
c009d2535896
very rough example file for how nominal2 specification can be parsed
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 3 | begin | 
| 
c009d2535896
very rough example file for how nominal2 specification can be parsed
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 4 | |
| 1428 
4029105011ca
Show that the new types are in finite support typeclass.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1418diff
changeset | 5 | atom_decl name | 
| 
4029105011ca
Show that the new types are in finite support typeclass.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1418diff
changeset | 6 | |
| 1465 
4de35639fef0
added another supp-proof for the non-recursive case
 Christian Urban <urbanc@in.tum.de> parents: 
1461diff
changeset | 7 | ML {* val _ = recursive := false  *}
 | 
| 1367 
9bbf56cdeb88
added compat definitions to some examples
 Christian Urban <urbanc@in.tum.de> parents: 
1365diff
changeset | 8 | |
| 
9bbf56cdeb88
added compat definitions to some examples
 Christian Urban <urbanc@in.tum.de> parents: 
1365diff
changeset | 9 | text {* example 3 from Terms.thy *}
 | 
| 1251 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 10 | |
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 11 | nominal_datatype trm3 = | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 12 | Vr3 "name" | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 13 | | Ap3 "trm3" "trm3" | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 14 | | Lm3 x::"name" t::"trm3" bind x in t | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 15 | | Lt3 r::"rassigns3" t::"trm3" bind "bv3 r" in t | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 16 | and rassigns3 = | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 17 | ANil | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 18 | | ACons "name" "trm3" "rassigns3" | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 19 | binder | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 20 | bv3 | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 21 | where | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 22 |   "bv3 ANil = {}"
 | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 23 | | "bv3 (ACons x t as) = {atom x} \<union> (bv3 as)"
 | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 24 | |
| 1515 
76fa21f27f22
Rename "_property" to ".property"
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1510diff
changeset | 25 | thm trm3_rassigns3.fv | 
| 
76fa21f27f22
Rename "_property" to ".property"
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1510diff
changeset | 26 | thm trm3_rassigns3.eq_iff | 
| 
76fa21f27f22
Rename "_property" to ".property"
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1510diff
changeset | 27 | thm trm3_rassigns3.bn | 
| 
76fa21f27f22
Rename "_property" to ".property"
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1510diff
changeset | 28 | thm trm3_rassigns3.perm | 
| 
76fa21f27f22
Rename "_property" to ".property"
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1510diff
changeset | 29 | thm trm3_rassigns3.induct | 
| 
76fa21f27f22
Rename "_property" to ".property"
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1510diff
changeset | 30 | thm trm3_rassigns3.distinct | 
| 1553 
4355eb3b7161
Automatically derive support for datatypes with at-most one binding per constructor.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1549diff
changeset | 31 | thm trm3_rassigns3.fv[simplified trm3_rassigns3.supp] | 
| 1367 
9bbf56cdeb88
added compat definitions to some examples
 Christian Urban <urbanc@in.tum.de> parents: 
1365diff
changeset | 32 | |
| 1251 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 33 | (* example 5 from Terms.thy *) | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 34 | |
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 35 | nominal_datatype trm5 = | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 36 | Vr5 "name" | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 37 | | Ap5 "trm5" "trm5" | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 38 | | Lt5 l::"lts" t::"trm5" bind "bv5 l" in t | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 39 | and lts = | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 40 | Lnil | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 41 | | Lcons "name" "trm5" "lts" | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 42 | binder | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 43 | bv5 | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 44 | where | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 45 |   "bv5 Lnil = {}"
 | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 46 | | "bv5 (Lcons n t ltl) = {atom n} \<union> (bv5 ltl)"
 | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 47 | |
| 1515 
76fa21f27f22
Rename "_property" to ".property"
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1510diff
changeset | 48 | thm trm5_lts.fv | 
| 
76fa21f27f22
Rename "_property" to ".property"
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1510diff
changeset | 49 | thm trm5_lts.eq_iff | 
| 
76fa21f27f22
Rename "_property" to ".property"
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1510diff
changeset | 50 | thm trm5_lts.bn | 
| 
76fa21f27f22
Rename "_property" to ".property"
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1510diff
changeset | 51 | thm trm5_lts.perm | 
| 
76fa21f27f22
Rename "_property" to ".property"
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1510diff
changeset | 52 | thm trm5_lts.induct | 
| 
76fa21f27f22
Rename "_property" to ".property"
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1510diff
changeset | 53 | thm trm5_lts.distinct | 
| 1553 
4355eb3b7161
Automatically derive support for datatypes with at-most one binding per constructor.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1549diff
changeset | 54 | thm trm5_lts.fv[simplified trm5_lts.supp] | 
| 1285 
e3ac56d3b7af
added example from my phd
 Christian Urban <urbanc@in.tum.de> parents: 
1284diff
changeset | 55 | |
| 
e3ac56d3b7af
added example from my phd
 Christian Urban <urbanc@in.tum.de> parents: 
1284diff
changeset | 56 | (* example form Leroy 96 about modules; OTT *) | 
| 1284 | 57 | |
| 1272 
6d140b2c751f
added ott-example about Leroy96
 Christian Urban <urbanc@in.tum.de> parents: 
1265diff
changeset | 58 | nominal_datatype mexp = | 
| 1296 | 59 | Acc "path" | 
| 60 | | Stru "body" | |
| 61 | | Funct x::"name" "sexp" m::"mexp" bind x in m | |
| 62 | | FApp "mexp" "path" | |
| 63 | | Ascr "mexp" "sexp" | |
| 1272 
6d140b2c751f
added ott-example about Leroy96
 Christian Urban <urbanc@in.tum.de> parents: 
1265diff
changeset | 64 | and body = | 
| 
6d140b2c751f
added ott-example about Leroy96
 Christian Urban <urbanc@in.tum.de> parents: 
1265diff
changeset | 65 | Empty | 
| 1296 | 66 | | Seq c::defn d::"body" bind "cbinders c" in d | 
| 1398 
1f3c01345c9e
Reordered examples in Test.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1396diff
changeset | 67 | and defn = | 
| 1296 | 68 | Type "name" "tyty" | 
| 69 | | Dty "name" | |
| 70 | | DStru "name" "mexp" | |
| 71 | | Val "name" "trmtrm" | |
| 1272 
6d140b2c751f
added ott-example about Leroy96
 Christian Urban <urbanc@in.tum.de> parents: 
1265diff
changeset | 72 | and sexp = | 
| 
6d140b2c751f
added ott-example about Leroy96
 Christian Urban <urbanc@in.tum.de> parents: 
1265diff
changeset | 73 | Sig sbody | 
| 1296 | 74 | | SFunc "name" "sexp" "sexp" | 
| 1398 
1f3c01345c9e
Reordered examples in Test.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1396diff
changeset | 75 | and sbody = | 
| 1272 
6d140b2c751f
added ott-example about Leroy96
 Christian Urban <urbanc@in.tum.de> parents: 
1265diff
changeset | 76 | SEmpty | 
| 1296 | 77 | | SSeq C::spec D::sbody bind "Cbinders C" in D | 
| 1272 
6d140b2c751f
added ott-example about Leroy96
 Christian Urban <urbanc@in.tum.de> parents: 
1265diff
changeset | 78 | and spec = | 
| 1398 
1f3c01345c9e
Reordered examples in Test.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1396diff
changeset | 79 | Type1 "name" | 
| 1296 | 80 | | Type2 "name" "tyty" | 
| 81 | | SStru "name" "sexp" | |
| 82 | | SVal "name" "tyty" | |
| 1272 
6d140b2c751f
added ott-example about Leroy96
 Christian Urban <urbanc@in.tum.de> parents: 
1265diff
changeset | 83 | and tyty = | 
| 1296 | 84 | Tyref1 "name" | 
| 85 | | Tyref2 "path" "tyty" | |
| 86 | | Fun "tyty" "tyty" | |
| 1272 
6d140b2c751f
added ott-example about Leroy96
 Christian Urban <urbanc@in.tum.de> parents: 
1265diff
changeset | 87 | and path = | 
| 1296 | 88 | Sref1 "name" | 
| 89 | | Sref2 "path" "name" | |
| 1272 
6d140b2c751f
added ott-example about Leroy96
 Christian Urban <urbanc@in.tum.de> parents: 
1265diff
changeset | 90 | and trmtrm = | 
| 1296 | 91 | Tref1 "name" | 
| 92 | | Tref2 "path" "name" | |
| 1340 
f201eb6acafc
Lift BV,FV,Permutations and injection :).
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1338diff
changeset | 93 | | Lam' v::"name" "tyty" M::"trmtrm" bind v in M | 
| 
f201eb6acafc
Lift BV,FV,Permutations and injection :).
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1338diff
changeset | 94 | | App' "trmtrm" "trmtrm" | 
| 
f201eb6acafc
Lift BV,FV,Permutations and injection :).
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1338diff
changeset | 95 | | Let' "body" "trmtrm" | 
| 1272 
6d140b2c751f
added ott-example about Leroy96
 Christian Urban <urbanc@in.tum.de> parents: 
1265diff
changeset | 96 | binder | 
| 
6d140b2c751f
added ott-example about Leroy96
 Christian Urban <urbanc@in.tum.de> parents: 
1265diff
changeset | 97 | cbinders :: "defn \<Rightarrow> atom set" | 
| 
6d140b2c751f
added ott-example about Leroy96
 Christian Urban <urbanc@in.tum.de> parents: 
1265diff
changeset | 98 | and Cbinders :: "spec \<Rightarrow> atom set" | 
| 
6d140b2c751f
added ott-example about Leroy96
 Christian Urban <urbanc@in.tum.de> parents: 
1265diff
changeset | 99 | where | 
| 
6d140b2c751f
added ott-example about Leroy96
 Christian Urban <urbanc@in.tum.de> parents: 
1265diff
changeset | 100 |   "cbinders (Type t T) = {atom t}"
 | 
| 
6d140b2c751f
added ott-example about Leroy96
 Christian Urban <urbanc@in.tum.de> parents: 
1265diff
changeset | 101 | | "cbinders (Dty t) = {atom t}"
 | 
| 
6d140b2c751f
added ott-example about Leroy96
 Christian Urban <urbanc@in.tum.de> parents: 
1265diff
changeset | 102 | | "cbinders (DStru x s) = {atom x}"
 | 
| 1396 
08294f4d6c08
Linked parser to fv and alpha.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1392diff
changeset | 103 | | "cbinders (Val v M) = {atom v}"
 | 
| 1272 
6d140b2c751f
added ott-example about Leroy96
 Christian Urban <urbanc@in.tum.de> parents: 
1265diff
changeset | 104 | | "Cbinders (Type1 t) = {atom t}"
 | 
| 
6d140b2c751f
added ott-example about Leroy96
 Christian Urban <urbanc@in.tum.de> parents: 
1265diff
changeset | 105 | | "Cbinders (Type2 t T) = {atom t}"
 | 
| 
6d140b2c751f
added ott-example about Leroy96
 Christian Urban <urbanc@in.tum.de> parents: 
1265diff
changeset | 106 | | "Cbinders (SStru x S) = {atom x}"
 | 
| 1396 
08294f4d6c08
Linked parser to fv and alpha.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1392diff
changeset | 107 | | "Cbinders (SVal v T) = {atom v}"
 | 
| 1265 
fc8f5897b00a
first attempt to make sense out of the core-haskell definition
 Christian Urban <urbanc@in.tum.de> parents: 
1261diff
changeset | 108 | |
| 1515 
76fa21f27f22
Rename "_property" to ".property"
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1510diff
changeset | 109 | thm mexp_body_defn_sexp_sbody_spec_tyty_path_trmtrm.fv | 
| 
76fa21f27f22
Rename "_property" to ".property"
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1510diff
changeset | 110 | thm mexp_body_defn_sexp_sbody_spec_tyty_path_trmtrm.eq_iff | 
| 
76fa21f27f22
Rename "_property" to ".property"
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1510diff
changeset | 111 | thm mexp_body_defn_sexp_sbody_spec_tyty_path_trmtrm.bn | 
| 
76fa21f27f22
Rename "_property" to ".property"
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1510diff
changeset | 112 | thm mexp_body_defn_sexp_sbody_spec_tyty_path_trmtrm.perm | 
| 
76fa21f27f22
Rename "_property" to ".property"
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1510diff
changeset | 113 | thm mexp_body_defn_sexp_sbody_spec_tyty_path_trmtrm.induct | 
| 
76fa21f27f22
Rename "_property" to ".property"
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1510diff
changeset | 114 | thm mexp_body_defn_sexp_sbody_spec_tyty_path_trmtrm.inducts | 
| 
76fa21f27f22
Rename "_property" to ".property"
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1510diff
changeset | 115 | thm mexp_body_defn_sexp_sbody_spec_tyty_path_trmtrm.distinct | 
| 1553 
4355eb3b7161
Automatically derive support for datatypes with at-most one binding per constructor.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1549diff
changeset | 116 | thm mexp_body_defn_sexp_sbody_spec_tyty_path_trmtrm.supp | 
| 
4355eb3b7161
Automatically derive support for datatypes with at-most one binding per constructor.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1549diff
changeset | 117 | thm mexp_body_defn_sexp_sbody_spec_tyty_path_trmtrm.fv[simplified mexp_body_defn_sexp_sbody_spec_tyty_path_trmtrm.supp] | 
| 1392 
baa67b07f436
parser produces ordered bn-fun information
 Christian Urban <urbanc@in.tum.de> parents: 
1380diff
changeset | 118 | |
| 1312 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 119 | (* example 3 from Peter Sewell's bestiary *) | 
| 1418 
632b08744613
With the 4 cheats, all examples fully lift.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1416diff
changeset | 120 | |
| 1312 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 121 | nominal_datatype exp = | 
| 1340 
f201eb6acafc
Lift BV,FV,Permutations and injection :).
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1338diff
changeset | 122 | VarP "name" | 
| 
f201eb6acafc
Lift BV,FV,Permutations and injection :).
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1338diff
changeset | 123 | | AppP "exp" "exp" | 
| 
f201eb6acafc
Lift BV,FV,Permutations and injection :).
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1338diff
changeset | 124 | | LamP x::"name" e::"exp" bind x in e | 
| 1466 
d18defacda25
commented out examples that should not work; but for example type-scheme example should work
 Christian Urban <urbanc@in.tum.de> parents: 
1465diff
changeset | 125 | | LetP x::"name" p::"pat3" e1::"exp" e2::"exp" bind x in e2, bind "bp'' p" in e1 | 
| 1378 
5c6c950812b1
All examples should work.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1368diff
changeset | 126 | and pat3 = | 
| 1319 
d793ce9cd06f
potential problem with the phd-example, where two permutations are generated, but only one is used
 Christian Urban <urbanc@in.tum.de> parents: 
1313diff
changeset | 127 | PVar "name" | 
| 1312 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 128 | | PUnit | 
| 1378 
5c6c950812b1
All examples should work.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1368diff
changeset | 129 | | PPair "pat3" "pat3" | 
| 1312 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 130 | binder | 
| 1466 
d18defacda25
commented out examples that should not work; but for example type-scheme example should work
 Christian Urban <urbanc@in.tum.de> parents: 
1465diff
changeset | 131 | bp'' :: "pat3 \<Rightarrow> atom set" | 
| 1312 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 132 | where | 
| 1466 
d18defacda25
commented out examples that should not work; but for example type-scheme example should work
 Christian Urban <urbanc@in.tum.de> parents: 
1465diff
changeset | 133 |   "bp'' (PVar x) = {atom x}"
 | 
| 
d18defacda25
commented out examples that should not work; but for example type-scheme example should work
 Christian Urban <urbanc@in.tum.de> parents: 
1465diff
changeset | 134 | | "bp'' (PUnit) = {}"
 | 
| 
d18defacda25
commented out examples that should not work; but for example type-scheme example should work
 Christian Urban <urbanc@in.tum.de> parents: 
1465diff
changeset | 135 | | "bp'' (PPair p1 p2) = bp'' p1 \<union> bp'' p2" | 
| 1418 
632b08744613
With the 4 cheats, all examples fully lift.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1416diff
changeset | 136 | |
| 1515 
76fa21f27f22
Rename "_property" to ".property"
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1510diff
changeset | 137 | thm exp_pat3.fv | 
| 
76fa21f27f22
Rename "_property" to ".property"
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1510diff
changeset | 138 | thm exp_pat3.eq_iff | 
| 
76fa21f27f22
Rename "_property" to ".property"
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1510diff
changeset | 139 | thm exp_pat3.bn | 
| 
76fa21f27f22
Rename "_property" to ".property"
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1510diff
changeset | 140 | thm exp_pat3.perm | 
| 
76fa21f27f22
Rename "_property" to ".property"
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1510diff
changeset | 141 | thm exp_pat3.induct | 
| 
76fa21f27f22
Rename "_property" to ".property"
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1510diff
changeset | 142 | thm exp_pat3.distinct | 
| 1553 
4355eb3b7161
Automatically derive support for datatypes with at-most one binding per constructor.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1549diff
changeset | 143 | thm exp_pat3.fv[simplified exp_pat3.supp] | 
| 1312 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 144 | |
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 145 | (* example 6 from Peter Sewell's bestiary *) | 
| 1341 
c25f797c7e6e
Added lifting of pseudo-injectivity, commented out the code again and enabled the weird examples.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1340diff
changeset | 146 | nominal_datatype exp6 = | 
| 1312 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 147 | EVar name | 
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 148 | | EPair exp6 exp6 | 
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 149 | | ELetRec x::name p::pat6 e1::exp6 e2::exp6 bind x in e1, bind x in e2, bind "bp6 p" in e1 | 
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 150 | and pat6 = | 
| 1340 
f201eb6acafc
Lift BV,FV,Permutations and injection :).
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1338diff
changeset | 151 | PVar' name | 
| 
f201eb6acafc
Lift BV,FV,Permutations and injection :).
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1338diff
changeset | 152 | | PUnit' | 
| 
f201eb6acafc
Lift BV,FV,Permutations and injection :).
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1338diff
changeset | 153 | | PPair' pat6 pat6 | 
| 1312 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 154 | binder | 
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 155 | bp6 :: "pat6 \<Rightarrow> atom set" | 
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 156 | where | 
| 1340 
f201eb6acafc
Lift BV,FV,Permutations and injection :).
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1338diff
changeset | 157 |   "bp6 (PVar' x) = {atom x}"
 | 
| 
f201eb6acafc
Lift BV,FV,Permutations and injection :).
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1338diff
changeset | 158 | | "bp6 (PUnit') = {}"
 | 
| 1341 
c25f797c7e6e
Added lifting of pseudo-injectivity, commented out the code again and enabled the weird examples.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1340diff
changeset | 159 | | "bp6 (PPair' p1 p2) = bp6 p1 \<union> bp6 p2" | 
| 1418 
632b08744613
With the 4 cheats, all examples fully lift.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1416diff
changeset | 160 | |
| 1515 
76fa21f27f22
Rename "_property" to ".property"
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1510diff
changeset | 161 | thm exp6_pat6.fv | 
| 
76fa21f27f22
Rename "_property" to ".property"
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1510diff
changeset | 162 | thm exp6_pat6.eq_iff | 
| 
76fa21f27f22
Rename "_property" to ".property"
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1510diff
changeset | 163 | thm exp6_pat6.bn | 
| 
76fa21f27f22
Rename "_property" to ".property"
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1510diff
changeset | 164 | thm exp6_pat6.perm | 
| 
76fa21f27f22
Rename "_property" to ".property"
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1510diff
changeset | 165 | thm exp6_pat6.induct | 
| 
76fa21f27f22
Rename "_property" to ".property"
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1510diff
changeset | 166 | thm exp6_pat6.distinct | 
| 1312 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 167 | |
| 1553 
4355eb3b7161
Automatically derive support for datatypes with at-most one binding per constructor.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1549diff
changeset | 168 | |
| 1398 
1f3c01345c9e
Reordered examples in Test.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1396diff
changeset | 169 | (* THE REST ARE NOT SUPPOSED TO WORK YET *) | 
| 
1f3c01345c9e
Reordered examples in Test.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1396diff
changeset | 170 | |
| 1312 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 171 | (* example 7 from Peter Sewell's bestiary *) | 
| 1466 
d18defacda25
commented out examples that should not work; but for example type-scheme example should work
 Christian Urban <urbanc@in.tum.de> parents: 
1465diff
changeset | 172 | (* dest_Const raised | 
| 1341 
c25f797c7e6e
Added lifting of pseudo-injectivity, commented out the code again and enabled the weird examples.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1340diff
changeset | 173 | nominal_datatype exp7 = | 
| 1466 
d18defacda25
commented out examples that should not work; but for example type-scheme example should work
 Christian Urban <urbanc@in.tum.de> parents: 
1465diff
changeset | 174 | EVar' name | 
| 
d18defacda25
commented out examples that should not work; but for example type-scheme example should work
 Christian Urban <urbanc@in.tum.de> parents: 
1465diff
changeset | 175 | | EUnit' | 
| 
d18defacda25
commented out examples that should not work; but for example type-scheme example should work
 Christian Urban <urbanc@in.tum.de> parents: 
1465diff
changeset | 176 | | EPair' exp7 exp7 | 
| 
d18defacda25
commented out examples that should not work; but for example type-scheme example should work
 Christian Urban <urbanc@in.tum.de> parents: 
1465diff
changeset | 177 | | ELetRec' l::lrbs e::exp7 bind "b7s l" in e, bind "b7s l" in l | 
| 1312 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 178 | and lrb = | 
| 1466 
d18defacda25
commented out examples that should not work; but for example type-scheme example should work
 Christian Urban <urbanc@in.tum.de> parents: 
1465diff
changeset | 179 | Assign' name exp7 | 
| 1312 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 180 | and lrbs = | 
| 1466 
d18defacda25
commented out examples that should not work; but for example type-scheme example should work
 Christian Urban <urbanc@in.tum.de> parents: 
1465diff
changeset | 181 | Single' lrb | 
| 
d18defacda25
commented out examples that should not work; but for example type-scheme example should work
 Christian Urban <urbanc@in.tum.de> parents: 
1465diff
changeset | 182 | | More' lrb lrbs | 
| 1312 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 183 | binder | 
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 184 | b7 :: "lrb \<Rightarrow> atom set" and | 
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 185 | b7s :: "lrbs \<Rightarrow> atom set" | 
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 186 | where | 
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 187 |   "b7 (Assign x e) = {atom x}"
 | 
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 188 | | "b7s (Single a) = b7 a" | 
| 1341 
c25f797c7e6e
Added lifting of pseudo-injectivity, commented out the code again and enabled the weird examples.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1340diff
changeset | 189 | | "b7s (More a as) = (b7 a) \<union> (b7s as)" | 
| 1361 
1e811e3424f3
With restricted_nominal=1, exp7 and exp8 work. Not sure about proving bn_rsp there.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1355diff
changeset | 190 | thm alpha_exp7_raw_alpha_lrb_raw_alpha_lrbs_raw.intros | 
| 1466 
d18defacda25
commented out examples that should not work; but for example type-scheme example should work
 Christian Urban <urbanc@in.tum.de> parents: 
1465diff
changeset | 191 | *) | 
| 1312 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 192 | |
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 193 | (* example 8 from Peter Sewell's bestiary *) | 
| 1466 
d18defacda25
commented out examples that should not work; but for example type-scheme example should work
 Christian Urban <urbanc@in.tum.de> parents: 
1465diff
changeset | 194 | (* | 
| 
d18defacda25
commented out examples that should not work; but for example type-scheme example should work
 Christian Urban <urbanc@in.tum.de> parents: 
1465diff
changeset | 195 | *** fv_bn: recursive argument, but wrong datatype. | 
| 
d18defacda25
commented out examples that should not work; but for example type-scheme example should work
 Christian Urban <urbanc@in.tum.de> parents: 
1465diff
changeset | 196 | *** At command "nominal_datatype". | 
| 1341 
c25f797c7e6e
Added lifting of pseudo-injectivity, commented out the code again and enabled the weird examples.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1340diff
changeset | 197 | nominal_datatype exp8 = | 
| 1466 
d18defacda25
commented out examples that should not work; but for example type-scheme example should work
 Christian Urban <urbanc@in.tum.de> parents: 
1465diff
changeset | 198 | EVar' name | 
| 
d18defacda25
commented out examples that should not work; but for example type-scheme example should work
 Christian Urban <urbanc@in.tum.de> parents: 
1465diff
changeset | 199 | | EUnit' | 
| 
d18defacda25
commented out examples that should not work; but for example type-scheme example should work
 Christian Urban <urbanc@in.tum.de> parents: 
1465diff
changeset | 200 | | EPair' exp8 exp8 | 
| 
d18defacda25
commented out examples that should not work; but for example type-scheme example should work
 Christian Urban <urbanc@in.tum.de> parents: 
1465diff
changeset | 201 | | ELetRec' l::lrbs8 e::exp8 bind "b_lrbs8 l" in e, bind "b_lrbs8 l" in l | 
| 1312 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 202 | and fnclause = | 
| 1466 
d18defacda25
commented out examples that should not work; but for example type-scheme example should work
 Christian Urban <urbanc@in.tum.de> parents: 
1465diff
changeset | 203 | K' x::name p::pat8 e::exp8 bind "b_pat p" in e | 
| 1312 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 204 | and fnclauses = | 
| 1466 
d18defacda25
commented out examples that should not work; but for example type-scheme example should work
 Christian Urban <urbanc@in.tum.de> parents: 
1465diff
changeset | 205 | S' fnclause | 
| 
d18defacda25
commented out examples that should not work; but for example type-scheme example should work
 Christian Urban <urbanc@in.tum.de> parents: 
1465diff
changeset | 206 | | ORs' fnclause fnclauses | 
| 1312 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 207 | and lrb8 = | 
| 1466 
d18defacda25
commented out examples that should not work; but for example type-scheme example should work
 Christian Urban <urbanc@in.tum.de> parents: 
1465diff
changeset | 208 | Clause' fnclauses | 
| 1312 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 209 | and lrbs8 = | 
| 1466 
d18defacda25
commented out examples that should not work; but for example type-scheme example should work
 Christian Urban <urbanc@in.tum.de> parents: 
1465diff
changeset | 210 | Single' lrb8 | 
| 
d18defacda25
commented out examples that should not work; but for example type-scheme example should work
 Christian Urban <urbanc@in.tum.de> parents: 
1465diff
changeset | 211 | | More' lrb8 lrbs8 | 
| 1312 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 212 | and pat8 = | 
| 1466 
d18defacda25
commented out examples that should not work; but for example type-scheme example should work
 Christian Urban <urbanc@in.tum.de> parents: 
1465diff
changeset | 213 | PVar'' name | 
| 
d18defacda25
commented out examples that should not work; but for example type-scheme example should work
 Christian Urban <urbanc@in.tum.de> parents: 
1465diff
changeset | 214 | | PUnit'' | 
| 
d18defacda25
commented out examples that should not work; but for example type-scheme example should work
 Christian Urban <urbanc@in.tum.de> parents: 
1465diff
changeset | 215 | | PPair'' pat8 pat8 | 
| 1312 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 216 | binder | 
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 217 | b_lrbs8 :: "lrbs8 \<Rightarrow> atom set" and | 
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 218 | b_pat :: "pat8 \<Rightarrow> atom set" and | 
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 219 | b_fnclauses :: "fnclauses \<Rightarrow> atom set" and | 
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 220 | b_fnclause :: "fnclause \<Rightarrow> atom set" and | 
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 221 | b_lrb8 :: "lrb8 \<Rightarrow> atom set" | 
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 222 | where | 
| 1466 
d18defacda25
commented out examples that should not work; but for example type-scheme example should work
 Christian Urban <urbanc@in.tum.de> parents: 
1465diff
changeset | 223 | "b_lrbs8 (Single' l) = b_lrb8 l" | 
| 
d18defacda25
commented out examples that should not work; but for example type-scheme example should work
 Christian Urban <urbanc@in.tum.de> parents: 
1465diff
changeset | 224 | | "b_lrbs8 (More' l ls) = b_lrb8 l \<union> b_lrbs8 ls" | 
| 
d18defacda25
commented out examples that should not work; but for example type-scheme example should work
 Christian Urban <urbanc@in.tum.de> parents: 
1465diff
changeset | 225 | | "b_pat (PVar'' x) = {atom x}"
 | 
| 
d18defacda25
commented out examples that should not work; but for example type-scheme example should work
 Christian Urban <urbanc@in.tum.de> parents: 
1465diff
changeset | 226 | | "b_pat (PUnit'') = {}"
 | 
| 
d18defacda25
commented out examples that should not work; but for example type-scheme example should work
 Christian Urban <urbanc@in.tum.de> parents: 
1465diff
changeset | 227 | | "b_pat (PPair'' p1 p2) = b_pat p1 \<union> b_pat p2" | 
| 
d18defacda25
commented out examples that should not work; but for example type-scheme example should work
 Christian Urban <urbanc@in.tum.de> parents: 
1465diff
changeset | 228 | | "b_fnclauses (S' fc) = (b_fnclause fc)" | 
| 
d18defacda25
commented out examples that should not work; but for example type-scheme example should work
 Christian Urban <urbanc@in.tum.de> parents: 
1465diff
changeset | 229 | | "b_fnclauses (ORs' fc fcs) = (b_fnclause fc) \<union> (b_fnclauses fcs)" | 
| 
d18defacda25
commented out examples that should not work; but for example type-scheme example should work
 Christian Urban <urbanc@in.tum.de> parents: 
1465diff
changeset | 230 | | "b_lrb8 (Clause' fcs) = (b_fnclauses fcs)" | 
| 
d18defacda25
commented out examples that should not work; but for example type-scheme example should work
 Christian Urban <urbanc@in.tum.de> parents: 
1465diff
changeset | 231 | | "b_fnclause (K' x pat exp8) = {atom x}"
 | 
| 1361 
1e811e3424f3
With restricted_nominal=1, exp7 and exp8 work. Not sure about proving bn_rsp there.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1355diff
changeset | 232 | thm alpha_exp8_raw_alpha_fnclause_raw_alpha_fnclauses_raw_alpha_lrb8_raw_alpha_lrbs8_raw_alpha_pat8_raw.intros | 
| 1466 
d18defacda25
commented out examples that should not work; but for example type-scheme example should work
 Christian Urban <urbanc@in.tum.de> parents: 
1465diff
changeset | 233 | *) | 
| 1398 
1f3c01345c9e
Reordered examples in Test.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1396diff
changeset | 234 | (* example 4 from Terms.thy *) | 
| 
1f3c01345c9e
Reordered examples in Test.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1396diff
changeset | 235 | (* fv_eqvt does not work, we need to repaire defined permute functions | 
| 
1f3c01345c9e
Reordered examples in Test.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1396diff
changeset | 236 | defined fv and defined alpha... *) | 
| 1466 
d18defacda25
commented out examples that should not work; but for example type-scheme example should work
 Christian Urban <urbanc@in.tum.de> parents: 
1465diff
changeset | 237 | (* lists-datastructure does not work yet | 
| 1398 
1f3c01345c9e
Reordered examples in Test.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1396diff
changeset | 238 | nominal_datatype trm4 = | 
| 
1f3c01345c9e
Reordered examples in Test.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1396diff
changeset | 239 | Vr4 "name" | 
| 
1f3c01345c9e
Reordered examples in Test.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1396diff
changeset | 240 | | Ap4 "trm4" "trm4 list" | 
| 
1f3c01345c9e
Reordered examples in Test.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1396diff
changeset | 241 | | Lm4 x::"name" t::"trm4" bind x in t | 
| 
1f3c01345c9e
Reordered examples in Test.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1396diff
changeset | 242 | |
| 
1f3c01345c9e
Reordered examples in Test.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1396diff
changeset | 243 | thm alpha_trm4_raw_alpha_trm4_raw_list.intros[no_vars] | 
| 
1f3c01345c9e
Reordered examples in Test.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1396diff
changeset | 244 | thm fv_trm4_raw_fv_trm4_raw_list.simps[no_vars] | 
| 1466 
d18defacda25
commented out examples that should not work; but for example type-scheme example should work
 Christian Urban <urbanc@in.tum.de> parents: 
1465diff
changeset | 245 | *) | 
| 1398 
1f3c01345c9e
Reordered examples in Test.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1396diff
changeset | 246 | (* core haskell *) | 
| 
1f3c01345c9e
Reordered examples in Test.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1396diff
changeset | 247 | atom_decl var | 
| 
1f3c01345c9e
Reordered examples in Test.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1396diff
changeset | 248 | atom_decl tvar | 
| 
1f3c01345c9e
Reordered examples in Test.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1396diff
changeset | 249 | |
| 
1f3c01345c9e
Reordered examples in Test.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1396diff
changeset | 250 | (* there are types, coercion types and regular types *) | 
| 1466 
d18defacda25
commented out examples that should not work; but for example type-scheme example should work
 Christian Urban <urbanc@in.tum.de> parents: 
1465diff
changeset | 251 | (* list-data-structure | 
| 1398 
1f3c01345c9e
Reordered examples in Test.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1396diff
changeset | 252 | nominal_datatype tkind = | 
| 
1f3c01345c9e
Reordered examples in Test.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1396diff
changeset | 253 | KStar | 
| 
1f3c01345c9e
Reordered examples in Test.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1396diff
changeset | 254 | | KFun "tkind" "tkind" | 
| 
1f3c01345c9e
Reordered examples in Test.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1396diff
changeset | 255 | and ckind = | 
| 
1f3c01345c9e
Reordered examples in Test.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1396diff
changeset | 256 | CKEq "ty" "ty" | 
| 
1f3c01345c9e
Reordered examples in Test.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1396diff
changeset | 257 | and ty = | 
| 
1f3c01345c9e
Reordered examples in Test.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1396diff
changeset | 258 | TVar "tvar" | 
| 
1f3c01345c9e
Reordered examples in Test.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1396diff
changeset | 259 | | TC "string" | 
| 
1f3c01345c9e
Reordered examples in Test.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1396diff
changeset | 260 | | TApp "ty" "ty" | 
| 
1f3c01345c9e
Reordered examples in Test.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1396diff
changeset | 261 | | TFun "string" "ty list" | 
| 
1f3c01345c9e
Reordered examples in Test.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1396diff
changeset | 262 | | TAll tv::"tvar" "tkind" T::"ty" bind tv in T | 
| 
1f3c01345c9e
Reordered examples in Test.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1396diff
changeset | 263 | | TEq "ty" "ty" "ty" | 
| 
1f3c01345c9e
Reordered examples in Test.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1396diff
changeset | 264 | and co = | 
| 
1f3c01345c9e
Reordered examples in Test.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1396diff
changeset | 265 | CC "string" | 
| 
1f3c01345c9e
Reordered examples in Test.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1396diff
changeset | 266 | | CApp "co" "co" | 
| 
1f3c01345c9e
Reordered examples in Test.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1396diff
changeset | 267 | | CFun "string" "co list" | 
| 
1f3c01345c9e
Reordered examples in Test.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1396diff
changeset | 268 | | CAll tv::"tvar" "ckind" C::"co" bind tv in C | 
| 
1f3c01345c9e
Reordered examples in Test.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1396diff
changeset | 269 | | CEq "co" "co" "co" | 
| 
1f3c01345c9e
Reordered examples in Test.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1396diff
changeset | 270 | | CSym "co" | 
| 
1f3c01345c9e
Reordered examples in Test.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1396diff
changeset | 271 | | CCir "co" "co" | 
| 
1f3c01345c9e
Reordered examples in Test.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1396diff
changeset | 272 | | CLeft "co" | 
| 
1f3c01345c9e
Reordered examples in Test.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1396diff
changeset | 273 | | CRight "co" | 
| 
1f3c01345c9e
Reordered examples in Test.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1396diff
changeset | 274 | | CSim "co" | 
| 
1f3c01345c9e
Reordered examples in Test.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1396diff
changeset | 275 | | CRightc "co" | 
| 
1f3c01345c9e
Reordered examples in Test.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1396diff
changeset | 276 | | CLeftc "co" | 
| 
1f3c01345c9e
Reordered examples in Test.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1396diff
changeset | 277 | | CCoe "co" "co" | 
| 
1f3c01345c9e
Reordered examples in Test.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1396diff
changeset | 278 | |
| 
1f3c01345c9e
Reordered examples in Test.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1396diff
changeset | 279 | abbreviation | 
| 
1f3c01345c9e
Reordered examples in Test.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1396diff
changeset | 280 | "atoms A \<equiv> atom ` A" | 
| 
1f3c01345c9e
Reordered examples in Test.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1396diff
changeset | 281 | |
| 
1f3c01345c9e
Reordered examples in Test.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1396diff
changeset | 282 | nominal_datatype trm = | 
| 
1f3c01345c9e
Reordered examples in Test.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1396diff
changeset | 283 | Var "var" | 
| 
1f3c01345c9e
Reordered examples in Test.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1396diff
changeset | 284 | | C "string" | 
| 
1f3c01345c9e
Reordered examples in Test.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1396diff
changeset | 285 | | LAM tv::"tvar" "kind" t::"trm" bind tv in t | 
| 
1f3c01345c9e
Reordered examples in Test.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1396diff
changeset | 286 | | APP "trm" "ty" | 
| 
1f3c01345c9e
Reordered examples in Test.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1396diff
changeset | 287 | | Lam v::"var" "ty" t::"trm" bind v in t | 
| 
1f3c01345c9e
Reordered examples in Test.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1396diff
changeset | 288 | | App "trm" "trm" | 
| 
1f3c01345c9e
Reordered examples in Test.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1396diff
changeset | 289 | | Let x::"var" "ty" "trm" t::"trm" bind x in t | 
| 
1f3c01345c9e
Reordered examples in Test.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1396diff
changeset | 290 | | Case "trm" "assoc list" | 
| 
1f3c01345c9e
Reordered examples in Test.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1396diff
changeset | 291 | | Cast "trm" "ty" --"ty is supposed to be a coercion type only" | 
| 
1f3c01345c9e
Reordered examples in Test.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1396diff
changeset | 292 | and assoc = | 
| 
1f3c01345c9e
Reordered examples in Test.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1396diff
changeset | 293 | A p::"pat" t::"trm" bind "bv p" in t | 
| 
1f3c01345c9e
Reordered examples in Test.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1396diff
changeset | 294 | and pat = | 
| 
1f3c01345c9e
Reordered examples in Test.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1396diff
changeset | 295 | K "string" "(tvar \<times> kind) list" "(var \<times> ty) list" | 
| 
1f3c01345c9e
Reordered examples in Test.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1396diff
changeset | 296 | binder | 
| 
1f3c01345c9e
Reordered examples in Test.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1396diff
changeset | 297 | bv :: "pat \<Rightarrow> atom set" | 
| 
1f3c01345c9e
Reordered examples in Test.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1396diff
changeset | 298 | where | 
| 
1f3c01345c9e
Reordered examples in Test.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1396diff
changeset | 299 | "bv (K s ts vs) = (atoms (set (map fst ts))) \<union> (atoms (set (map fst vs)))" | 
| 
1f3c01345c9e
Reordered examples in Test.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1396diff
changeset | 300 | *) | 
| 
1f3c01345c9e
Reordered examples in Test.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1396diff
changeset | 301 | |
| 
1f3c01345c9e
Reordered examples in Test.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1396diff
changeset | 302 | (* example 8 from Terms.thy *) | 
| 
1f3c01345c9e
Reordered examples in Test.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1396diff
changeset | 303 | |
| 
1f3c01345c9e
Reordered examples in Test.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1396diff
changeset | 304 | (* Binding in a term under a bn, needs to fail *) | 
| 1466 
d18defacda25
commented out examples that should not work; but for example type-scheme example should work
 Christian Urban <urbanc@in.tum.de> parents: 
1465diff
changeset | 305 | (* | 
| 1398 
1f3c01345c9e
Reordered examples in Test.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1396diff
changeset | 306 | nominal_datatype foo8 = | 
| 
1f3c01345c9e
Reordered examples in Test.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1396diff
changeset | 307 | Foo0 "name" | 
| 
1f3c01345c9e
Reordered examples in Test.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1396diff
changeset | 308 | | Foo1 b::"bar8" f::"foo8" bind "bv8 b" in f --"check fo error if this is called foo" | 
| 
1f3c01345c9e
Reordered examples in Test.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1396diff
changeset | 309 | and bar8 = | 
| 
1f3c01345c9e
Reordered examples in Test.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1396diff
changeset | 310 | Bar0 "name" | 
| 
1f3c01345c9e
Reordered examples in Test.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1396diff
changeset | 311 | | Bar1 "name" s::"name" b::"bar8" bind s in b | 
| 
1f3c01345c9e
Reordered examples in Test.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1396diff
changeset | 312 | binder | 
| 
1f3c01345c9e
Reordered examples in Test.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1396diff
changeset | 313 | bv8 | 
| 
1f3c01345c9e
Reordered examples in Test.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1396diff
changeset | 314 | where | 
| 
1f3c01345c9e
Reordered examples in Test.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1396diff
changeset | 315 |   "bv8 (Bar0 x) = {}"
 | 
| 
1f3c01345c9e
Reordered examples in Test.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1396diff
changeset | 316 | | "bv8 (Bar1 v x b) = {atom v}"
 | 
| 1466 
d18defacda25
commented out examples that should not work; but for example type-scheme example should work
 Christian Urban <urbanc@in.tum.de> parents: 
1465diff
changeset | 317 | *) | 
| 1398 
1f3c01345c9e
Reordered examples in Test.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1396diff
changeset | 318 | |
| 1312 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 319 | (* example 9 from Peter Sewell's bestiary *) | 
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 320 | (* run out of steam at the moment *) | 
| 1265 
fc8f5897b00a
first attempt to make sense out of the core-haskell definition
 Christian Urban <urbanc@in.tum.de> parents: 
1261diff
changeset | 321 | |
| 961 
0f88e04eb486
Variable takes a 'name'.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
954diff
changeset | 322 | end | 
| 1223 
160343d86a6f
"raw"-ified the term-constructors and types given in the specification
 Christian Urban <urbanc@in.tum.de> parents: 
1194diff
changeset | 323 | |
| 
160343d86a6f
"raw"-ified the term-constructors and types given in the specification
 Christian Urban <urbanc@in.tum.de> parents: 
1194diff
changeset | 324 | |
| 1228 
c179ad9d2446
declarartion of the raw datatype already works; raw binding functions throw an exception about mutual recursive types
 Christian Urban <urbanc@in.tum.de> parents: 
1223diff
changeset | 325 |