| author | huffman | 
| Mon, 28 May 2007 03:45:41 +0200 | |
| changeset 23112 | 2bc882fbe51c | 
| parent 22790 | e1cff9268177 | 
| child 23577 | c5b93c69afd3 | 
| permissions | -rw-r--r-- | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 1 | (* Title: HOL/Tools/inductive_realizer.ML | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 2 | ID: $Id$ | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 3 | Author: Stefan Berghofer, TU Muenchen | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 4 | |
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 5 | Porgram extraction from proofs involving inductive predicates: | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 6 | Realizers for induction and elimination rules | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 7 | *) | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 8 | |
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 9 | signature INDUCTIVE_REALIZER = | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 10 | sig | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 11 | val add_ind_realizers: string -> string list -> theory -> theory | 
| 18708 | 12 | val setup: theory -> theory | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 13 | end; | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 14 | |
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 15 | structure InductiveRealizer : INDUCTIVE_REALIZER = | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 16 | struct | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 17 | |
| 22271 | 18 | (* FIXME: LocalTheory.note should return theorems with proper names! *) | 
| 22606 
962f824c2df9
- Tried to make name_of_thm more robust against changes of the
 berghofe parents: 
22596diff
changeset | 19 | fun name_of_thm thm = | 
| 
962f824c2df9
- Tried to make name_of_thm more robust against changes of the
 berghofe parents: 
22596diff
changeset | 20 | (case Symtab.dest (Proofterm.thms_of_proof' (proof_of thm) Symtab.empty) of | 
| 
962f824c2df9
- Tried to make name_of_thm more robust against changes of the
 berghofe parents: 
22596diff
changeset | 21 | [(name, _)] => name | 
| 
962f824c2df9
- Tried to make name_of_thm more robust against changes of the
 berghofe parents: 
22596diff
changeset | 22 |    | _ => error ("name_of_thm: bad proof of theorem\n" ^ string_of_thm thm));
 | 
| 22271 | 23 | |
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 24 | val all_simps = map (symmetric o mk_meta_eq) (thms "HOL.all_simps"); | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 25 | |
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 26 | fun prf_of thm = | 
| 22596 | 27 |   let val {thy, prop, der = (_, prf), ...} = rep_thm thm
 | 
| 28 |   in Reconstruct.expand_proof thy [("", NONE)] (Reconstruct.reconstruct_proof thy prop prf) end; (* FIXME *)
 | |
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 29 | |
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 30 | fun forall_intr_prf (t, prf) = | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 31 | let val (a, T) = (case t of Var ((a, _), T) => (a, T) | Free p => p) | 
| 15531 | 32 | in Abst (a, SOME T, Proofterm.prf_abstract_over t prf) end; | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 33 | |
| 22271 | 34 | fun forall_intr_term (t, u) = | 
| 35 | let val (a, T) = (case t of Var ((a, _), T) => (a, T) | Free p => p) | |
| 36 | in all T $ Abs (a, T, abstract_over (t, u)) end; | |
| 37 | ||
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 38 | fun subsets [] = [[]] | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 39 | | subsets (x::xs) = | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 40 | let val ys = subsets xs | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 41 | in ys @ map (cons x) ys end; | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 42 | |
| 22271 | 43 | val pred_of = fst o dest_Const o head_of; | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 44 | |
| 22271 | 45 | fun strip_all' used names (Const ("all", _) $ Abs (s, T, t)) =
 | 
| 46 | let val (s', names') = (case names of | |
| 47 | [] => (Name.variant used s, []) | |
| 48 | | name :: names' => (name, names')) | |
| 49 | in strip_all' (s'::used) names' (subst_bound (Free (s', T), t)) end | |
| 50 |   | strip_all' used names ((t as Const ("==>", _) $ P) $ Q) =
 | |
| 51 | t $ strip_all' used names Q | |
| 52 | | strip_all' _ _ t = t; | |
| 53 | ||
| 54 | fun strip_all t = strip_all' (add_term_free_names (t, [])) [] t; | |
| 55 | ||
| 56 | fun strip_one name (Const ("all", _) $ Abs (s, T, Const ("==>", _) $ P $ Q)) =
 | |
| 57 | (subst_bound (Free (name, T), P), subst_bound (Free (name, T), Q)) | |
| 58 |   | strip_one _ (Const ("==>", _) $ P $ Q) = (P, Q);
 | |
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 59 | |
| 15574 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 skalberg parents: 
15570diff
changeset | 60 | fun relevant_vars prop = foldr (fn | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 61 | (Var ((a, i), T), vs) => (case strip_type T of | 
| 22271 | 62 | (_, Type (s, _)) => if s mem ["bool"] then (a, T) :: vs else vs | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 63 | | _ => vs) | 
| 15574 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 skalberg parents: 
15570diff
changeset | 64 | | (_, vs) => vs) [] (term_vars prop); | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 65 | |
| 22271 | 66 | fun dt_of_intrs thy vs nparms intrs = | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 67 | let | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 68 | val iTs = term_tvars (prop_of (hd intrs)); | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 69 | val Tvs = map TVar iTs; | 
| 22271 | 70 | val (Const (s, _), ts) = strip_comb (HOLogic.dest_Trueprop | 
| 71 | (Logic.strip_imp_concl (prop_of (hd intrs)))); | |
| 72 | val params = map dest_Var (Library.take (nparms, ts)); | |
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 73 | val tname = space_implode "_" (Sign.base_name s ^ "T" :: vs); | 
| 22271 | 74 | fun constr_of_intr intr = (Sign.base_name (name_of_thm intr), | 
| 19806 | 75 | map (Logic.unvarifyT o snd) (rev (Term.add_vars (prop_of intr) []) \\ params) @ | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 76 | filter_out (equal Extraction.nullT) (map | 
| 19806 | 77 | (Logic.unvarifyT o Extraction.etype_of thy vs []) (prems_of intr)), | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 78 | NoSyn); | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 79 | in (map (fn a => "'" ^ a) vs @ map (fst o fst) iTs, tname, NoSyn, | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 80 | map constr_of_intr intrs) | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 81 | end; | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 82 | |
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 83 | fun mk_rlz T = Const ("realizes", [T, HOLogic.boolT] ---> HOLogic.boolT);
 | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 84 | |
| 22271 | 85 | (** turn "P" into "%r x. realizes r (P x)" **) | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 86 | |
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 87 | fun gen_rvar vs (t as Var ((a, 0), T)) = | 
| 22271 | 88 | if body_type T <> HOLogic.boolT then t else | 
| 89 | let | |
| 90 |           val U = TVar (("'" ^ a, 0), HOLogic.typeS)
 | |
| 91 | val Ts = binder_types T; | |
| 92 | val i = length Ts; | |
| 93 | val xs = map (pair "x") Ts; | |
| 94 | val u = list_comb (t, map Bound (i - 1 downto 0)) | |
| 95 | in | |
| 96 | if a mem vs then | |
| 97 |             list_abs (("r", U) :: xs, mk_rlz U $ Bound i $ u)
 | |
| 98 | else list_abs (xs, mk_rlz Extraction.nullT $ Extraction.nullt $ u) | |
| 99 | end | |
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 100 | | gen_rvar _ t = t; | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 101 | |
| 22271 | 102 | fun mk_realizes_eqn n vs nparms intrs = | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 103 | let | 
| 22271 | 104 | val concl = HOLogic.dest_Trueprop (concl_of (hd intrs)); | 
| 105 | val iTs = term_tvars concl; | |
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 106 | val Tvs = map TVar iTs; | 
| 22271 | 107 | val (h as Const (s, T), us) = strip_comb concl; | 
| 108 | val params = List.take (us, nparms); | |
| 109 | val elTs = List.drop (binder_types T, nparms); | |
| 110 | val predT = elTs ---> HOLogic.boolT; | |
| 111 | val used = map (fst o fst o dest_Var) params; | |
| 112 | val xs = map (Var o apfst (rpair 0)) | |
| 113 | (Name.variant_list used (replicate (length elTs) "x") ~~ elTs); | |
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 114 | val rT = if n then Extraction.nullT | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 115 | else Type (space_implode "_" (s ^ "T" :: vs), | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 116 |         map (fn a => TVar (("'" ^ a, 0), HOLogic.typeS)) vs @ Tvs);
 | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 117 | val r = if n then Extraction.nullt else Var ((Sign.base_name s, 0), rT); | 
| 22271 | 118 | val S = list_comb (h, params @ xs); | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 119 | val rvs = relevant_vars S; | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 120 | val vs' = map fst rvs \\ vs; | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 121 | val rname = space_implode "_" (s ^ "R" :: vs); | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 122 | |
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 123 | fun mk_Tprem n v = | 
| 17485 | 124 | let val T = (the o AList.lookup (op =) rvs) v | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 125 |       in (Const ("typeof", T --> Type ("Type", [])) $ Var ((v, 0), T),
 | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 126 | Extraction.mk_typ (if n then Extraction.nullT | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 127 |           else TVar (("'" ^ v, 0), HOLogic.typeS)))
 | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 128 | end; | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 129 | |
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 130 | val prems = map (mk_Tprem true) vs' @ map (mk_Tprem false) vs; | 
| 22271 | 131 | val ts = map (gen_rvar vs) params; | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 132 | val argTs = map fastype_of ts; | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 133 | |
| 22271 | 134 |   in ((prems, (Const ("typeof", HOLogic.boolT --> Type ("Type", [])) $ S,
 | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 135 | Extraction.mk_typ rT)), | 
| 22271 | 136 | (prems, (mk_rlz rT $ r $ S, | 
| 137 | if n then list_comb (Const (rname, argTs ---> predT), ts @ xs) | |
| 138 | else list_comb (Const (rname, argTs @ [rT] ---> predT), ts @ [r] @ xs)))) | |
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 139 | end; | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 140 | |
| 22271 | 141 | fun fun_of_prem thy rsets vs params rule ivs intr = | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 142 | let | 
| 22271 | 143 | val ctxt = ProofContext.init thy | 
| 144 | val args = map (Free o apfst fst o dest_Var) ivs; | |
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 145 | val args' = map (Free o apfst fst) | 
| 16861 | 146 | (Term.add_vars (prop_of intr) [] \\ params); | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 147 | val rule' = strip_all rule; | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 148 | val conclT = Extraction.etype_of thy vs [] (Logic.strip_imp_concl rule'); | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 149 | val used = map (fst o dest_Free) args; | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 150 | |
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 151 | fun is_rec t = not (null (term_consts t inter rsets)); | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 152 | |
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 153 |     fun is_meta (Const ("all", _) $ Abs (s, _, P)) = is_meta P
 | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 154 |       | is_meta (Const ("==>", _) $ _ $ Q) = is_meta Q
 | 
| 22271 | 155 |       | is_meta (Const ("Trueprop", _) $ t) = (case head_of t of
 | 
| 156 | Const (s, _) => can (InductivePackage.the_inductive ctxt) s | |
| 157 | | _ => true) | |
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 158 | | is_meta _ = false; | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 159 | |
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 160 | fun fun_of ts rts args used (prem :: prems) = | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 161 | let | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 162 | val T = Extraction.etype_of thy vs [] prem; | 
| 20071 
8f3e1ddb50e6
replaced Term.variant(list) by Name.variant(_list);
 wenzelm parents: 
19806diff
changeset | 163 | val [x, r] = Name.variant_list used ["x", "r"] | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 164 | in if T = Extraction.nullT | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 165 | then fun_of ts rts args used prems | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 166 | else if is_rec prem then | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 167 | if is_meta prem then | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 168 | let | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 169 | val prem' :: prems' = prems; | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 170 | val U = Extraction.etype_of thy vs [] prem'; | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 171 | in if U = Extraction.nullT | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 172 | then fun_of (Free (x, T) :: ts) | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 173 | (Free (r, binder_types T ---> HOLogic.unitT) :: rts) | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 174 | (Free (x, T) :: args) (x :: r :: used) prems' | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 175 | else fun_of (Free (x, T) :: ts) (Free (r, U) :: rts) | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 176 | (Free (r, U) :: Free (x, T) :: args) (x :: r :: used) prems' | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 177 | end | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 178 | else (case strip_type T of | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 179 |                   (Ts, Type ("*", [T1, T2])) =>
 | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 180 | let | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 181 | val fx = Free (x, Ts ---> T1); | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 182 | val fr = Free (r, Ts ---> T2); | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 183 | val bs = map Bound (length Ts - 1 downto 0); | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 184 | val t = list_abs (map (pair "z") Ts, | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 185 | HOLogic.mk_prod (list_comb (fx, bs), list_comb (fr, bs))) | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 186 | in fun_of (fx :: ts) (fr :: rts) (t::args) | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 187 | (x :: r :: used) prems | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 188 | end | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 189 | | (Ts, U) => fun_of (Free (x, T) :: ts) | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 190 | (Free (r, binder_types T ---> HOLogic.unitT) :: rts) | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 191 | (Free (x, T) :: args) (x :: r :: used) prems) | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 192 | else fun_of (Free (x, T) :: ts) rts (Free (x, T) :: args) | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 193 | (x :: used) prems | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 194 | end | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 195 | | fun_of ts rts args used [] = | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 196 | let val xs = rev (rts @ ts) | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 197 | in if conclT = Extraction.nullT | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 198 | then list_abs_free (map dest_Free xs, HOLogic.unit) | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 199 | else list_abs_free (map dest_Free xs, list_comb | 
| 22271 | 200 |               (Free ("r" ^ Sign.base_name (name_of_thm intr),
 | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 201 | map fastype_of (rev args) ---> conclT), rev args)) | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 202 | end | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 203 | |
| 13921 
69c627b6b28d
Fixed problem in add_elim_realizer which caused bound variables to
 berghofe parents: 
13725diff
changeset | 204 | in fun_of args' [] (rev args) used (Logic.strip_imp_prems rule') end; | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 205 | |
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 206 | fun indrule_realizer thy induct raw_induct rsets params vs rec_names rss intrs dummies = | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 207 | let | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 208 | val concls = HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of raw_induct)); | 
| 15570 | 209 | val premss = List.mapPartial (fn (s, rs) => if s mem rsets then | 
| 22271 | 210 | SOME (rs, map (fn (_, r) => List.nth (prems_of raw_induct, | 
| 15570 | 211 | find_index_eq (prop_of r) (map prop_of intrs))) rs) else NONE) rss; | 
| 22271 | 212 | val fs = maps (fn ((intrs, prems), dummy) => | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 213 | let | 
| 22271 | 214 | val fs = map (fn (rule, (ivs, intr)) => | 
| 215 | fun_of_prem thy rsets vs params rule ivs intr) (prems ~~ intrs) | |
| 216 |       in if dummy then Const ("arbitrary",
 | |
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 217 | HOLogic.unitT --> body_type (fastype_of (hd fs))) :: fs | 
| 22271 | 218 | else fs | 
| 219 | end) (premss ~~ dummies); | |
| 16861 | 220 | val frees = fold Term.add_frees fs []; | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 221 | val Ts = map fastype_of fs; | 
| 22271 | 222 | fun name_of_fn intr = "r" ^ Sign.base_name (name_of_thm intr) | 
| 223 | in | |
| 224 | fst (fold_map (fn concl => fn names => | |
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 225 | let val T = Extraction.etype_of thy vs [] concl | 
| 22271 | 226 | in if T = Extraction.nullT then (Extraction.nullt, names) else | 
| 227 | let | |
| 228 |           val Type ("fun", [U, _]) = T;
 | |
| 229 | val a :: names' = names | |
| 230 |         in (list_abs_free (("x", U) :: List.mapPartial (fn intr =>
 | |
| 231 | Option.map (pair (name_of_fn intr)) | |
| 232 | (AList.lookup (op =) frees (name_of_fn intr))) intrs, | |
| 233 |           list_comb (Const (a, Ts ---> T), fs) $ Free ("x", U)), names')
 | |
| 234 | end | |
| 235 | end) concls rec_names) | |
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 236 | end; | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 237 | |
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 238 | fun add_dummy name dname (x as (_, (vs, s, mfx, cs))) = | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 239 | if name = s then (true, (vs, s, mfx, (dname, [HOLogic.unitT], NoSyn) :: cs)) | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 240 | else x; | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 241 | |
| 18314 | 242 | fun add_dummies f [] _ thy = | 
| 243 | (([], NONE), thy) | |
| 244 | | add_dummies f dts used thy = | |
| 245 | thy | |
| 246 | |> f (map snd dts) | |
| 247 | |-> (fn dtinfo => pair ((map fst dts), SOME dtinfo)) | |
| 248 | handle DatatypeAux.Datatype_Empty name' => | |
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 249 | let | 
| 14888 
99ac3eb0f84e
add_dummies no longer uses transform_error but handles specific
 berghofe parents: 
13928diff
changeset | 250 | val name = Sign.base_name name'; | 
| 20071 
8f3e1ddb50e6
replaced Term.variant(list) by Name.variant(_list);
 wenzelm parents: 
19806diff
changeset | 251 | val dname = Name.variant used "Dummy" | 
| 18314 | 252 | in | 
| 253 | thy | |
| 254 | |> add_dummies f (map (add_dummy name dname) dts) (dname :: used) | |
| 14888 
99ac3eb0f84e
add_dummies no longer uses transform_error but handles specific
 berghofe parents: 
13928diff
changeset | 255 | end; | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 256 | |
| 22271 | 257 | fun mk_realizer thy vs (name, rule, rrule, rlz, rt) = | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 258 | let | 
| 13725 
12404b452034
Changed format of realizers / correctness proofs.
 berghofe parents: 
13710diff
changeset | 259 | val rvs = map fst (relevant_vars (prop_of rule)); | 
| 16861 | 260 | val xs = rev (Term.add_vars (prop_of rule) []); | 
| 13725 
12404b452034
Changed format of realizers / correctness proofs.
 berghofe parents: 
13710diff
changeset | 261 | val vs1 = map Var (filter_out (fn ((a, _), _) => a mem rvs) xs); | 
| 16861 | 262 | val rlzvs = rev (Term.add_vars (prop_of rrule) []); | 
| 17485 | 263 | val vs2 = map (fn (ixn, _) => Var (ixn, (the o AList.lookup (op =) rlzvs) ixn)) xs; | 
| 22271 | 264 | val rs = map Var (subtract (op = o pairself fst) xs rlzvs); | 
| 265 | val rlz' = foldr forall_intr_term (prop_of rrule) (vs2 @ rs); | |
| 266 | val rlz'' = foldr forall_intr_term rlz vs2 | |
| 267 | in (name, (vs, | |
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 268 | if rt = Extraction.nullt then rt else | 
| 15574 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 skalberg parents: 
15570diff
changeset | 269 | foldr (uncurry lambda) rt vs1, | 
| 22271 | 270 | ProofRewriteRules.un_hhf_proof rlz' rlz'' | 
| 271 | (foldr forall_intr_prf (prf_of rrule) (vs2 @ rs)))) | |
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 272 | end; | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 273 | |
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 274 | fun add_ind_realizer rsets intrs induct raw_induct elims (thy, vs) = | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 275 | let | 
| 22271 | 276 | val qualifier = NameSpace.qualifier (name_of_thm induct); | 
| 277 | val inducts = PureThy.get_thms thy (Name | |
| 278 | (NameSpace.qualified qualifier "inducts")); | |
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 279 | val iTs = term_tvars (prop_of (hd intrs)); | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 280 | val ar = length vs + length iTs; | 
| 22790 
e1cff9268177
Moved functions infer_intro_vars, arities_of, params_of, and
 berghofe parents: 
22606diff
changeset | 281 | val params = InductivePackage.params_of raw_induct; | 
| 
e1cff9268177
Moved functions infer_intro_vars, arities_of, params_of, and
 berghofe parents: 
22606diff
changeset | 282 | val arities = InductivePackage.arities_of raw_induct; | 
| 22271 | 283 | val nparms = length params; | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 284 | val params' = map dest_Var params; | 
| 22790 
e1cff9268177
Moved functions infer_intro_vars, arities_of, params_of, and
 berghofe parents: 
22606diff
changeset | 285 | val rss = InductivePackage.partition_rules raw_induct intrs; | 
| 22271 | 286 | val rss' = map (fn (((s, rs), (_, arity)), elim) => | 
| 22790 
e1cff9268177
Moved functions infer_intro_vars, arities_of, params_of, and
 berghofe parents: 
22606diff
changeset | 287 | (s, (InductivePackage.infer_intro_vars elim arity rs ~~ rs))) | 
| 
e1cff9268177
Moved functions infer_intro_vars, arities_of, params_of, and
 berghofe parents: 
22606diff
changeset | 288 | (rss ~~ arities ~~ elims); | 
| 21858 
05f57309170c
avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
 wenzelm parents: 
21646diff
changeset | 289 | val (prfx, _) = split_last (NameSpace.explode (fst (hd rss))); | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 290 | val tnames = map (fn s => space_implode "_" (s ^ "T" :: vs)) rsets; | 
| 16123 | 291 | |
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 292 | val thy1 = thy |> | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 293 | Theory.root_path |> | 
| 21858 
05f57309170c
avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
 wenzelm parents: 
21646diff
changeset | 294 | Theory.add_path (NameSpace.implode prfx); | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 295 | val (ty_eqs, rlz_eqs) = split_list | 
| 22271 | 296 | (map (fn (s, rs) => mk_realizes_eqn (not (s mem rsets)) vs nparms rs) rss); | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 297 | |
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 298 | val thy1' = thy1 |> | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 299 | Theory.copy |> | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 300 | Theory.add_types (map (fn s => (Sign.base_name s, ar, NoSyn)) tnames) |> | 
| 19510 | 301 | fold (fn s => AxClass.axiomatize_arity_i | 
| 302 | (s, replicate ar HOLogic.typeS, HOLogic.typeS)) tnames |> | |
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 303 | Extraction.add_typeof_eqns_i ty_eqs; | 
| 15570 | 304 | val dts = List.mapPartial (fn (s, rs) => if s mem rsets then | 
| 22271 | 305 | SOME (dt_of_intrs thy1' vs nparms rs) else NONE) rss; | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 306 | |
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 307 | (** datatype representing computational content of inductive set **) | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 308 | |
| 18314 | 309 | val ((dummies, dt_info), thy2) = | 
| 18008 | 310 | thy1 | 
| 18314 | 311 | |> add_dummies | 
| 312 | (DatatypePackage.add_datatype_i false false (map #2 dts)) | |
| 313 | (map (pair false) dts) [] | |
| 314 | ||> Extraction.add_typeof_eqns_i ty_eqs | |
| 315 | ||> Extraction.add_realizes_eqns_i rlz_eqs; | |
| 316 | fun get f = (these oo Option.map) f; | |
| 19046 
bc5c6c9b114e
removed distinct, renamed gen_distinct to distinct;
 wenzelm parents: 
18929diff
changeset | 317 | val rec_names = distinct (op =) (map (fst o dest_Const o head_of o fst o | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 318 | HOLogic.dest_eq o HOLogic.dest_Trueprop o prop_of) (get #rec_thms dt_info)); | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 319 | val (_, constrss) = foldl_map (fn ((recs, dummies), (s, rs)) => | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 320 | if s mem rsets then | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 321 | let | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 322 | val (d :: dummies') = dummies; | 
| 19473 | 323 | val (recs1, recs2) = chop (length rs) (if d then tl recs else recs) | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 324 | in ((recs2, dummies'), map (head_of o hd o rev o snd o strip_comb o | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 325 | fst o HOLogic.dest_eq o HOLogic.dest_Trueprop o prop_of) recs1) | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 326 | end | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 327 | else ((recs, dummies), replicate (length rs) Extraction.nullt)) | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 328 | ((get #rec_thms dt_info, dummies), rss); | 
| 18929 | 329 | val rintrs = map (fn (intr, c) => Envir.eta_contract | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 330 | (Extraction.realizes_of thy2 vs | 
| 22271 | 331 | (if c = Extraction.nullt then c else list_comb (c, map Var (rev | 
| 332 | (Term.add_vars (prop_of intr) []) \\ params'))) (prop_of intr))) | |
| 333 | (maps snd rss ~~ List.concat constrss); | |
| 334 | val (rlzpreds, rlzpreds') = split_list | |
| 335 | (distinct (op = o pairself (#1 o #1)) (map (fn rintr => | |
| 336 | let | |
| 337 | val Const (s, T) = head_of (HOLogic.dest_Trueprop | |
| 338 | (Logic.strip_assums_concl rintr)); | |
| 339 | val s' = Sign.base_name s; | |
| 340 | val T' = Logic.unvarifyT T | |
| 341 | in ((s', SOME T', NoSyn), | |
| 342 | (Const (s, T'), Free (s', T'))) | |
| 343 | end) rintrs)); | |
| 344 | val rlzparams = map (fn Var ((s, _), T) => (s, SOME (Logic.unvarifyT T))) | |
| 345 | (List.take (snd (strip_comb | |
| 346 | (HOLogic.dest_Trueprop (Logic.strip_assums_concl (hd rintrs)))), nparms)); | |
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 347 | |
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 348 | (** realizability predicate **) | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 349 | |
| 22271 | 350 | val (ind_info, thy3') = thy2 |> | 
| 22606 
962f824c2df9
- Tried to make name_of_thm more robust against changes of the
 berghofe parents: 
22596diff
changeset | 351 | InductivePackage.add_inductive_global false "" false false false | 
| 22271 | 352 | rlzpreds rlzparams (map (fn (rintr, intr) => | 
| 353 | ((Sign.base_name (name_of_thm intr), []), | |
| 354 | subst_atomic rlzpreds' (Logic.unvarify rintr))) | |
| 355 | (rintrs ~~ maps snd rss)) [] ||> | |
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 356 | Theory.absolute_path; | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 357 | val thy3 = PureThy.hide_thms false | 
| 22271 | 358 | (map name_of_thm (#intrs ind_info)) thy3'; | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 359 | |
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 360 | (** realizer for induction rule **) | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 361 | |
| 22271 | 362 | val Ps = List.mapPartial (fn _ $ M $ P => if pred_of M mem rsets then | 
| 15531 | 363 | SOME (fst (fst (dest_Var (head_of P)))) else NONE) | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 364 | (HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of raw_induct))); | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 365 | |
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 366 | fun add_ind_realizer (thy, Ps) = | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 367 | let | 
| 22271 | 368 | val rs = indrule_realizer thy induct raw_induct rsets params' | 
| 369 | (vs @ Ps) rec_names rss' intrs dummies; | |
| 370 | val rlzs = map (fn (r, ind) => Extraction.realizes_of thy (vs @ Ps) r | |
| 371 | (prop_of ind)) (rs ~~ inducts); | |
| 372 | val used = foldr add_term_free_names [] rlzs; | |
| 373 | val rnames = Name.variant_list used (replicate (length inducts) "r"); | |
| 374 | val rnames' = Name.variant_list | |
| 375 | (used @ rnames) (replicate (length intrs) "s"); | |
| 376 | val rlzs' as (prems, _, _) :: _ = map (fn (rlz, name) => | |
| 377 | let | |
| 378 | val (P, Q) = strip_one name (Logic.unvarify rlz); | |
| 379 | val Q' = strip_all' [] rnames' Q | |
| 380 | in | |
| 381 | (Logic.strip_imp_prems Q', P, Logic.strip_imp_concl Q') | |
| 382 | end) (rlzs ~~ rnames); | |
| 383 | val concl = HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj (map | |
| 384 | (fn (_, _ $ P, _ $ Q) => HOLogic.mk_imp (P, Q)) rlzs')); | |
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 385 | val rews = map mk_meta_eq | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 386 | (fst_conv :: snd_conv :: get #rec_thms dt_info); | 
| 22271 | 387 | val thm = Goal.prove_global thy [] prems concl (fn prems => EVERY | 
| 388 | [rtac (#raw_induct ind_info) 1, | |
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 389 | rewrite_goals_tac rews, | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 390 | REPEAT ((resolve_tac prems THEN_ALL_NEW EVERY' | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 391 | [K (rewrite_goals_tac rews), ObjectLogic.atomize_tac, | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 392 | DEPTH_SOLVE_1 o FIRST' [atac, etac allE, etac impE]]) 1)]); | 
| 18358 | 393 | val (thm', thy') = PureThy.store_thm ((space_implode "_" | 
| 22271 | 394 | (NameSpace.qualified qualifier "induct" :: vs @ Ps @ | 
| 395 | ["correctness"]), thm), []) thy; | |
| 396 | val thms = map (fn th => zero_var_indexes (rotate_prems ~1 (th RS mp))) | |
| 397 | (DatatypeAux.split_conj_thm thm'); | |
| 398 | val ([thms'], thy'') = PureThy.add_thmss | |
| 399 | [((space_implode "_" | |
| 400 | (NameSpace.qualified qualifier "inducts" :: vs @ Ps @ | |
| 401 | ["correctness"]), thms), [])] thy'; | |
| 402 | val realizers = inducts ~~ thms' ~~ rlzs ~~ rs; | |
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 403 | in | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 404 | Extraction.add_realizers_i | 
| 22271 | 405 | (map (fn (((ind, corr), rlz), r) => | 
| 406 | mk_realizer thy' (vs @ Ps) (Thm.get_name ind, ind, corr, rlz, r)) | |
| 407 | realizers @ (case realizers of | |
| 408 | [(((ind, corr), rlz), r)] => | |
| 409 | [mk_realizer thy' (vs @ Ps) (NameSpace.qualified qualifier "induct", | |
| 410 | ind, corr, rlz, r)] | |
| 411 | | _ => [])) thy'' | |
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 412 | end; | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 413 | |
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 414 | (** realizer for elimination rules **) | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 415 | |
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 416 | val case_names = map (fst o dest_Const o head_of o fst o HOLogic.dest_eq o | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 417 | HOLogic.dest_Trueprop o prop_of o hd) (get #case_thms dt_info); | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 418 | |
| 13921 
69c627b6b28d
Fixed problem in add_elim_realizer which caused bound variables to
 berghofe parents: 
13725diff
changeset | 419 | fun add_elim_realizer Ps | 
| 
69c627b6b28d
Fixed problem in add_elim_realizer which caused bound variables to
 berghofe parents: 
13725diff
changeset | 420 | (((((elim, elimR), intrs), case_thms), case_name), dummy) thy = | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 421 | let | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 422 | val (prem :: prems) = prems_of elim; | 
| 22271 | 423 | fun reorder1 (p, (_, intr)) = | 
| 15570 | 424 | Library.foldl (fn (t, ((s, _), T)) => all T $ lambda (Free (s, T)) t) | 
| 16861 | 425 | (strip_all p, Term.add_vars (prop_of intr) [] \\ params'); | 
| 22271 | 426 | fun reorder2 ((ivs, intr), i) = | 
| 427 | let val fs = Term.add_vars (prop_of intr) [] \\ params' | |
| 15570 | 428 | in Library.foldl (fn (t, x) => lambda (Var x) t) | 
| 22271 | 429 | (list_comb (Bound (i + length ivs), ivs), fs) | 
| 13921 
69c627b6b28d
Fixed problem in add_elim_realizer which caused bound variables to
 berghofe parents: 
13725diff
changeset | 430 | end; | 
| 
69c627b6b28d
Fixed problem in add_elim_realizer which caused bound variables to
 berghofe parents: 
13725diff
changeset | 431 | val p = Logic.list_implies | 
| 
69c627b6b28d
Fixed problem in add_elim_realizer which caused bound variables to
 berghofe parents: 
13725diff
changeset | 432 | (map reorder1 (prems ~~ intrs) @ [prem], concl_of elim); | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 433 | val T' = Extraction.etype_of thy (vs @ Ps) [] p; | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 434 | val T = if dummy then (HOLogic.unitT --> body_type T') --> T' else T'; | 
| 13921 
69c627b6b28d
Fixed problem in add_elim_realizer which caused bound variables to
 berghofe parents: 
13725diff
changeset | 435 | val Ts = map (Extraction.etype_of thy (vs @ Ps) []) (prems_of elim); | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 436 | val r = if null Ps then Extraction.nullt | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 437 | else list_abs (map (pair "x") Ts, list_comb (Const (case_name, T), | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 438 | (if dummy then | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 439 |                [Abs ("x", HOLogic.unitT, Const ("arbitrary", body_type T))]
 | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 440 | else []) @ | 
| 13921 
69c627b6b28d
Fixed problem in add_elim_realizer which caused bound variables to
 berghofe parents: 
13725diff
changeset | 441 | map reorder2 (intrs ~~ (length prems - 1 downto 0)) @ | 
| 
69c627b6b28d
Fixed problem in add_elim_realizer which caused bound variables to
 berghofe parents: 
13725diff
changeset | 442 | [Bound (length prems)])); | 
| 22271 | 443 | val rlz = Extraction.realizes_of thy (vs @ Ps) r (prop_of elim); | 
| 444 | val rlz' = strip_all (Logic.unvarify rlz); | |
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 445 | val rews = map mk_meta_eq case_thms; | 
| 22271 | 446 | val thm = Goal.prove_global thy [] | 
| 447 | (Logic.strip_imp_prems rlz') (Logic.strip_imp_concl rlz') (fn prems => EVERY | |
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 448 | [cut_facts_tac [hd prems] 1, | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 449 | etac elimR 1, | 
| 22271 | 450 | ALLGOALS (asm_simp_tac HOL_basic_ss), | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 451 | rewrite_goals_tac rews, | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 452 | REPEAT ((resolve_tac prems THEN_ALL_NEW (ObjectLogic.atomize_tac THEN' | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 453 | DEPTH_SOLVE_1 o FIRST' [atac, etac allE, etac impE])) 1)]); | 
| 18358 | 454 | val (thm', thy') = PureThy.store_thm ((space_implode "_" | 
| 22271 | 455 | (name_of_thm elim :: vs @ Ps @ ["correctness"]), thm), []) thy | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 456 | in | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 457 | Extraction.add_realizers_i | 
| 22271 | 458 | [mk_realizer thy' (vs @ Ps) (name_of_thm elim, elim, thm', rlz, r)] thy' | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 459 | end; | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 460 | |
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 461 | (** add realizers to theory **) | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 462 | |
| 15570 | 463 | val thy4 = Library.foldl add_ind_realizer (thy3, subsets Ps); | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 464 | val thy5 = Extraction.add_realizers_i | 
| 22271 | 465 | (map (mk_realizer thy4 vs) (map (fn (((rule, rrule), rlz), c) => | 
| 466 | (name_of_thm rule, rule, rrule, rlz, | |
| 467 | list_comb (c, map Var (rev (Term.add_vars (prop_of rule) []) \\ params')))) | |
| 468 | (List.concat (map snd rss) ~~ #intrs ind_info ~~ rintrs ~~ | |
| 469 | List.concat constrss))) thy4; | |
| 470 | val elimps = List.mapPartial (fn ((s, intrs), p) => | |
| 471 | if s mem rsets then SOME (p, intrs) else NONE) | |
| 472 | (rss' ~~ (elims ~~ #elims ind_info)); | |
| 15570 | 473 | val thy6 = Library.foldl (fn (thy, p as (((((elim, _), _), _), _), _)) => thy |> | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 474 | add_elim_realizer [] p |> add_elim_realizer [fst (fst (dest_Var | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 475 | (HOLogic.dest_Trueprop (concl_of elim))))] p) (thy5, | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 476 | elimps ~~ get #case_thms dt_info ~~ case_names ~~ dummies) | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 477 | |
| 16123 | 478 | in Theory.restore_naming thy thy6 end; | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 479 | |
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 480 | fun add_ind_realizers name rsets thy = | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 481 | let | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 482 |     val (_, {intrs, induct, raw_induct, elims, ...}) =
 | 
| 22271 | 483 | InductivePackage.the_inductive (ProofContext.init thy) name; | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 484 | val vss = sort (int_ord o pairself length) | 
| 22271 | 485 | (subsets (map fst (relevant_vars (concl_of (hd intrs))))) | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 486 | in | 
| 15570 | 487 | Library.foldl (add_ind_realizer rsets intrs induct raw_induct elims) (thy, vss) | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 488 | end | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 489 | |
| 20897 | 490 | fun rlz_attrib arg = Thm.declaration_attribute (fn thm => Context.mapping | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 491 | let | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 492 | fun err () = error "ind_realizer: bad rule"; | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 493 | val sets = | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 494 | (case HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of thm)) of | 
| 22271 | 495 | [_] => [pred_of (HOLogic.dest_Trueprop (hd (prems_of thm)))] | 
| 496 | | xs => map (pred_of o fst o HOLogic.dest_imp) xs) | |
| 15570 | 497 | handle TERM _ => err () | Empty => err (); | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 498 | in | 
| 18728 | 499 | add_ind_realizers (hd sets) | 
| 500 | (case arg of | |
| 15531 | 501 | NONE => sets | SOME NONE => [] | 
| 15703 | 502 | | SOME (SOME sets') => sets \\ sets') | 
| 20897 | 503 | end I); | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 504 | |
| 18728 | 505 | val ind_realizer = Attrib.syntax | 
| 15703 | 506 | ((Scan.option (Scan.lift (Args.$$$ "irrelevant") |-- | 
| 507 | Scan.option (Scan.lift (Args.colon) |-- | |
| 18728 | 508 | Scan.repeat1 Args.const))) >> rlz_attrib); | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 509 | |
| 18708 | 510 | val setup = | 
| 18728 | 511 |   Attrib.add_attributes [("ind_realizer", ind_realizer, "add realizers for inductive set")];
 | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 512 | |
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 513 | end; | 
| 15706 | 514 |