| author | Christian Urban <urbanc@in.tum.de> | 
| Thu, 19 Apr 2018 13:58:22 +0100 | |
| branch | Nominal2-Isabelle2016-1 | 
| changeset 3246 | 66114fa3d2ee | 
| parent 3235 | 5ebd327ffb96 | 
| permissions | -rw-r--r-- | 
| 3109 
d79e936e30ea
commented out parts of TypeScheme1 in order to run all tests
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1 | theory FiniteType | 
| 
d79e936e30ea
commented out parts of TypeScheme1 in order to run all tests
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 2 | imports "../Nominal2" | 
| 
d79e936e30ea
commented out parts of TypeScheme1 in order to run all tests
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 3 | begin | 
| 
d79e936e30ea
commented out parts of TypeScheme1 in order to run all tests
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 4 | |
| 
d79e936e30ea
commented out parts of TypeScheme1 in order to run all tests
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 5 | typedef ('a::pt, 'b::fs) ffun = "{f::'a => 'b. finite (supp f)}"
 | 
| 
d79e936e30ea
commented out parts of TypeScheme1 in order to run all tests
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 6 | apply(rule_tac x="\<lambda>_. undefined::'b::fs" in exI) | 
| 
d79e936e30ea
commented out parts of TypeScheme1 in order to run all tests
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 7 | apply(auto) | 
| 
d79e936e30ea
commented out parts of TypeScheme1 in order to run all tests
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 8 | apply(rule_tac S="supp (undefined::'b::fs)" in supports_finite) | 
| 
d79e936e30ea
commented out parts of TypeScheme1 in order to run all tests
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 9 | apply(simp add: supports_def) | 
| 
d79e936e30ea
commented out parts of TypeScheme1 in order to run all tests
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 10 | apply(perm_simp) | 
| 
d79e936e30ea
commented out parts of TypeScheme1 in order to run all tests
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 11 | apply(simp add: fresh_def[symmetric]) | 
| 
d79e936e30ea
commented out parts of TypeScheme1 in order to run all tests
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 12 | apply(simp add: swap_fresh_fresh) | 
| 
d79e936e30ea
commented out parts of TypeScheme1 in order to run all tests
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 13 | apply(simp add: finite_supp) | 
| 
d79e936e30ea
commented out parts of TypeScheme1 in order to run all tests
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 14 | done | 
| 
d79e936e30ea
commented out parts of TypeScheme1 in order to run all tests
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 15 | |
| 
d79e936e30ea
commented out parts of TypeScheme1 in order to run all tests
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 16 | |
| 
d79e936e30ea
commented out parts of TypeScheme1 in order to run all tests
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 17 | |
| 
d79e936e30ea
commented out parts of TypeScheme1 in order to run all tests
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 18 | |
| 3235 
5ebd327ffb96
changed nominal_primrec into the more appropriate nominal_function
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: 
3109diff
changeset | 19 | end |