| author | Christian Urban <urbanc@in.tum.de> | 
| Tue, 02 Mar 2010 20:11:56 +0100 | |
| changeset 1321 | bfd9af005e23 | 
| parent 1320 | 447666754176 | 
| child 1322 | 12ce01673188 | 
| 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 | 
| 1261 
853abc14c5c6
added IsaMakefile...but so far included only a test for the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1258diff
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 | |
| 1320 
447666754176
Another problem with permutations in alpha and possibly also in fv
 Christian Urban <urbanc@in.tum.de> parents: 
1319diff
changeset | 5 | text {* weirdo example from Peter Sewell's bestiary *}
 | 
| 
447666754176
Another problem with permutations in alpha and possibly also in fv
 Christian Urban <urbanc@in.tum.de> parents: 
1319diff
changeset | 6 | |
| 
447666754176
Another problem with permutations in alpha and possibly also in fv
 Christian Urban <urbanc@in.tum.de> parents: 
1319diff
changeset | 7 | nominal_datatype weird = | 
| 
447666754176
Another problem with permutations in alpha and possibly also in fv
 Christian Urban <urbanc@in.tum.de> parents: 
1319diff
changeset | 8 | WBind x::"name" y::"name" p1::"weird" p2::"weird" p3::"weird" | 
| 
447666754176
Another problem with permutations in alpha and possibly also in fv
 Christian Urban <urbanc@in.tum.de> parents: 
1319diff
changeset | 9 | bind x in p1, bind x in p2, bind y in p2, bind y in p3 | 
| 1321 
bfd9af005e23
preliinary test about alpha-weirdo
 Christian Urban <urbanc@in.tum.de> parents: 
1320diff
changeset | 10 | | WV "name" | 
| 
bfd9af005e23
preliinary test about alpha-weirdo
 Christian Urban <urbanc@in.tum.de> parents: 
1320diff
changeset | 11 | | WP "weird" "weird" | 
| 1320 
447666754176
Another problem with permutations in alpha and possibly also in fv
 Christian Urban <urbanc@in.tum.de> parents: 
1319diff
changeset | 12 | |
| 1321 
bfd9af005e23
preliinary test about alpha-weirdo
 Christian Urban <urbanc@in.tum.de> parents: 
1320diff
changeset | 13 | thm permute_weird_raw.simps[no_vars] | 
| 1320 
447666754176
Another problem with permutations in alpha and possibly also in fv
 Christian Urban <urbanc@in.tum.de> parents: 
1319diff
changeset | 14 | thm alpha_weird_raw.intros[no_vars] | 
| 
447666754176
Another problem with permutations in alpha and possibly also in fv
 Christian Urban <urbanc@in.tum.de> parents: 
1319diff
changeset | 15 | thm fv_weird_raw.simps[no_vars] | 
| 
447666754176
Another problem with permutations in alpha and possibly also in fv
 Christian Urban <urbanc@in.tum.de> parents: 
1319diff
changeset | 16 | (* Potential problem: Is it correct that in the fv-function | 
| 
447666754176
Another problem with permutations in alpha and possibly also in fv
 Christian Urban <urbanc@in.tum.de> parents: 
1319diff
changeset | 17 | the first two terms are created? Should be ommitted. Also | 
| 
447666754176
Another problem with permutations in alpha and possibly also in fv
 Christian Urban <urbanc@in.tum.de> parents: 
1319diff
changeset | 18 | something wrong with the permutations. *) | 
| 
447666754176
Another problem with permutations in alpha and possibly also in fv
 Christian Urban <urbanc@in.tum.de> parents: 
1319diff
changeset | 19 | |
| 1321 
bfd9af005e23
preliinary test about alpha-weirdo
 Christian Urban <urbanc@in.tum.de> parents: 
1320diff
changeset | 20 | primrec | 
| 
bfd9af005e23
preliinary test about alpha-weirdo
 Christian Urban <urbanc@in.tum.de> parents: 
1320diff
changeset | 21 | fv_weird | 
| 
bfd9af005e23
preliinary test about alpha-weirdo
 Christian Urban <urbanc@in.tum.de> parents: 
1320diff
changeset | 22 | where | 
| 
bfd9af005e23
preliinary test about alpha-weirdo
 Christian Urban <urbanc@in.tum.de> parents: 
1320diff
changeset | 23 | "fv_weird (WBind_raw x y p1 p2 p3) = | 
| 
bfd9af005e23
preliinary test about alpha-weirdo
 Christian Urban <urbanc@in.tum.de> parents: 
1320diff
changeset | 24 |     (fv_weird p1 - {atom x}) \<union> (fv_weird p2 - ({atom x} \<union> {atom y})) \<union> (fv_weird p3 - {atom y})"
 | 
| 
bfd9af005e23
preliinary test about alpha-weirdo
 Christian Urban <urbanc@in.tum.de> parents: 
1320diff
changeset | 25 | | "fv_weird (WV_raw x) = {atom x}"
 | 
| 
bfd9af005e23
preliinary test about alpha-weirdo
 Christian Urban <urbanc@in.tum.de> parents: 
1320diff
changeset | 26 | | "fv_weird (WP_raw p1 p2) = (fv_weird p1 \<union> fv_weird p2)" | 
| 
bfd9af005e23
preliinary test about alpha-weirdo
 Christian Urban <urbanc@in.tum.de> parents: 
1320diff
changeset | 27 | |
| 
bfd9af005e23
preliinary test about alpha-weirdo
 Christian Urban <urbanc@in.tum.de> parents: 
1320diff
changeset | 28 | inductive | 
| 
bfd9af005e23
preliinary test about alpha-weirdo
 Christian Urban <urbanc@in.tum.de> parents: 
1320diff
changeset | 29 | alpha_weird | 
| 
bfd9af005e23
preliinary test about alpha-weirdo
 Christian Urban <urbanc@in.tum.de> parents: 
1320diff
changeset | 30 | where | 
| 
bfd9af005e23
preliinary test about alpha-weirdo
 Christian Urban <urbanc@in.tum.de> parents: 
1320diff
changeset | 31 | "\<exists>p1 p2. | 
| 
bfd9af005e23
preliinary test about alpha-weirdo
 Christian Urban <urbanc@in.tum.de> parents: 
1320diff
changeset | 32 |   ({atom y}, w3) \<approx>gen alpha_weird fv_weird p2 ({atom ya}, w3a) \<and>
 | 
| 
bfd9af005e23
preliinary test about alpha-weirdo
 Christian Urban <urbanc@in.tum.de> parents: 
1320diff
changeset | 33 |   ({atom x} \<union> {atom y}, w2) \<approx>gen alpha_weird fv_weird (p1 + p2) ({atom xa} \<union> {atom ya}, w2a) \<and>
 | 
| 
bfd9af005e23
preliinary test about alpha-weirdo
 Christian Urban <urbanc@in.tum.de> parents: 
1320diff
changeset | 34 |   ({atom x}, w1) \<approx>gen alpha_weird fv_weird p1 ({atom xa}, w1a) \<and>
 | 
| 
bfd9af005e23
preliinary test about alpha-weirdo
 Christian Urban <urbanc@in.tum.de> parents: 
1320diff
changeset | 35 |   supp p1 \<inter> supp p2 = {} \<Longrightarrow>
 | 
| 
bfd9af005e23
preliinary test about alpha-weirdo
 Christian Urban <urbanc@in.tum.de> parents: 
1320diff
changeset | 36 | alpha_weird (WBind_raw x y w1 w2 w3) (WBind_raw xa ya w1a w2a w3a)" | 
| 
bfd9af005e23
preliinary test about alpha-weirdo
 Christian Urban <urbanc@in.tum.de> parents: 
1320diff
changeset | 37 | | "x = xa \<Longrightarrow> alpha_weird (WV_raw x) (WV_raw xa)" | 
| 
bfd9af005e23
preliinary test about alpha-weirdo
 Christian Urban <urbanc@in.tum.de> parents: 
1320diff
changeset | 38 | | "alpha_weird w2 w2a \<and> alpha_weird w1 w1a \<Longrightarrow> alpha_weird (WP_raw w1 w2) (WP_raw w1a w2a)" | 
| 
bfd9af005e23
preliinary test about alpha-weirdo
 Christian Urban <urbanc@in.tum.de> parents: 
1320diff
changeset | 39 | |
| 
bfd9af005e23
preliinary test about alpha-weirdo
 Christian Urban <urbanc@in.tum.de> parents: 
1320diff
changeset | 40 | (* | 
| 
bfd9af005e23
preliinary test about alpha-weirdo
 Christian Urban <urbanc@in.tum.de> parents: 
1320diff
changeset | 41 | abbreviation "WBind \<equiv> WBind_raw" | 
| 
bfd9af005e23
preliinary test about alpha-weirdo
 Christian Urban <urbanc@in.tum.de> parents: 
1320diff
changeset | 42 | abbreviation "WP \<equiv> WP_raw" | 
| 
bfd9af005e23
preliinary test about alpha-weirdo
 Christian Urban <urbanc@in.tum.de> parents: 
1320diff
changeset | 43 | abbreviation "WV \<equiv> WV_raw" | 
| 
bfd9af005e23
preliinary test about alpha-weirdo
 Christian Urban <urbanc@in.tum.de> parents: 
1320diff
changeset | 44 | |
| 
bfd9af005e23
preliinary test about alpha-weirdo
 Christian Urban <urbanc@in.tum.de> parents: 
1320diff
changeset | 45 | lemma test: | 
| 
bfd9af005e23
preliinary test about alpha-weirdo
 Christian Urban <urbanc@in.tum.de> parents: 
1320diff
changeset | 46 | assumes a: "distinct [x, y, z]" | 
| 
bfd9af005e23
preliinary test about alpha-weirdo
 Christian Urban <urbanc@in.tum.de> parents: 
1320diff
changeset | 47 | shows "alpha_weird (WBind x y (WP (WV x) (WV z)) (WP (WV x) (WV y)) (WP (WV z) (WV y))) | 
| 
bfd9af005e23
preliinary test about alpha-weirdo
 Christian Urban <urbanc@in.tum.de> parents: 
1320diff
changeset | 48 | (WBind y x (WP (WV y) (WV z)) (WP (WV y) (WV x)) (WP (WV z) (WV x)))" | 
| 
bfd9af005e23
preliinary test about alpha-weirdo
 Christian Urban <urbanc@in.tum.de> parents: 
1320diff
changeset | 49 | apply(rule_tac alpha_weird.intros) | 
| 
bfd9af005e23
preliinary test about alpha-weirdo
 Christian Urban <urbanc@in.tum.de> parents: 
1320diff
changeset | 50 | unfolding alpha_gen | 
| 
bfd9af005e23
preliinary test about alpha-weirdo
 Christian Urban <urbanc@in.tum.de> parents: 
1320diff
changeset | 51 | using a | 
| 
bfd9af005e23
preliinary test about alpha-weirdo
 Christian Urban <urbanc@in.tum.de> parents: 
1320diff
changeset | 52 | apply(simp) | 
| 
bfd9af005e23
preliinary test about alpha-weirdo
 Christian Urban <urbanc@in.tum.de> parents: 
1320diff
changeset | 53 | apply(rule_tac x="(x \<leftrightarrow> y)" in exI) | 
| 
bfd9af005e23
preliinary test about alpha-weirdo
 Christian Urban <urbanc@in.tum.de> parents: 
1320diff
changeset | 54 | apply(rule_tac x="(x \<leftrightarrow> y)" in exI) | 
| 
bfd9af005e23
preliinary test about alpha-weirdo
 Christian Urban <urbanc@in.tum.de> parents: 
1320diff
changeset | 55 | apply(simp add: fresh_star_def) | 
| 
bfd9af005e23
preliinary test about alpha-weirdo
 Christian Urban <urbanc@in.tum.de> parents: 
1320diff
changeset | 56 | apply(simp add: flip_def) | 
| 
bfd9af005e23
preliinary test about alpha-weirdo
 Christian Urban <urbanc@in.tum.de> parents: 
1320diff
changeset | 57 | apply(auto) | 
| 
bfd9af005e23
preliinary test about alpha-weirdo
 Christian Urban <urbanc@in.tum.de> parents: 
1320diff
changeset | 58 | prefer 2 | 
| 
bfd9af005e23
preliinary test about alpha-weirdo
 Christian Urban <urbanc@in.tum.de> parents: 
1320diff
changeset | 59 | apply(rule alpha_weird.intros) | 
| 
bfd9af005e23
preliinary test about alpha-weirdo
 Christian Urban <urbanc@in.tum.de> parents: 
1320diff
changeset | 60 | apply(simp add: alpha_weird.intros(2)) | 
| 
bfd9af005e23
preliinary test about alpha-weirdo
 Christian Urban <urbanc@in.tum.de> parents: 
1320diff
changeset | 61 | prefer 2 | 
| 
bfd9af005e23
preliinary test about alpha-weirdo
 Christian Urban <urbanc@in.tum.de> parents: 
1320diff
changeset | 62 | apply(rule alpha_weird.intros) | 
| 
bfd9af005e23
preliinary test about alpha-weirdo
 Christian Urban <urbanc@in.tum.de> parents: 
1320diff
changeset | 63 | apply(simp add: alpha_weird.intros(2)) | 
| 
bfd9af005e23
preliinary test about alpha-weirdo
 Christian Urban <urbanc@in.tum.de> parents: 
1320diff
changeset | 64 | apply(simp) | 
| 
bfd9af005e23
preliinary test about alpha-weirdo
 Christian Urban <urbanc@in.tum.de> parents: 
1320diff
changeset | 65 | apply(rule alpha_weird.intros) | 
| 
bfd9af005e23
preliinary test about alpha-weirdo
 Christian Urban <urbanc@in.tum.de> parents: 
1320diff
changeset | 66 | apply(simp) | 
| 
bfd9af005e23
preliinary test about alpha-weirdo
 Christian Urban <urbanc@in.tum.de> parents: 
1320diff
changeset | 67 | apply(simp add: alpha_gen) | 
| 
bfd9af005e23
preliinary test about alpha-weirdo
 Christian Urban <urbanc@in.tum.de> parents: 
1320diff
changeset | 68 | using a | 
| 
bfd9af005e23
preliinary test about alpha-weirdo
 Christian Urban <urbanc@in.tum.de> parents: 
1320diff
changeset | 69 | apply(simp) | 
| 
bfd9af005e23
preliinary test about alpha-weirdo
 Christian Urban <urbanc@in.tum.de> parents: 
1320diff
changeset | 70 | *) | 
| 
bfd9af005e23
preliinary test about alpha-weirdo
 Christian Urban <urbanc@in.tum.de> parents: 
1320diff
changeset | 71 | |
| 954 
c009d2535896
very rough example file for how nominal2 specification can be parsed
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 72 | text {* example 1 *}
 | 
| 1251 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 73 | |
| 954 
c009d2535896
very rough example file for how nominal2 specification can be parsed
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 74 | nominal_datatype lam = | 
| 
c009d2535896
very rough example file for how nominal2 specification can be parsed
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 75 | VAR "name" | 
| 
c009d2535896
very rough example file for how nominal2 specification can be parsed
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 76 | | APP "lam" "lam" | 
| 1287 
8557af71724e
slight simplification of the raw-decl generation
 Christian Urban <urbanc@in.tum.de> parents: 
1285diff
changeset | 77 | | LET bp::"bp" t::"lam" bind "bi bp" in t | 
| 954 
c009d2535896
very rough example file for how nominal2 specification can be parsed
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 78 | and bp = | 
| 1287 
8557af71724e
slight simplification of the raw-decl generation
 Christian Urban <urbanc@in.tum.de> parents: 
1285diff
changeset | 79 | BP "name" "lam" | 
| 954 
c009d2535896
very rough example file for how nominal2 specification can be parsed
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 80 | binder | 
| 1295 
0ecc775e5fce
Link calls to Raw permutations, FV definition and alpha_definition into the parser.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1290diff
changeset | 81 | bi::"bp \<Rightarrow> atom set" | 
| 954 
c009d2535896
very rough example file for how nominal2 specification can be parsed
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 82 | where | 
| 1295 
0ecc775e5fce
Link calls to Raw permutations, FV definition and alpha_definition into the parser.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1290diff
changeset | 83 |   "bi (BP x t) = {atom x}"
 | 
| 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 | 84 | |
| 
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 | 85 | typ lam_raw | 
| 
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 | 86 | term VAR_raw | 
| 1283 
6a133abb7633
generated the "binding list" from the input; at the moment it is only printed out as tracing; does not yet include the "bind itself binders"
 Christian Urban <urbanc@in.tum.de> parents: 
1272diff
changeset | 87 | term APP_raw | 
| 
6a133abb7633
generated the "binding list" from the input; at the moment it is only printed out as tracing; does not yet include the "bind itself binders"
 Christian Urban <urbanc@in.tum.de> parents: 
1272diff
changeset | 88 | term LET_raw | 
| 1251 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 89 | term Test.BP_raw | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 90 | thm bi_raw.simps | 
| 1295 
0ecc775e5fce
Link calls to Raw permutations, FV definition and alpha_definition into the parser.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1290diff
changeset | 91 | thm permute_lam_raw_permute_bp_raw.simps | 
| 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 | 92 | thm alpha_lam_raw_alpha_bp_raw.intros[no_vars] | 
| 
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 | 93 | thm fv_lam_raw_fv_bp_raw.simps[no_vars] | 
| 1309 | 94 | thm eqvts | 
| 1299 
cbcd4997dac5
most tests work - the ones that do not I commented out
 Christian Urban <urbanc@in.tum.de> parents: 
1298diff
changeset | 95 | |
| 961 
0f88e04eb486
Variable takes a 'name'.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
954diff
changeset | 96 | print_theorems | 
| 954 
c009d2535896
very rough example file for how nominal2 specification can be parsed
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 97 | |
| 1261 
853abc14c5c6
added IsaMakefile...but so far included only a test for the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1258diff
changeset | 98 | text {* example 2 *}
 | 
| 
853abc14c5c6
added IsaMakefile...but so far included only a test for the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1258diff
changeset | 99 | |
| 1265 
fc8f5897b00a
first attempt to make sense out of the core-haskell definition
 Christian Urban <urbanc@in.tum.de> parents: 
1261diff
changeset | 100 | nominal_datatype trm' = | 
| 961 
0f88e04eb486
Variable takes a 'name'.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
954diff
changeset | 101 | Var "name" | 
| 1265 
fc8f5897b00a
first attempt to make sense out of the core-haskell definition
 Christian Urban <urbanc@in.tum.de> parents: 
1261diff
changeset | 102 | | App "trm'" "trm'" | 
| 1272 
6d140b2c751f
added ott-example about Leroy96
 Christian Urban <urbanc@in.tum.de> parents: 
1265diff
changeset | 103 | | Lam x::"name" t::"trm'" bind x in t | 
| 1265 
fc8f5897b00a
first attempt to make sense out of the core-haskell definition
 Christian Urban <urbanc@in.tum.de> parents: 
1261diff
changeset | 104 | | Let p::"pat'" "trm'" t::"trm'" bind "f p" in t | 
| 
fc8f5897b00a
first attempt to make sense out of the core-haskell definition
 Christian Urban <urbanc@in.tum.de> parents: 
1261diff
changeset | 105 | and pat' = | 
| 954 
c009d2535896
very rough example file for how nominal2 specification can be parsed
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 106 | PN | 
| 
c009d2535896
very rough example file for how nominal2 specification can be parsed
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 107 | | PS "name" | 
| 1251 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 108 | | PD "name" "name" | 
| 954 
c009d2535896
very rough example file for how nominal2 specification can be parsed
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 109 | binder | 
| 1299 
cbcd4997dac5
most tests work - the ones that do not I commented out
 Christian Urban <urbanc@in.tum.de> parents: 
1298diff
changeset | 110 | f::"pat' \<Rightarrow> atom set" | 
| 954 
c009d2535896
very rough example file for how nominal2 specification can be parsed
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 111 | where | 
| 978 
b44592adf235
Improper interface for datatype and function packages and proper interface lateron.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
973diff
changeset | 112 |   "f PN = {}"
 | 
| 1299 
cbcd4997dac5
most tests work - the ones that do not I commented out
 Christian Urban <urbanc@in.tum.de> parents: 
1298diff
changeset | 113 | | "f (PS x) = {atom x}"
 | 
| 
cbcd4997dac5
most tests work - the ones that do not I commented out
 Christian Urban <urbanc@in.tum.de> parents: 
1298diff
changeset | 114 | | "f (PD x y) = {atom x, atom y}"
 | 
| 1251 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 115 | |
| 1320 
447666754176
Another problem with permutations in alpha and possibly also in fv
 Christian Urban <urbanc@in.tum.de> parents: 
1319diff
changeset | 116 | |
| 
447666754176
Another problem with permutations in alpha and possibly also in fv
 Christian Urban <urbanc@in.tum.de> parents: 
1319diff
changeset | 117 | thm alpha_trm'_raw_alpha_pat'_raw.intros[no_vars] | 
| 
447666754176
Another problem with permutations in alpha and possibly also in fv
 Christian Urban <urbanc@in.tum.de> parents: 
1319diff
changeset | 118 | thm fv_trm'_raw_fv_pat'_raw.simps[no_vars] | 
| 1251 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 119 | thm f_raw.simps | 
| 954 
c009d2535896
very rough example file for how nominal2 specification can be parsed
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 120 | |
| 1251 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 121 | nominal_datatype trm0 = | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 122 | Var0 "name" | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 123 | | App0 "trm0" "trm0" | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 124 | | Lam0 x::"name" t::"trm0" 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 | 125 | | Let0 p::"pat0" "trm0" t::"trm0" bind "f0 p" 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 | 126 | and pat0 = | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 127 | PN0 | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 128 | | PS0 "name" | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 129 | | PD0 "pat0" "pat0" | 
| 954 
c009d2535896
very rough example file for how nominal2 specification can be parsed
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 130 | binder | 
| 1299 
cbcd4997dac5
most tests work - the ones that do not I commented out
 Christian Urban <urbanc@in.tum.de> parents: 
1298diff
changeset | 131 | f0::"pat0 \<Rightarrow> atom set" | 
| 954 
c009d2535896
very rough example file for how nominal2 specification can be parsed
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 132 | where | 
| 1251 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 133 |   "f0 PN0 = {}"
 | 
| 1299 
cbcd4997dac5
most tests work - the ones that do not I commented out
 Christian Urban <urbanc@in.tum.de> parents: 
1298diff
changeset | 134 | | "f0 (PS0 x) = {atom x}"
 | 
| 1251 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 135 | | "f0 (PD0 p1 p2) = (f0 p1) \<union> (f0 p2)" | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 136 | |
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 137 | thm f0_raw.simps | 
| 954 
c009d2535896
very rough example file for how nominal2 specification can be parsed
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 138 | |
| 
c009d2535896
very rough example file for how nominal2 specification can be parsed
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 139 | text {* example type schemes *}
 | 
| 1285 
e3ac56d3b7af
added example from my phd
 Christian Urban <urbanc@in.tum.de> parents: 
1284diff
changeset | 140 | |
| 1304 
dc65319809cc
Change type schemes to name set.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1302diff
changeset | 141 | (* does not work yet | 
| 1302 
d3101a5b9c4d
Length fix for nested recursions.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1299diff
changeset | 142 | nominal_datatype t = | 
| 
d3101a5b9c4d
Length fix for nested recursions.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1299diff
changeset | 143 | Var "name" | 
| 1265 
fc8f5897b00a
first attempt to make sense out of the core-haskell definition
 Christian Urban <urbanc@in.tum.de> parents: 
1261diff
changeset | 144 | | Fun "t" "t" | 
| 954 
c009d2535896
very rough example file for how nominal2 specification can be parsed
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 145 | |
| 1302 
d3101a5b9c4d
Length fix for nested recursions.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1299diff
changeset | 146 | nominal_datatype tyS = | 
| 1299 
cbcd4997dac5
most tests work - the ones that do not I commented out
 Christian Urban <urbanc@in.tum.de> parents: 
1298diff
changeset | 147 | All xs::"name list" ty::"t_raw" bind xs in ty | 
| 
cbcd4997dac5
most tests work - the ones that do not I commented out
 Christian Urban <urbanc@in.tum.de> parents: 
1298diff
changeset | 148 | *) | 
| 954 
c009d2535896
very rough example file for how nominal2 specification can be parsed
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 149 | |
| 1299 
cbcd4997dac5
most tests work - the ones that do not I commented out
 Christian Urban <urbanc@in.tum.de> parents: 
1298diff
changeset | 150 | nominal_datatype t = | 
| 
cbcd4997dac5
most tests work - the ones that do not I commented out
 Christian Urban <urbanc@in.tum.de> parents: 
1298diff
changeset | 151 | Var "name" | 
| 
cbcd4997dac5
most tests work - the ones that do not I commented out
 Christian Urban <urbanc@in.tum.de> parents: 
1298diff
changeset | 152 | | Fun "t" "t" | 
| 
cbcd4997dac5
most tests work - the ones that do not I commented out
 Christian Urban <urbanc@in.tum.de> parents: 
1298diff
changeset | 153 | and tyS = | 
| 1304 
dc65319809cc
Change type schemes to name set.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1302diff
changeset | 154 | All xs::"name set" ty::"t" bind xs in ty | 
| 954 
c009d2535896
very rough example file for how nominal2 specification can be parsed
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 155 | |
| 1251 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 156 | (* example 1 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 | 157 | |
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 158 | nominal_datatype trm1 = | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 159 | Vr1 "name" | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 160 | | Ap1 "trm1" "trm1" | 
| 1290 
a7e7ffb7f362
modified for new binding format - hope it is the intended one
 Christian Urban <urbanc@in.tum.de> parents: 
1287diff
changeset | 161 | | Lm1 x::"name" t::"trm1" bind x in t | 
| 1251 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 162 | | Lt1 p::"bp1" "trm1" t::"trm1" bind "bv1 p" 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 | 163 | and bp1 = | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 164 | BUnit1 | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 165 | | BV1 "name" | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 166 | | BP1 "bp1" "bp1" | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 167 | binder | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 168 | bv1 | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 169 | where | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 170 |   "bv1 (BUnit1) = {}"
 | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 171 | | "bv1 (BV1 x) = {atom x}"
 | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 172 | | "bv1 (BP1 bp1 bp2) = (bv1 bp1) \<union> (bv1 bp2)" | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 173 | |
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 174 | thm bv1_raw.simps | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 175 | |
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 176 | (* example 2 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 | 177 | |
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 178 | nominal_datatype trm2 = | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 179 | Vr2 "name" | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 180 | | Ap2 "trm2" "trm2" | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 181 | | Lm2 x::"name" t::"trm2" 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 | 182 | | Lt2 r::"rassign" t::"trm2" bind "bv2 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 | 183 | and rassign = | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 184 | As "name" "trm2" | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 185 | binder | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 186 | bv2 | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 187 | where | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 188 |   "bv2 (As x t) = {atom x}"
 | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 189 | |
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 190 | (* example 3 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 | 191 | |
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 192 | 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 | 193 | 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 | 194 | | 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 | 195 | | 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 | 196 | | 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 | 197 | 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 | 198 | ANil | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 199 | | 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 | 200 | binder | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 201 | bv3 | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 202 | where | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 203 |   "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 | 204 | | "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 | 205 | |
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 206 | (* example 4 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 | 207 | |
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 208 | nominal_datatype trm4 = | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 209 | Vr4 "name" | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 210 | | Ap4 "trm4" "trm4 list" | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 211 | | Lm4 x::"name" t::"trm4" bind x in t | 
| 1302 
d3101a5b9c4d
Length fix for nested recursions.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1299diff
changeset | 212 | |
| 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 | 213 | thm alpha_trm4_raw_alpha_trm4_raw_list.intros[no_vars] | 
| 
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 | 214 | thm fv_trm4_raw_fv_trm4_raw_list.simps[no_vars] | 
| 1251 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 215 | |
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 216 | (* 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 | 217 | |
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 218 | 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 | 219 | 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 | 220 | | 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 | 221 | | 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 | 222 | 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 | 223 | Lnil | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 224 | | 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 | 225 | binder | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 226 | bv5 | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 227 | where | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 228 |   "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 | 229 | | "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 | 230 | |
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 231 | (* example 6 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 | 232 | |
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 233 | nominal_datatype trm6 = | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 234 | Vr6 "name" | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 235 | | Lm6 x::"name" t::"trm6" 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 | 236 | | Lt6 left::"trm6" right::"trm6" bind "bv6 left" in right | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 237 | binder | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 238 | bv6 | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 239 | where | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 240 |   "bv6 (Vr6 n) = {}"
 | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 241 | | "bv6 (Lm6 n t) = {atom n} \<union> bv6 t"
 | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 242 | | "bv6 (Lt6 l r) = bv6 l \<union> bv6 r" | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 243 | |
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 244 | (* example 7 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 | 245 | |
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 246 | nominal_datatype trm7 = | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 247 | Vr7 "name" | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 248 | | Lm7 l::"name" r::"trm7" bind l in r | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 249 | | Lt7 l::"trm7" r::"trm7" bind "bv7 l" in r | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 250 | binder | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 251 | bv7 | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 252 | where | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 253 |   "bv7 (Vr7 n) = {atom n}"
 | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 254 | | "bv7 (Lm7 n t) = bv7 t - {atom n}"
 | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 255 | | "bv7 (Lt7 l r) = bv7 l \<union> bv7 r" | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 256 | |
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 257 | (* example 8 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 | 258 | |
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 259 | nominal_datatype foo8 = | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 260 | Foo0 "name" | 
| 1283 
6a133abb7633
generated the "binding list" from the input; at the moment it is only printed out as tracing; does not yet include the "bind itself binders"
 Christian Urban <urbanc@in.tum.de> parents: 
1272diff
changeset | 261 | | Foo1 b::"bar8" f::"foo8" bind "bv8 b" in f --"check fo error if this is called foo" | 
| 1251 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 262 | and bar8 = | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 263 | Bar0 "name" | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 264 | | Bar1 "name" s::"name" b::"bar8" bind s in b | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 265 | binder | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 266 | bv8 | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 267 | where | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 268 |   "bv8 (Bar0 x) = {}"
 | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 269 | | "bv8 (Bar1 v x b) = {atom v}"
 | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 270 | |
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 271 | (* example 9 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 | 272 | |
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 273 | nominal_datatype lam9 = | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 274 | Var9 "name" | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 275 | | Lam9 n::"name" l::"lam9" bind n in l | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 276 | and bla9 = | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 277 | Bla9 f::"lam9" s::"lam9" bind "bv9 f" in s | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 278 | binder | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 279 | bv9 | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 280 | where | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 281 |   "bv9 (Var9 x) = {}"
 | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 282 | | "bv9 (Lam9 x b) = {atom x}"
 | 
| 954 
c009d2535896
very rough example file for how nominal2 specification can be parsed
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 283 | |
| 1285 
e3ac56d3b7af
added example from my phd
 Christian Urban <urbanc@in.tum.de> parents: 
1284diff
changeset | 284 | (* example from my PHD *) | 
| 
e3ac56d3b7af
added example from my phd
 Christian Urban <urbanc@in.tum.de> parents: 
1284diff
changeset | 285 | |
| 
e3ac56d3b7af
added example from my phd
 Christian Urban <urbanc@in.tum.de> parents: 
1284diff
changeset | 286 | atom_decl coname | 
| 1272 
6d140b2c751f
added ott-example about Leroy96
 Christian Urban <urbanc@in.tum.de> parents: 
1265diff
changeset | 287 | |
| 1285 
e3ac56d3b7af
added example from my phd
 Christian Urban <urbanc@in.tum.de> parents: 
1284diff
changeset | 288 | nominal_datatype phd = | 
| 1290 
a7e7ffb7f362
modified for new binding format - hope it is the intended one
 Christian Urban <urbanc@in.tum.de> parents: 
1287diff
changeset | 289 | Ax "name" "coname" | 
| 
a7e7ffb7f362
modified for new binding format - hope it is the intended one
 Christian Urban <urbanc@in.tum.de> parents: 
1287diff
changeset | 290 | | Cut n::"coname" t1::"phd" c::"coname" t2::"phd" bind n in t1, bind c in t2 | 
| 
a7e7ffb7f362
modified for new binding format - hope it is the intended one
 Christian Urban <urbanc@in.tum.de> parents: 
1287diff
changeset | 291 | | AndR c1::"coname" t1::"phd" c2::"coname" t2::"phd" "coname" bind c1 in t1, bind c2 in t2 | 
| 
a7e7ffb7f362
modified for new binding format - hope it is the intended one
 Christian Urban <urbanc@in.tum.de> parents: 
1287diff
changeset | 292 | | AndL1 n::"name" t::"phd" "name" bind n in t | 
| 
a7e7ffb7f362
modified for new binding format - hope it is the intended one
 Christian Urban <urbanc@in.tum.de> parents: 
1287diff
changeset | 293 | | AndL2 n::"name" t::"phd" "name" bind n in t | 
| 
a7e7ffb7f362
modified for new binding format - hope it is the intended one
 Christian Urban <urbanc@in.tum.de> parents: 
1287diff
changeset | 294 | | ImpL c::"coname" t1::"phd" n::"name" t2::"phd" "name" bind c in t1, bind n in t2 | 
| 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 | 295 | | ImpR c::"coname" n::"name" t::"phd" "coname" bind n in t, bind c in t | 
| 
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 | 296 | |
| 
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 | 297 | (* PROBLEM?: why does it create for the Cut AndR ImpL cases | 
| 
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 | 298 | two permutations, but only one is used *) | 
| 
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 | 299 | thm alpha_phd_raw.intros[no_vars] | 
| 
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 | 300 | thm fv_phd_raw.simps[no_vars] | 
| 1285 
e3ac56d3b7af
added example from my phd
 Christian Urban <urbanc@in.tum.de> parents: 
1284diff
changeset | 301 | |
| 
e3ac56d3b7af
added example from my phd
 Christian Urban <urbanc@in.tum.de> parents: 
1284diff
changeset | 302 | (* example form Leroy 96 about modules; OTT *) | 
| 1284 | 303 | |
| 1272 
6d140b2c751f
added ott-example about Leroy96
 Christian Urban <urbanc@in.tum.de> parents: 
1265diff
changeset | 304 | nominal_datatype mexp = | 
| 1296 | 305 | Acc "path" | 
| 306 | | Stru "body" | |
| 307 | | Funct x::"name" "sexp" m::"mexp" bind x in m | |
| 308 | | FApp "mexp" "path" | |
| 309 | | Ascr "mexp" "sexp" | |
| 1272 
6d140b2c751f
added ott-example about Leroy96
 Christian Urban <urbanc@in.tum.de> parents: 
1265diff
changeset | 310 | and body = | 
| 
6d140b2c751f
added ott-example about Leroy96
 Christian Urban <urbanc@in.tum.de> parents: 
1265diff
changeset | 311 | Empty | 
| 1296 | 312 | | Seq c::defn d::"body" bind "cbinders c" in d | 
| 1272 
6d140b2c751f
added ott-example about Leroy96
 Christian Urban <urbanc@in.tum.de> parents: 
1265diff
changeset | 313 | and defn = | 
| 1296 | 314 | Type "name" "tyty" | 
| 315 | | Dty "name" | |
| 316 | | DStru "name" "mexp" | |
| 317 | | Val "name" "trmtrm" | |
| 1272 
6d140b2c751f
added ott-example about Leroy96
 Christian Urban <urbanc@in.tum.de> parents: 
1265diff
changeset | 318 | and sexp = | 
| 
6d140b2c751f
added ott-example about Leroy96
 Christian Urban <urbanc@in.tum.de> parents: 
1265diff
changeset | 319 | Sig sbody | 
| 1296 | 320 | | SFunc "name" "sexp" "sexp" | 
| 1272 
6d140b2c751f
added ott-example about Leroy96
 Christian Urban <urbanc@in.tum.de> parents: 
1265diff
changeset | 321 | and sbody = | 
| 
6d140b2c751f
added ott-example about Leroy96
 Christian Urban <urbanc@in.tum.de> parents: 
1265diff
changeset | 322 | SEmpty | 
| 1296 | 323 | | 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 | 324 | and spec = | 
| 1296 | 325 | Type1 "name" | 
| 326 | | Type2 "name" "tyty" | |
| 327 | | SStru "name" "sexp" | |
| 328 | | SVal "name" "tyty" | |
| 1272 
6d140b2c751f
added ott-example about Leroy96
 Christian Urban <urbanc@in.tum.de> parents: 
1265diff
changeset | 329 | and tyty = | 
| 1296 | 330 | Tyref1 "name" | 
| 331 | | Tyref2 "path" "tyty" | |
| 332 | | Fun "tyty" "tyty" | |
| 1272 
6d140b2c751f
added ott-example about Leroy96
 Christian Urban <urbanc@in.tum.de> parents: 
1265diff
changeset | 333 | and path = | 
| 1296 | 334 | Sref1 "name" | 
| 335 | | Sref2 "path" "name" | |
| 1272 
6d140b2c751f
added ott-example about Leroy96
 Christian Urban <urbanc@in.tum.de> parents: 
1265diff
changeset | 336 | and trmtrm = | 
| 1296 | 337 | Tref1 "name" | 
| 338 | | Tref2 "path" "name" | |
| 339 | | Lam v::"name" "tyty" M::"trmtrm" bind v in M | |
| 340 | | App "trmtrm" "trmtrm" | |
| 341 | | Let "body" "trmtrm" | |
| 1272 
6d140b2c751f
added ott-example about Leroy96
 Christian Urban <urbanc@in.tum.de> parents: 
1265diff
changeset | 342 | binder | 
| 
6d140b2c751f
added ott-example about Leroy96
 Christian Urban <urbanc@in.tum.de> parents: 
1265diff
changeset | 343 | cbinders :: "defn \<Rightarrow> atom set" | 
| 
6d140b2c751f
added ott-example about Leroy96
 Christian Urban <urbanc@in.tum.de> parents: 
1265diff
changeset | 344 | and Cbinders :: "spec \<Rightarrow> atom set" | 
| 
6d140b2c751f
added ott-example about Leroy96
 Christian Urban <urbanc@in.tum.de> parents: 
1265diff
changeset | 345 | where | 
| 
6d140b2c751f
added ott-example about Leroy96
 Christian Urban <urbanc@in.tum.de> parents: 
1265diff
changeset | 346 |   "cbinders (Type t T) = {atom t}"
 | 
| 
6d140b2c751f
added ott-example about Leroy96
 Christian Urban <urbanc@in.tum.de> parents: 
1265diff
changeset | 347 | | "cbinders (Dty t) = {atom t}"
 | 
| 
6d140b2c751f
added ott-example about Leroy96
 Christian Urban <urbanc@in.tum.de> parents: 
1265diff
changeset | 348 | | "cbinders (DStru x s) = {atom x}"
 | 
| 
6d140b2c751f
added ott-example about Leroy96
 Christian Urban <urbanc@in.tum.de> parents: 
1265diff
changeset | 349 | | "cbinders (Val v M) = {atom v}"
 | 
| 
6d140b2c751f
added ott-example about Leroy96
 Christian Urban <urbanc@in.tum.de> parents: 
1265diff
changeset | 350 | | "Cbinders (Type1 t) = {atom t}"
 | 
| 
6d140b2c751f
added ott-example about Leroy96
 Christian Urban <urbanc@in.tum.de> parents: 
1265diff
changeset | 351 | | "Cbinders (Type2 t T) = {atom t}"
 | 
| 
6d140b2c751f
added ott-example about Leroy96
 Christian Urban <urbanc@in.tum.de> parents: 
1265diff
changeset | 352 | | "Cbinders (SStru x S) = {atom x}"
 | 
| 
6d140b2c751f
added ott-example about Leroy96
 Christian Urban <urbanc@in.tum.de> parents: 
1265diff
changeset | 353 | | "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 | 354 | |
| 
fc8f5897b00a
first attempt to make sense out of the core-haskell definition
 Christian Urban <urbanc@in.tum.de> parents: 
1261diff
changeset | 355 | (* core haskell *) | 
| 
fc8f5897b00a
first attempt to make sense out of the core-haskell definition
 Christian Urban <urbanc@in.tum.de> parents: 
1261diff
changeset | 356 | |
| 
fc8f5897b00a
first attempt to make sense out of the core-haskell definition
 Christian Urban <urbanc@in.tum.de> parents: 
1261diff
changeset | 357 | atom_decl var | 
| 
fc8f5897b00a
first attempt to make sense out of the core-haskell definition
 Christian Urban <urbanc@in.tum.de> parents: 
1261diff
changeset | 358 | atom_decl tvar | 
| 
fc8f5897b00a
first attempt to make sense out of the core-haskell definition
 Christian Urban <urbanc@in.tum.de> parents: 
1261diff
changeset | 359 | atom_decl co | 
| 
fc8f5897b00a
first attempt to make sense out of the core-haskell definition
 Christian Urban <urbanc@in.tum.de> parents: 
1261diff
changeset | 360 | |
| 
fc8f5897b00a
first attempt to make sense out of the core-haskell definition
 Christian Urban <urbanc@in.tum.de> parents: 
1261diff
changeset | 361 | datatype sort = | 
| 
fc8f5897b00a
first attempt to make sense out of the core-haskell definition
 Christian Urban <urbanc@in.tum.de> parents: 
1261diff
changeset | 362 | TY tvar | 
| 
fc8f5897b00a
first attempt to make sense out of the core-haskell definition
 Christian Urban <urbanc@in.tum.de> parents: 
1261diff
changeset | 363 | | CO co | 
| 
fc8f5897b00a
first attempt to make sense out of the core-haskell definition
 Christian Urban <urbanc@in.tum.de> parents: 
1261diff
changeset | 364 | |
| 1299 
cbcd4997dac5
most tests work - the ones that do not I commented out
 Christian Urban <urbanc@in.tum.de> parents: 
1298diff
changeset | 365 | nominal_datatype kind = | 
| 1265 
fc8f5897b00a
first attempt to make sense out of the core-haskell definition
 Christian Urban <urbanc@in.tum.de> parents: 
1261diff
changeset | 366 | KStar | 
| 
fc8f5897b00a
first attempt to make sense out of the core-haskell definition
 Christian Urban <urbanc@in.tum.de> parents: 
1261diff
changeset | 367 | | KFun kind kind | 
| 
fc8f5897b00a
first attempt to make sense out of the core-haskell definition
 Christian Urban <urbanc@in.tum.de> parents: 
1261diff
changeset | 368 | | KEq kind kind | 
| 
fc8f5897b00a
first attempt to make sense out of the core-haskell definition
 Christian Urban <urbanc@in.tum.de> parents: 
1261diff
changeset | 369 | |
| 
fc8f5897b00a
first attempt to make sense out of the core-haskell definition
 Christian Urban <urbanc@in.tum.de> parents: 
1261diff
changeset | 370 | (* there are types, coercion types and regular types *) | 
| 1299 
cbcd4997dac5
most tests work - the ones that do not I commented out
 Christian Urban <urbanc@in.tum.de> parents: 
1298diff
changeset | 371 | (* | 
| 1265 
fc8f5897b00a
first attempt to make sense out of the core-haskell definition
 Christian Urban <urbanc@in.tum.de> parents: 
1261diff
changeset | 372 | nominal_datatype ty = | 
| 
fc8f5897b00a
first attempt to make sense out of the core-haskell definition
 Christian Urban <urbanc@in.tum.de> parents: 
1261diff
changeset | 373 | TVar tvar | 
| 
fc8f5897b00a
first attempt to make sense out of the core-haskell definition
 Christian Urban <urbanc@in.tum.de> parents: 
1261diff
changeset | 374 | | TFun string "ty list" | 
| 1299 
cbcd4997dac5
most tests work - the ones that do not I commented out
 Christian Urban <urbanc@in.tum.de> parents: 
1298diff
changeset | 375 | | TAll tvar kind_raw ty --"some binding" | 
| 1265 
fc8f5897b00a
first attempt to make sense out of the core-haskell definition
 Christian Urban <urbanc@in.tum.de> parents: 
1261diff
changeset | 376 | | TSym ty | 
| 
fc8f5897b00a
first attempt to make sense out of the core-haskell definition
 Christian Urban <urbanc@in.tum.de> parents: 
1261diff
changeset | 377 | | TCir ty ty | 
| 
fc8f5897b00a
first attempt to make sense out of the core-haskell definition
 Christian Urban <urbanc@in.tum.de> parents: 
1261diff
changeset | 378 | | TApp ty ty | 
| 
fc8f5897b00a
first attempt to make sense out of the core-haskell definition
 Christian Urban <urbanc@in.tum.de> parents: 
1261diff
changeset | 379 | | TLeft ty | 
| 
fc8f5897b00a
first attempt to make sense out of the core-haskell definition
 Christian Urban <urbanc@in.tum.de> parents: 
1261diff
changeset | 380 | | TRight ty | 
| 
fc8f5897b00a
first attempt to make sense out of the core-haskell definition
 Christian Urban <urbanc@in.tum.de> parents: 
1261diff
changeset | 381 | | TEq ty | 
| 
fc8f5897b00a
first attempt to make sense out of the core-haskell definition
 Christian Urban <urbanc@in.tum.de> parents: 
1261diff
changeset | 382 | | TRightc ty | 
| 
fc8f5897b00a
first attempt to make sense out of the core-haskell definition
 Christian Urban <urbanc@in.tum.de> parents: 
1261diff
changeset | 383 | | TLeftc ty | 
| 
fc8f5897b00a
first attempt to make sense out of the core-haskell definition
 Christian Urban <urbanc@in.tum.de> parents: 
1261diff
changeset | 384 | | TCoe ty ty | 
| 1299 
cbcd4997dac5
most tests work - the ones that do not I commented out
 Christian Urban <urbanc@in.tum.de> parents: 
1298diff
changeset | 385 | *) | 
| 1265 
fc8f5897b00a
first attempt to make sense out of the core-haskell definition
 Christian Urban <urbanc@in.tum.de> parents: 
1261diff
changeset | 386 | typedecl ty --"hack since ty is not yet defined" | 
| 
fc8f5897b00a
first attempt to make sense out of the core-haskell definition
 Christian Urban <urbanc@in.tum.de> parents: 
1261diff
changeset | 387 | |
| 
fc8f5897b00a
first attempt to make sense out of the core-haskell definition
 Christian Urban <urbanc@in.tum.de> parents: 
1261diff
changeset | 388 | abbreviation | 
| 
fc8f5897b00a
first attempt to make sense out of the core-haskell definition
 Christian Urban <urbanc@in.tum.de> parents: 
1261diff
changeset | 389 | "atoms A \<equiv> atom ` A" | 
| 
fc8f5897b00a
first attempt to make sense out of the core-haskell definition
 Christian Urban <urbanc@in.tum.de> parents: 
1261diff
changeset | 390 | |
| 1299 
cbcd4997dac5
most tests work - the ones that do not I commented out
 Christian Urban <urbanc@in.tum.de> parents: 
1298diff
changeset | 391 | (* does not work yet | 
| 1265 
fc8f5897b00a
first attempt to make sense out of the core-haskell definition
 Christian Urban <urbanc@in.tum.de> parents: 
1261diff
changeset | 392 | nominal_datatype trm = | 
| 
fc8f5897b00a
first attempt to make sense out of the core-haskell definition
 Christian Urban <urbanc@in.tum.de> parents: 
1261diff
changeset | 393 | Var var | 
| 1299 
cbcd4997dac5
most tests work - the ones that do not I commented out
 Christian Urban <urbanc@in.tum.de> parents: 
1298diff
changeset | 394 | | LAM tv::tvar kind_raw t::trm bind tv in t | 
| 1265 
fc8f5897b00a
first attempt to make sense out of the core-haskell definition
 Christian Urban <urbanc@in.tum.de> parents: 
1261diff
changeset | 395 | | APP trm ty | 
| 
fc8f5897b00a
first attempt to make sense out of the core-haskell definition
 Christian Urban <urbanc@in.tum.de> parents: 
1261diff
changeset | 396 | | Lam v::var ty t::trm bind v in t | 
| 
fc8f5897b00a
first attempt to make sense out of the core-haskell definition
 Christian Urban <urbanc@in.tum.de> parents: 
1261diff
changeset | 397 | | App trm trm | 
| 
fc8f5897b00a
first attempt to make sense out of the core-haskell definition
 Christian Urban <urbanc@in.tum.de> parents: 
1261diff
changeset | 398 | | Let x::var ty trm t::trm bind x in t | 
| 
fc8f5897b00a
first attempt to make sense out of the core-haskell definition
 Christian Urban <urbanc@in.tum.de> parents: 
1261diff
changeset | 399 | | Case trm "assoc list" | 
| 
fc8f5897b00a
first attempt to make sense out of the core-haskell definition
 Christian Urban <urbanc@in.tum.de> parents: 
1261diff
changeset | 400 | | Cast trm ty --"ty is supposed to be a coercion type only" | 
| 
fc8f5897b00a
first attempt to make sense out of the core-haskell definition
 Christian Urban <urbanc@in.tum.de> parents: 
1261diff
changeset | 401 | and assoc = | 
| 
fc8f5897b00a
first attempt to make sense out of the core-haskell definition
 Christian Urban <urbanc@in.tum.de> parents: 
1261diff
changeset | 402 | A p::pat t::trm bind "bv p" in t | 
| 
fc8f5897b00a
first attempt to make sense out of the core-haskell definition
 Christian Urban <urbanc@in.tum.de> parents: 
1261diff
changeset | 403 | and pat = | 
| 1299 
cbcd4997dac5
most tests work - the ones that do not I commented out
 Christian Urban <urbanc@in.tum.de> parents: 
1298diff
changeset | 404 | K string "(tvar \<times> kind_raw) list" "(var \<times> ty) list" | 
| 1265 
fc8f5897b00a
first attempt to make sense out of the core-haskell definition
 Christian Urban <urbanc@in.tum.de> parents: 
1261diff
changeset | 405 | binder | 
| 
fc8f5897b00a
first attempt to make sense out of the core-haskell definition
 Christian Urban <urbanc@in.tum.de> parents: 
1261diff
changeset | 406 | bv :: "pat \<Rightarrow> atom set" | 
| 
fc8f5897b00a
first attempt to make sense out of the core-haskell definition
 Christian Urban <urbanc@in.tum.de> parents: 
1261diff
changeset | 407 | where | 
| 
fc8f5897b00a
first attempt to make sense out of the core-haskell definition
 Christian Urban <urbanc@in.tum.de> parents: 
1261diff
changeset | 408 | "bv (K s ts vs) = (atoms (set (map fst ts))) \<union> (atoms (set (map fst vs)))" | 
| 1299 
cbcd4997dac5
most tests work - the ones that do not I commented out
 Christian Urban <urbanc@in.tum.de> parents: 
1298diff
changeset | 409 | *) | 
| 1265 
fc8f5897b00a
first attempt to make sense out of the core-haskell definition
 Christian Urban <urbanc@in.tum.de> parents: 
1261diff
changeset | 410 | |
| 1312 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 411 | (*thm bv_raw.simps*) | 
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 412 | |
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 413 | (* example 3 from Peter Sewell's bestiary *) | 
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 414 | nominal_datatype exp = | 
| 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 | 415 | Var "name" | 
| 
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 | 416 | | App "exp" "exp" | 
| 
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 | 417 | | Lam x::"name" e::"exp" bind x in e | 
| 
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 | 418 | | Let x::"name" p::"pat" e1::"exp" e2::"exp" bind x in e2, bind "bp p" in e1 | 
| 1312 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 419 | and pat = | 
| 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 | 420 | PVar "name" | 
| 1312 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 421 | | PUnit | 
| 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 | 422 | | PPair "pat" "pat" | 
| 1312 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 423 | binder | 
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 424 | bp :: "pat \<Rightarrow> atom set" | 
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 425 | where | 
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 426 |   "bp (PVar x) = {atom x}"
 | 
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 427 | | "bp (PUnit) = {}"
 | 
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 428 | | "bp (PPair p1 p2) = bp p1 \<union> bp p2" | 
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 429 | |
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 430 | (* example 6 from Peter Sewell's bestiary *) | 
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 431 | nominal_datatype exp6 = | 
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 432 | EVar name | 
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 433 | | EPair exp6 exp6 | 
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 434 | | 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 | 435 | and pat6 = | 
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 436 | PVar name | 
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 437 | | PUnit | 
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 438 | | PPair pat6 pat6 | 
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 439 | binder | 
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 440 | bp6 :: "pat6 \<Rightarrow> atom set" | 
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 441 | where | 
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 442 |   "bp6 (PVar x) = {atom x}"
 | 
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 443 | | "bp6 (PUnit) = {}"
 | 
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 444 | | "bp6 (PPair p1 p2) = bp6 p1 \<union> bp6 p2" | 
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 445 | |
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 446 | (* example 7 from Peter Sewell's bestiary *) | 
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 447 | nominal_datatype exp7 = | 
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 448 | EVar name | 
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 449 | | EUnit | 
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 450 | | EPair exp7 exp7 | 
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 451 | | ELetRec l::lrbs e::exp7 bind "b7s l" in e, bind "b7s l" in l | 
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 452 | and lrb = | 
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 453 | Assign name exp7 | 
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 454 | and lrbs = | 
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 455 | Single lrb | 
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 456 | | More lrb lrbs | 
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 457 | binder | 
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 458 | 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 | 459 | b7s :: "lrbs \<Rightarrow> atom set" | 
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 460 | where | 
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 461 |   "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 | 462 | | "b7s (Single a) = b7 a" | 
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 463 | | "b7s (More a as) = (b7 a) \<union> (b7s as)" | 
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 464 | |
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 465 | (* example 8 from Peter Sewell's bestiary *) | 
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 466 | nominal_datatype exp8 = | 
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 467 | EVar name | 
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 468 | | EUnit | 
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 469 | | EPair exp8 exp8 | 
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 470 | | ELetRec l::lrbs8 e::exp8 bind "b_lrbs8 l" in e, bind "b_lrbs8 l" in l | 
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 471 | and fnclause = | 
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 472 | K x::name p::pat8 e::exp8 bind "b_pat p" in e | 
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 473 | and fnclauses = | 
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 474 | S fnclause | 
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 475 | | ORs fnclause fnclauses | 
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 476 | and lrb8 = | 
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 477 | Clause fnclauses | 
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 478 | and lrbs8 = | 
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 479 | Single lrb8 | 
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 480 | | More lrb8 lrbs8 | 
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 481 | and pat8 = | 
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 482 | PVar name | 
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 483 | | PUnit | 
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 484 | | PPair pat8 pat8 | 
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 485 | binder | 
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 486 | 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 | 487 | 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 | 488 | 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 | 489 | 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 | 490 | 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 | 491 | where | 
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 492 | "b_lrbs8 (Single l) = b_lrb8 l" | 
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 493 | | "b_lrbs8 (More l ls) = b_lrb8 l \<union> b_lrbs8 ls" | 
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 494 | | "b_pat (PVar x) = {atom x}"
 | 
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 495 | | "b_pat (PUnit) = {}"
 | 
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 496 | | "b_pat (PPair p1 p2) = b_pat p1 \<union> b_pat p2" | 
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 497 | | "b_fnclauses (S fc) = (b_fnclause fc)" | 
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 498 | | "b_fnclauses (ORs fc fcs) = (b_fnclause fc) \<union> (b_fnclauses fcs)" | 
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 499 | | "b_lrb8 (Clause fcs) = (b_fnclauses fcs)" | 
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 500 | | "b_fnclause (K x pat exp8) = {atom x}"
 | 
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 501 | |
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 502 | |
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 503 | |
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 504 | |
| 
b0eae8c93314
added some more examples from Peter Sewell's bestiary
 Christian Urban <urbanc@in.tum.de> parents: 
1304diff
changeset | 505 | (* 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 | 506 | (* 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 | 507 | |
| 961 
0f88e04eb486
Variable takes a 'name'.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
954diff
changeset | 508 | end | 
| 1223 
160343d86a6f
"raw"-ified the term-constructors and types given in the specification
 Christian Urban <urbanc@in.tum.de> parents: 
1194diff
changeset | 509 | |
| 
160343d86a6f
"raw"-ified the term-constructors and types given in the specification
 Christian Urban <urbanc@in.tum.de> parents: 
1194diff
changeset | 510 | |
| 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 | 511 |