| author | wenzelm | 
| Thu, 02 Jul 2015 12:39:08 +0200 | |
| changeset 60630 | fc7625ec7427 | 
| parent 55968 | 94242fa87638 | 
| child 67149 | e61557884799 | 
| permissions | -rw-r--r-- | 
| 31775 | 1 | (* Title: HOL/Tools/Function/sum_tree.ML | 
| 25555 | 2 | Author: Alexander Krauss, TU Muenchen | 
| 3 | ||
| 4 | Some common tools for working with sum types in balanced tree form. | |
| 5 | *) | |
| 6 | ||
| 49965 | 7 | signature SUM_TREE = | 
| 8 | sig | |
| 9 | val sumcase_split_ss: simpset | |
| 10 |   val access_top_down: {init: 'a, left: 'a -> 'a, right: 'a -> 'a} -> int -> int -> 'a
 | |
| 11 | val mk_sumT: typ -> typ -> typ | |
| 12 | val mk_sumcase: typ -> typ -> typ -> term -> term -> term | |
| 13 | val App: term -> term -> term | |
| 14 | val mk_inj: typ -> int -> int -> term -> term | |
| 15 | val mk_proj: typ -> int -> int -> term -> term | |
| 16 | val mk_sumcases: typ -> term list -> term | |
| 17 | end | |
| 18 | ||
| 55968 | 19 | structure Sum_Tree: SUM_TREE = | 
| 25555 | 20 | struct | 
| 21 | ||
| 22 | (* Theory dependencies *) | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33961diff
changeset | 23 | val sumcase_split_ss = | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49965diff
changeset | 24 |   simpset_of (put_simpset HOL_basic_ss @{context}
 | 
| 55642 
63beb38e9258
adapted to renaming of datatype 'cases' and 'recs' to 'case' and 'rec'
 blanchet parents: 
55414diff
changeset | 25 |     addsimps (@{thm Product_Type.split} :: @{thms sum.case}))
 | 
| 25555 | 26 | |
| 27 | (* top-down access in balanced tree *) | |
| 28 | fun access_top_down {left, right, init} len i =
 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33961diff
changeset | 29 | Balanced_Tree.access | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33961diff
changeset | 30 |     {left = (fn f => f o left), right = (fn f => f o right), init = I} len i init
 | 
| 25555 | 31 | |
| 32 | (* Sum types *) | |
| 37678 
0040bafffdef
"prod" and "sum" replace "*" and "+" respectively
 haftmann parents: 
37387diff
changeset | 33 | fun mk_sumT LT RT = Type (@{type_name Sum_Type.sum}, [LT, RT])
 | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33961diff
changeset | 34 | fun mk_sumcase TL TR T l r = | 
| 55414 
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
 blanchet parents: 
55393diff
changeset | 35 |   Const (@{const_name sum.case_sum}, (TL --> T) --> (TR --> T) --> mk_sumT TL TR --> T) $ l $ r
 | 
| 25555 | 36 | |
| 37 | val App = curry op $ | |
| 38 | ||
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33961diff
changeset | 39 | fun mk_inj ST n i = | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33961diff
changeset | 40 | access_top_down | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33961diff
changeset | 41 |   { init = (ST, I : term -> term),
 | 
| 37678 
0040bafffdef
"prod" and "sum" replace "*" and "+" respectively
 haftmann parents: 
37387diff
changeset | 42 |     left = (fn (T as Type (@{type_name Sum_Type.sum}, [LT, RT]), inj) =>
 | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33961diff
changeset | 43 |       (LT, inj o App (Const (@{const_name Inl}, LT --> T)))),
 | 
| 37678 
0040bafffdef
"prod" and "sum" replace "*" and "+" respectively
 haftmann parents: 
37387diff
changeset | 44 |     right =(fn (T as Type (@{type_name Sum_Type.sum}, [LT, RT]), inj) =>
 | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33961diff
changeset | 45 |       (RT, inj o App (Const (@{const_name Inr}, RT --> T))))} n i
 | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33961diff
changeset | 46 | |> snd | 
| 25555 | 47 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33961diff
changeset | 48 | fun mk_proj ST n i = | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33961diff
changeset | 49 | access_top_down | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33961diff
changeset | 50 |   { init = (ST, I : term -> term),
 | 
| 37678 
0040bafffdef
"prod" and "sum" replace "*" and "+" respectively
 haftmann parents: 
37387diff
changeset | 51 |     left = (fn (T as Type (@{type_name Sum_Type.sum}, [LT, RT]), proj) =>
 | 
| 55393 
ce5cebfaedda
 se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
 blanchet parents: 
51717diff
changeset | 52 |       (LT, App (Const (@{const_name Sum_Type.projl}, T --> LT)) o proj)),
 | 
| 37678 
0040bafffdef
"prod" and "sum" replace "*" and "+" respectively
 haftmann parents: 
37387diff
changeset | 53 |     right =(fn (T as Type (@{type_name Sum_Type.sum}, [LT, RT]), proj) =>
 | 
| 55393 
ce5cebfaedda
 se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
 blanchet parents: 
51717diff
changeset | 54 |       (RT, App (Const (@{const_name Sum_Type.projr}, T --> RT)) o proj))} n i
 | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33961diff
changeset | 55 | |> snd | 
| 25555 | 56 | |
| 57 | fun mk_sumcases T fs = | |
| 49965 | 58 | Balanced_Tree.make (fn ((f, fT), (g, gT)) => (mk_sumcase fT gT T f g, mk_sumT fT gT)) | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33961diff
changeset | 59 | (map (fn f => (f, domain_type (fastype_of f))) fs) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33961diff
changeset | 60 | |> fst | 
| 25555 | 61 | |
| 62 | end |