| author | wenzelm | 
| Tue, 14 Nov 2017 13:56:38 +0100 | |
| changeset 67070 | 85e6c1ff5be3 | 
| parent 63617 | 3646e2ba554c | 
| child 69593 | 3dda49e08b9d | 
| permissions | -rw-r--r-- | 
| 36032 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 1 | (* Title: HOL/Tools/Predicate_Compile/predicate_compile_specialisation.ML | 
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 2 | Author: Lukas Bulwahn, TU Muenchen | 
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 3 | |
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 4 | Deriving specialised predicates and their intro rules | 
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 5 | *) | 
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 6 | |
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 7 | signature PREDICATE_COMPILE_SPECIALISATION = | 
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 8 | sig | 
| 55437 | 9 | val find_specialisations : string list -> (string * thm list) list -> | 
| 10 | theory -> (string * thm list) list * theory | |
| 36032 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 11 | end; | 
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 12 | |
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 13 | structure Predicate_Compile_Specialisation : PREDICATE_COMPILE_SPECIALISATION = | 
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 14 | struct | 
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 15 | |
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 16 | open Predicate_Compile_Aux; | 
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 17 | |
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 18 | (* table of specialisations *) | 
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 19 | structure Specialisations = Theory_Data | 
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 20 | ( | 
| 55437 | 21 | type T = (term * term) Item_Net.T | 
| 59058 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 wenzelm parents: 
58950diff
changeset | 22 | val empty : T = Item_Net.init (op aconv o apply2 fst) (single o fst) | 
| 55437 | 23 | val extend = I | 
| 24 | val merge = Item_Net.merge | |
| 36032 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 25 | ) | 
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 26 | |
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 27 | fun specialisation_of thy atom = | 
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 28 | Item_Net.retrieve (Specialisations.get thy) atom | 
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 29 | |
| 50056 | 30 | fun import (_, intros) args ctxt = | 
| 36032 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 31 | let | 
| 50056 | 32 | val ((_, intros'), ctxt') = Variable.importT intros ctxt | 
| 55437 | 33 | val pred' = | 
| 59582 | 34 | fst (strip_comb (HOLogic.dest_Trueprop (Logic.strip_imp_concl (Thm.prop_of (hd intros'))))) | 
| 36032 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 35 | val Ts = binder_types (fastype_of pred') | 
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 36 | val argTs = map fastype_of args | 
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 37 | val Tsubst = Type.raw_matches (argTs, Ts) Vartab.empty | 
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 38 | val args' = map (Envir.subst_term_types Tsubst) args | 
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 39 | in | 
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 40 | (((pred', intros'), args'), ctxt') | 
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 41 | end | 
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 42 | |
| 36036 
ea7d0df15be0
no specialisation for patterns with only tuples in the predicate compiler
 bulwahn parents: 
36032diff
changeset | 43 | (* patterns only constructed of variables and pairs/tuples are trivial constructor terms*) | 
| 62581 
fc5198b44314
clarified: constructors in the sense of the code generator are not invertible;
 haftmann parents: 
62391diff
changeset | 44 | fun is_nontrivial_constrt ctxt t = | 
| 36036 
ea7d0df15be0
no specialisation for patterns with only tuples in the predicate compiler
 bulwahn parents: 
36032diff
changeset | 45 | let | 
| 62581 
fc5198b44314
clarified: constructors in the sense of the code generator are not invertible;
 haftmann parents: 
62391diff
changeset | 46 | val lookup_constr = lookup_constr ctxt | 
| 55437 | 47 | fun check t = | 
| 48 | (case strip_comb t of | |
| 36036 
ea7d0df15be0
no specialisation for patterns with only tuples in the predicate compiler
 bulwahn parents: 
36032diff
changeset | 49 | (Var _, []) => (true, true) | 
| 
ea7d0df15be0
no specialisation for patterns with only tuples in the predicate compiler
 bulwahn parents: 
36032diff
changeset | 50 | | (Free _, []) => (true, true) | 
| 
ea7d0df15be0
no specialisation for patterns with only tuples in the predicate compiler
 bulwahn parents: 
36032diff
changeset | 51 |       | (Const (@{const_name Pair}, _), ts) =>
 | 
| 59058 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 wenzelm parents: 
58950diff
changeset | 52 | apply2 (forall I) (split_list (map check ts)) | 
| 62581 
fc5198b44314
clarified: constructors in the sense of the code generator are not invertible;
 haftmann parents: 
62391diff
changeset | 53 | | (Const cT, ts) => | 
| 
fc5198b44314
clarified: constructors in the sense of the code generator are not invertible;
 haftmann parents: 
62391diff
changeset | 54 | (case lookup_constr cT of | 
| 
fc5198b44314
clarified: constructors in the sense of the code generator are not invertible;
 haftmann parents: 
62391diff
changeset | 55 | SOME i => (false, | 
| 
fc5198b44314
clarified: constructors in the sense of the code generator are not invertible;
 haftmann parents: 
62391diff
changeset | 56 | length ts = i andalso forall (snd o check) ts) | 
| 36036 
ea7d0df15be0
no specialisation for patterns with only tuples in the predicate compiler
 bulwahn parents: 
36032diff
changeset | 57 | | _ => (false, false)) | 
| 
ea7d0df15be0
no specialisation for patterns with only tuples in the predicate compiler
 bulwahn parents: 
36032diff
changeset | 58 | | _ => (false, false)) | 
| 55437 | 59 | in check t = (false, true) end | 
| 36032 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 60 | |
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 61 | fun specialise_intros black_list (pred, intros) pats thy = | 
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 62 | let | 
| 42361 | 63 | val ctxt = Proof_Context.init_global thy | 
| 59582 | 64 | val maxidx = fold (Term.maxidx_term o Thm.prop_of) intros ~1 | 
| 59787 | 65 | val pats = map (Logic.incr_indexes ([], [], maxidx + 1)) pats | 
| 36032 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 66 | val (((pred, intros), pats), ctxt') = import (pred, intros) pats ctxt | 
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 67 | val result_pats = map Var (fold_rev Term.add_vars pats []) | 
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 68 | fun mk_fresh_name names = | 
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 69 | let | 
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 70 | val name = | 
| 43324 
2b47822868e4
discontinued Name.variant to emphasize that this is old-style / indirect;
 wenzelm parents: 
42361diff
changeset | 71 | singleton (Name.variant_list names) | 
| 
2b47822868e4
discontinued Name.variant to emphasize that this is old-style / indirect;
 wenzelm parents: 
42361diff
changeset | 72 |             ("specialised_" ^ Long_Name.base_name (fst (dest_Const pred)))
 | 
| 36032 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 73 | val bname = Sign.full_bname thy name | 
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 74 | in | 
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 75 | if Sign.declared_const thy bname then | 
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 76 | mk_fresh_name (name :: names) | 
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 77 | else | 
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 78 | bname | 
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 79 | end | 
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 80 | val constname = mk_fresh_name [] | 
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 81 |     val constT = map fastype_of result_pats ---> @{typ bool}
 | 
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 82 | val specialised_const = Const (constname, constT) | 
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 83 | fun specialise_intro intro = | 
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 84 | (let | 
| 59582 | 85 | val (prems, concl) = Logic.strip_horn (Thm.prop_of intro) | 
| 58950 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 wenzelm parents: 
56239diff
changeset | 86 | val env = Pattern.unify (Context.Theory thy) | 
| 63617 | 87 | (HOLogic.mk_Trueprop (list_comb (pred, pats)), concl) Envir.init | 
| 36032 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 88 | val prems = map (Envir.norm_term env) prems | 
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 89 | val args = map (Envir.norm_term env) result_pats | 
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 90 | val concl = HOLogic.mk_Trueprop (list_comb (specialised_const, args)) | 
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 91 | val intro = Logic.list_implies (prems, concl) | 
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 92 | in | 
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 93 | SOME intro | 
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 94 | end handle Pattern.Unif => NONE) | 
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 95 | val specialised_intros_t = map_filter I (map specialise_intro intros) | 
| 55437 | 96 | val thy' = | 
| 56239 | 97 | Sign.add_consts [(Binding.name (Long_Name.base_name constname), constT, NoSyn)] thy | 
| 36032 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 98 | val specialised_intros = map (Skip_Proof.make_thm thy') specialised_intros_t | 
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 99 | val exported_intros = Variable.exportT ctxt' ctxt specialised_intros | 
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 100 | val [t, specialised_t] = Variable.exportT_terms ctxt' ctxt | 
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 101 | [list_comb (pred, pats), list_comb (specialised_const, result_pats)] | 
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 102 | val thy'' = Specialisations.map (Item_Net.update (t, specialised_t)) thy' | 
| 36246 
43fecedff8cf
added peephole optimisations to the predicate compiler; added structure Predicate_Compile_Simps for peephole optimisations
 bulwahn parents: 
36052diff
changeset | 103 | val optimised_intros = | 
| 
43fecedff8cf
added peephole optimisations to the predicate compiler; added structure Predicate_Compile_Simps for peephole optimisations
 bulwahn parents: 
36052diff
changeset | 104 | map_filter (Predicate_Compile_Aux.peephole_optimisation thy'') exported_intros | 
| 
43fecedff8cf
added peephole optimisations to the predicate compiler; added structure Predicate_Compile_Simps for peephole optimisations
 bulwahn parents: 
36052diff
changeset | 105 | val ([spec], thy''') = find_specialisations black_list [(constname, optimised_intros)] thy'' | 
| 40053 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
38759diff
changeset | 106 | val thy'''' = Core_Data.register_intros spec thy''' | 
| 36032 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 107 | in | 
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 108 | thy'''' | 
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 109 | end | 
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 110 | |
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 111 | and find_specialisations black_list specs thy = | 
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 112 | let | 
| 55399 | 113 | val ctxt = Proof_Context.init_global thy | 
| 36032 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 114 | val add_vars = fold_aterms (fn Var v => cons v | _ => I); | 
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 115 | fun fresh_free T free_names = | 
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 116 | let | 
| 43324 
2b47822868e4
discontinued Name.variant to emphasize that this is old-style / indirect;
 wenzelm parents: 
42361diff
changeset | 117 | val free_name = singleton (Name.variant_list free_names) "x" | 
| 36032 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 118 | in | 
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 119 | (Free (free_name, T), free_name :: free_names) | 
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 120 | end | 
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 121 | fun replace_term_and_restrict thy T t Tts free_names = | 
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 122 | let | 
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 123 | val (free, free_names') = fresh_free T free_names | 
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 124 | val Tts' = map (apsnd (Pattern.rewrite_term thy [(t, free)] [])) Tts | 
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 125 | val (ts', free_names'') = restrict_pattern' thy Tts' free_names' | 
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 126 | in | 
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 127 | (free :: ts', free_names'') | 
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 128 | end | 
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 129 | and restrict_pattern' thy [] free_names = ([], free_names) | 
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 130 | | restrict_pattern' thy ((T, Free (x, _)) :: Tts) free_names = | 
| 55437 | 131 | let | 
| 132 | val (ts', free_names') = restrict_pattern' thy Tts free_names | |
| 133 | in | |
| 134 | (Free (x, T) :: ts', free_names') | |
| 135 | end | |
| 36032 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 136 | | restrict_pattern' thy ((T as TFree _, t) :: Tts) free_names = | 
| 55437 | 137 | replace_term_and_restrict thy T t Tts free_names | 
| 50056 | 138 | | restrict_pattern' thy ((T as Type (Tcon, _), t) :: Tts) free_names = | 
| 55399 | 139 | case Ctr_Sugar.ctr_sugar_of ctxt Tcon of | 
| 36032 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 140 | NONE => replace_term_and_restrict thy T t Tts free_names | 
| 55440 | 141 |         | SOME {ctrs, ...} =>
 | 
| 142 | (case strip_comb t of | |
| 143 | (Const (s, _), ats) => | |
| 144 | (case AList.lookup (op =) (map_filter (try dest_Const) ctrs) s of | |
| 145 | SOME constr_T => | |
| 146 | let | |
| 147 | val (Ts', T') = strip_type constr_T | |
| 148 | val Tsubst = Type.raw_match (T', T) Vartab.empty | |
| 149 | val Ts = map (Envir.subst_type Tsubst) Ts' | |
| 150 | val (bts', free_names') = restrict_pattern' thy ((Ts ~~ ats) @ Tts) free_names | |
| 151 | val (ats', ts') = chop (length ats) bts' | |
| 152 | in | |
| 153 | (list_comb (Const (s, map fastype_of ats' ---> T), ats') :: ts', free_names') | |
| 154 | end | |
| 155 | | NONE => replace_term_and_restrict thy T t Tts free_names)) | |
| 36032 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 156 | fun restrict_pattern thy Ts args = | 
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 157 | let | 
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 158 | val args = map Logic.unvarify_global args | 
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 159 | val Ts = map Logic.unvarifyT_global Ts | 
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 160 | val free_names = fold Term.add_free_names args [] | 
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 161 | val (pat, _) = restrict_pattern' thy (Ts ~~ args) free_names | 
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 162 | in map Logic.varify_global pat end | 
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 163 | fun detect' atom thy = | 
| 55437 | 164 | (case strip_comb atom of | 
| 36032 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 165 | (pred as Const (pred_name, _), args) => | 
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 166 | let | 
| 55437 | 167 | val Ts = binder_types (Sign.the_const_type thy pred_name) | 
| 168 | val pats = restrict_pattern thy Ts args | |
| 169 | in | |
| 62581 
fc5198b44314
clarified: constructors in the sense of the code generator are not invertible;
 haftmann parents: 
62391diff
changeset | 170 | if (exists (is_nontrivial_constrt ctxt) pats) | 
| 55437 | 171 | orelse (has_duplicates (op =) (fold add_vars pats [])) then | 
| 172 | let | |
| 173 | val thy' = | |
| 174 | (case specialisation_of thy atom of | |
| 175 | [] => | |
| 176 | if member (op =) ((map fst specs) @ black_list) pred_name then | |
| 177 | thy | |
| 178 | else | |
| 179 | (case try (Core_Data.intros_of (Proof_Context.init_global thy)) pred_name of | |
| 180 | NONE => thy | |
| 181 | | SOME [] => thy | |
| 182 | | SOME intros => | |
| 183 | specialise_intros ((map fst specs) @ (pred_name :: black_list)) | |
| 184 | (pred, intros) pats thy) | |
| 185 | | _ :: _ => thy) | |
| 36032 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 186 | val atom' = | 
| 55437 | 187 | (case specialisation_of thy' atom of | 
| 36032 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 188 | [] => atom | 
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 189 | | (t, specialised_t) :: _ => | 
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 190 | let | 
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 191 | val subst = Pattern.match thy' (t, atom) (Vartab.empty, Vartab.empty) | 
| 55437 | 192 | in Envir.subst_term subst specialised_t end handle Pattern.MATCH => atom) | 
| 193 | (*FIXME: this exception could be handled earlier in specialisation_of *) | |
| 194 | in | |
| 195 | (atom', thy') | |
| 196 | end | |
| 197 | else (atom, thy) | |
| 198 | end | |
| 199 | | _ => (atom, thy)) | |
| 36032 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 200 | fun specialise' (constname, intros) thy = | 
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 201 | let | 
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 202 | (* FIXME: only necessary because of sloppy Logic.unvarify in restrict_pattern *) | 
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 203 | val intros = Drule.zero_var_indexes_list intros | 
| 59582 | 204 | val (intros_t', thy') = (fold_map o fold_map_atoms) detect' (map Thm.prop_of intros) thy | 
| 36032 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 205 | in | 
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 206 | ((constname, map (Skip_Proof.make_thm thy') intros_t'), thy') | 
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 207 | end | 
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 208 | in | 
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 209 | fold_map specialise' specs thy | 
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 210 | end | 
| 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 bulwahn parents: diff
changeset | 211 | |
| 62391 | 212 | end |