| author | wenzelm | 
| Sun, 12 Jan 2020 23:29:35 +0100 | |
| changeset 71371 | 1c4ec697bee5 | 
| parent 69597 | ff784d5a5bfb | 
| child 80634 | a90ab1ea6458 | 
| permissions | -rw-r--r-- | 
| 63731 
9f906a2eb0e7
added theory to provide workaround to support nested datatypes in quickcheck (until quickcheck is generalized to support it with new datatypes)
 blanchet parents: diff
changeset | 1 | theory Quickcheck_Nesting | 
| 
9f906a2eb0e7
added theory to provide workaround to support nested datatypes in quickcheck (until quickcheck is generalized to support it with new datatypes)
 blanchet parents: diff
changeset | 2 | imports Main | 
| 
9f906a2eb0e7
added theory to provide workaround to support nested datatypes in quickcheck (until quickcheck is generalized to support it with new datatypes)
 blanchet parents: diff
changeset | 3 | begin | 
| 
9f906a2eb0e7
added theory to provide workaround to support nested datatypes in quickcheck (until quickcheck is generalized to support it with new datatypes)
 blanchet parents: diff
changeset | 4 | |
| 
9f906a2eb0e7
added theory to provide workaround to support nested datatypes in quickcheck (until quickcheck is generalized to support it with new datatypes)
 blanchet parents: diff
changeset | 5 | ML \<open> | 
| 
9f906a2eb0e7
added theory to provide workaround to support nested datatypes in quickcheck (until quickcheck is generalized to support it with new datatypes)
 blanchet parents: diff
changeset | 6 | let | 
| 
9f906a2eb0e7
added theory to provide workaround to support nested datatypes in quickcheck (until quickcheck is generalized to support it with new datatypes)
 blanchet parents: diff
changeset | 7 | open BNF_FP_Def_Sugar | 
| 
9f906a2eb0e7
added theory to provide workaround to support nested datatypes in quickcheck (until quickcheck is generalized to support it with new datatypes)
 blanchet parents: diff
changeset | 8 | open BNF_LFP_Compat | 
| 
9f906a2eb0e7
added theory to provide workaround to support nested datatypes in quickcheck (until quickcheck is generalized to support it with new datatypes)
 blanchet parents: diff
changeset | 9 | |
| 69597 | 10 | val compat_plugin = Plugin_Name.declare_setup \<^binding>\<open>compat\<close>; | 
| 63731 
9f906a2eb0e7
added theory to provide workaround to support nested datatypes in quickcheck (until quickcheck is generalized to support it with new datatypes)
 blanchet parents: diff
changeset | 11 | |
| 
9f906a2eb0e7
added theory to provide workaround to support nested datatypes in quickcheck (until quickcheck is generalized to support it with new datatypes)
 blanchet parents: diff
changeset | 12 | fun compat fp_sugars = | 
| 
9f906a2eb0e7
added theory to provide workaround to support nested datatypes in quickcheck (until quickcheck is generalized to support it with new datatypes)
 blanchet parents: diff
changeset | 13 | perhaps (try (datatype_compat (map (fst o dest_Type o #T) fp_sugars))); | 
| 
9f906a2eb0e7
added theory to provide workaround to support nested datatypes in quickcheck (until quickcheck is generalized to support it with new datatypes)
 blanchet parents: diff
changeset | 14 | in | 
| 
9f906a2eb0e7
added theory to provide workaround to support nested datatypes in quickcheck (until quickcheck is generalized to support it with new datatypes)
 blanchet parents: diff
changeset | 15 | Theory.setup (fp_sugars_interpretation compat_plugin compat) | 
| 
9f906a2eb0e7
added theory to provide workaround to support nested datatypes in quickcheck (until quickcheck is generalized to support it with new datatypes)
 blanchet parents: diff
changeset | 16 | end | 
| 
9f906a2eb0e7
added theory to provide workaround to support nested datatypes in quickcheck (until quickcheck is generalized to support it with new datatypes)
 blanchet parents: diff
changeset | 17 | \<close> | 
| 
9f906a2eb0e7
added theory to provide workaround to support nested datatypes in quickcheck (until quickcheck is generalized to support it with new datatypes)
 blanchet parents: diff
changeset | 18 | |
| 
9f906a2eb0e7
added theory to provide workaround to support nested datatypes in quickcheck (until quickcheck is generalized to support it with new datatypes)
 blanchet parents: diff
changeset | 19 | end |