| author | wenzelm | 
| Wed, 26 Nov 2014 20:05:34 +0100 | |
| changeset 59058 | a78612c67ec0 | 
| parent 58963 | 26bf09b95dda | 
| child 59498 | 50b60f501b05 | 
| 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 | Author: Stefan Berghofer, TU Muenchen | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 3 | |
| 36043 | 4 | Program extraction from proofs involving inductive predicates: | 
| 29265 
5b4247055bd7
moved old add_term_vars, add_term_frees etc. to structure OldTerm;
 wenzelm parents: 
28965diff
changeset | 5 | Realizers for induction and elimination rules. | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 6 | *) | 
| 
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 | signature INDUCTIVE_REALIZER = | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 9 | sig | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 10 | val add_ind_realizers: string -> string list -> theory -> theory | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 11 | end; | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 12 | |
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 13 | structure InductiveRealizer : INDUCTIVE_REALIZER = | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 14 | struct | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 15 | |
| 22606 
962f824c2df9
- Tried to make name_of_thm more robust against changes of the
 berghofe parents: 
22596diff
changeset | 16 | fun name_of_thm thm = | 
| 28800 | 17 | (case Proofterm.fold_proof_atoms false (fn PThm (_, ((name, _, _), _)) => cons name | _ => I) | 
| 28814 | 18 | [Thm.proof_of thm] [] of | 
| 28800 | 19 | [name] => name | 
| 55235 | 20 |   | _ => raise THM ("name_of_thm: bad proof of theorem", 0, [thm]));
 | 
| 22271 | 21 | |
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 22 | fun prf_of thm = | 
| 44060 | 23 | Reconstruct.proof_of thm | 
| 24 |   |> Reconstruct.expand_proof (Thm.theory_of_thm thm) [("", NONE)];  (* FIXME *)
 | |
| 13710 
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 subsets [] = [[]] | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 27 | | subsets (x::xs) = | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 28 | let val ys = subsets xs | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 29 | in ys @ map (cons x) ys end; | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 30 | |
| 22271 | 31 | 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 | 32 | |
| 56245 | 33 | fun strip_all' used names (Const (@{const_name Pure.all}, _) $ Abs (s, T, t)) =
 | 
| 22271 | 34 | let val (s', names') = (case names of | 
| 43324 
2b47822868e4
discontinued Name.variant to emphasize that this is old-style / indirect;
 wenzelm parents: 
42375diff
changeset | 35 | [] => (singleton (Name.variant_list used) s, []) | 
| 22271 | 36 | | name :: names' => (name, names')) | 
| 37 | in strip_all' (s'::used) names' (subst_bound (Free (s', T), t)) end | |
| 56245 | 38 |   | strip_all' used names ((t as Const (@{const_name Pure.imp}, _) $ P) $ Q) =
 | 
| 22271 | 39 | t $ strip_all' used names Q | 
| 40 | | strip_all' _ _ t = t; | |
| 41 | ||
| 29281 | 42 | fun strip_all t = strip_all' (Term.add_free_names t []) [] t; | 
| 22271 | 43 | |
| 56245 | 44 | fun strip_one name | 
| 45 |     (Const (@{const_name Pure.all}, _) $ Abs (s, T, Const (@{const_name Pure.imp}, _) $ P $ Q)) =
 | |
| 22271 | 46 | (subst_bound (Free (name, T), P), subst_bound (Free (name, T), Q)) | 
| 56245 | 47 |   | strip_one _ (Const (@{const_name Pure.imp}, _) $ P $ Q) = (P, Q);
 | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 48 | |
| 37233 
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
 berghofe parents: 
36945diff
changeset | 49 | fun relevant_vars prop = fold (fn ((a, i), T) => fn vs => | 
| 
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
 berghofe parents: 
36945diff
changeset | 50 | (case strip_type T of | 
| 36692 
54b64d4ad524
farewell to old-style mem infixes -- type inference in situations with mem_int and mem_string should provide enough information to resolve the type of (op =)
 haftmann parents: 
36610diff
changeset | 51 |         (_, Type (s, _)) => if s = @{type_name bool} then (a, T) :: vs else vs
 | 
| 37233 
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
 berghofe parents: 
36945diff
changeset | 52 | | _ => vs)) (Term.add_vars prop []) []; | 
| 
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
 berghofe parents: 
36945diff
changeset | 53 | |
| 
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
 berghofe parents: 
36945diff
changeset | 54 | val attach_typeS = map_types (map_atyps | 
| 56254 | 55 |   (fn TFree (s, []) => TFree (s, @{sort type})
 | 
| 56 |     | TVar (ixn, []) => TVar (ixn, @{sort type})
 | |
| 37233 
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
 berghofe parents: 
36945diff
changeset | 57 | | T => T)); | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 58 | |
| 22271 | 59 | fun dt_of_intrs thy vs nparms intrs = | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 60 | let | 
| 37233 
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
 berghofe parents: 
36945diff
changeset | 61 | val iTs = rev (Term.add_tvars (prop_of (hd intrs)) []); | 
| 22271 | 62 | val (Const (s, _), ts) = strip_comb (HOLogic.dest_Trueprop | 
| 63 | (Logic.strip_imp_concl (prop_of (hd intrs)))); | |
| 33957 | 64 | val params = map dest_Var (take nparms ts); | 
| 30364 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 wenzelm parents: 
30345diff
changeset | 65 | val tname = Binding.name (space_implode "_" (Long_Name.base_name s ^ "T" :: vs)); | 
| 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 wenzelm parents: 
30345diff
changeset | 66 | fun constr_of_intr intr = (Binding.name (Long_Name.base_name (name_of_thm intr)), | 
| 35845 
e5980f0ad025
renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
 wenzelm parents: 
35625diff
changeset | 67 | map (Logic.unvarifyT_global o snd) (subtract (op =) params (rev (Term.add_vars (prop_of intr) []))) @ | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 68 | filter_out (equal Extraction.nullT) (map | 
| 35845 
e5980f0ad025
renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
 wenzelm parents: 
35625diff
changeset | 69 | (Logic.unvarifyT_global o Extraction.etype_of thy vs []) (prems_of intr)), | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 70 | NoSyn); | 
| 45839 
43a5b86bc102
'datatype' specifications allow explicit sort constraints;
 wenzelm parents: 
45701diff
changeset | 71 | in | 
| 
43a5b86bc102
'datatype' specifications allow explicit sort constraints;
 wenzelm parents: 
45701diff
changeset | 72 | ((tname, map (rpair dummyS) (map (fn a => "'" ^ a) vs @ map (fst o fst) iTs), NoSyn), | 
| 
43a5b86bc102
'datatype' specifications allow explicit sort constraints;
 wenzelm parents: 
45701diff
changeset | 73 | map constr_of_intr intrs) | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 74 | end; | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 75 | |
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 76 | 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 | 77 | |
| 22271 | 78 | (** turn "P" into "%r x. realizes r (P x)" **) | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 79 | |
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 80 | fun gen_rvar vs (t as Var ((a, 0), T)) = | 
| 22271 | 81 | if body_type T <> HOLogic.boolT then t else | 
| 82 | let | |
| 37233 
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
 berghofe parents: 
36945diff
changeset | 83 |           val U = TVar (("'" ^ a, 0), [])
 | 
| 22271 | 84 | val Ts = binder_types T; | 
| 85 | val i = length Ts; | |
| 86 | val xs = map (pair "x") Ts; | |
| 87 | val u = list_comb (t, map Bound (i - 1 downto 0)) | |
| 88 | in | |
| 36692 
54b64d4ad524
farewell to old-style mem infixes -- type inference in situations with mem_int and mem_string should provide enough information to resolve the type of (op =)
 haftmann parents: 
36610diff
changeset | 89 | if member (op =) vs a then | 
| 46219 
426ed18eba43
discontinued old-style Term.list_abs in favour of plain Term.abs;
 wenzelm parents: 
45839diff
changeset | 90 |             fold_rev Term.abs (("r", U) :: xs) (mk_rlz U $ Bound i $ u)
 | 
| 
426ed18eba43
discontinued old-style Term.list_abs in favour of plain Term.abs;
 wenzelm parents: 
45839diff
changeset | 91 | else | 
| 
426ed18eba43
discontinued old-style Term.list_abs in favour of plain Term.abs;
 wenzelm parents: 
45839diff
changeset | 92 | fold_rev Term.abs xs (mk_rlz Extraction.nullT $ Extraction.nullt $ u) | 
| 22271 | 93 | end | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 94 | | gen_rvar _ t = t; | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 95 | |
| 22271 | 96 | fun mk_realizes_eqn n vs nparms intrs = | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 97 | let | 
| 37233 
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
 berghofe parents: 
36945diff
changeset | 98 | val intr = map_types Type.strip_sorts (prop_of (hd intrs)); | 
| 
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
 berghofe parents: 
36945diff
changeset | 99 | val concl = HOLogic.dest_Trueprop (Logic.strip_imp_concl intr); | 
| 
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
 berghofe parents: 
36945diff
changeset | 100 | val iTs = rev (Term.add_tvars intr []); | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 101 | val Tvs = map TVar iTs; | 
| 22271 | 102 | val (h as Const (s, T), us) = strip_comb concl; | 
| 103 | val params = List.take (us, nparms); | |
| 104 | val elTs = List.drop (binder_types T, nparms); | |
| 105 | val predT = elTs ---> HOLogic.boolT; | |
| 106 | val used = map (fst o fst o dest_Var) params; | |
| 107 | val xs = map (Var o apfst (rpair 0)) | |
| 108 | (Name.variant_list used (replicate (length elTs) "x") ~~ elTs); | |
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 109 | val rT = if n then Extraction.nullT | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 110 | else Type (space_implode "_" (s ^ "T" :: vs), | 
| 37233 
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
 berghofe parents: 
36945diff
changeset | 111 |         map (fn a => TVar (("'" ^ a, 0), [])) vs @ Tvs);
 | 
| 30364 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 wenzelm parents: 
30345diff
changeset | 112 | val r = if n then Extraction.nullt else Var ((Long_Name.base_name s, 0), rT); | 
| 22271 | 113 | val S = list_comb (h, params @ xs); | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 114 | val rvs = relevant_vars S; | 
| 33040 | 115 | val vs' = subtract (op =) vs (map fst rvs); | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 116 | val rname = space_implode "_" (s ^ "R" :: vs); | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 117 | |
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 118 | fun mk_Tprem n v = | 
| 17485 | 119 | 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 | 120 |       in (Const ("typeof", T --> Type ("Type", [])) $ Var ((v, 0), T),
 | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 121 | Extraction.mk_typ (if n then Extraction.nullT | 
| 37233 
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
 berghofe parents: 
36945diff
changeset | 122 |           else TVar (("'" ^ v, 0), [])))
 | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 123 | end; | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 124 | |
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 125 | val prems = map (mk_Tprem true) vs' @ map (mk_Tprem false) vs; | 
| 22271 | 126 | val ts = map (gen_rvar vs) params; | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 127 | val argTs = map fastype_of ts; | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 128 | |
| 22271 | 129 |   in ((prems, (Const ("typeof", HOLogic.boolT --> Type ("Type", [])) $ S,
 | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 130 | Extraction.mk_typ rT)), | 
| 22271 | 131 | (prems, (mk_rlz rT $ r $ S, | 
| 132 | if n then list_comb (Const (rname, argTs ---> predT), ts @ xs) | |
| 133 | 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 | 134 | end; | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 135 | |
| 22271 | 136 | 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 | 137 | let | 
| 42361 | 138 | val ctxt = Proof_Context.init_global thy | 
| 22271 | 139 | 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 | 140 | val args' = map (Free o apfst fst) | 
| 33040 | 141 | (subtract (op =) params (Term.add_vars (prop_of intr) [])); | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 142 | val rule' = strip_all rule; | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 143 | 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 | 144 | val used = map (fst o dest_Free) args; | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 145 | |
| 29271 
1d685baea08e
moved old add_type_XXX, add_term_XXX etc. to structure OldTerm;
 wenzelm parents: 
29265diff
changeset | 146 | val is_rec = exists_Const (fn (c, _) => member (op =) rsets c); | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 147 | |
| 56245 | 148 |     fun is_meta (Const (@{const_name Pure.all}, _) $ Abs (s, _, P)) = is_meta P
 | 
| 149 |       | is_meta (Const (@{const_name Pure.imp}, _) $ _ $ Q) = is_meta Q
 | |
| 35364 | 150 |       | is_meta (Const (@{const_name Trueprop}, _) $ t) =
 | 
| 151 | (case head_of t of | |
| 152 | Const (s, _) => can (Inductive.the_inductive ctxt) s | |
| 153 | | _ => true) | |
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 154 | | is_meta _ = false; | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 155 | |
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 156 | fun fun_of ts rts args used (prem :: prems) = | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 157 | let | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 158 | val T = Extraction.etype_of thy vs [] prem; | 
| 20071 
8f3e1ddb50e6
replaced Term.variant(list) by Name.variant(_list);
 wenzelm parents: 
19806diff
changeset | 159 | val [x, r] = Name.variant_list used ["x", "r"] | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 160 | in if T = Extraction.nullT | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 161 | then fun_of ts rts args used prems | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 162 | else if is_rec prem then | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 163 | if is_meta prem then | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 164 | let | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 165 | val prem' :: prems' = prems; | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 166 | val U = Extraction.etype_of thy vs [] prem'; | 
| 46219 
426ed18eba43
discontinued old-style Term.list_abs in favour of plain Term.abs;
 wenzelm parents: 
45839diff
changeset | 167 | in | 
| 
426ed18eba43
discontinued old-style Term.list_abs in favour of plain Term.abs;
 wenzelm parents: 
45839diff
changeset | 168 | if U = Extraction.nullT | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 169 | then fun_of (Free (x, T) :: ts) | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 170 | (Free (r, binder_types T ---> HOLogic.unitT) :: rts) | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 171 | (Free (x, T) :: args) (x :: r :: used) prems' | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 172 | 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 | 173 | (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 | 174 | end | 
| 46219 
426ed18eba43
discontinued old-style Term.list_abs in favour of plain Term.abs;
 wenzelm parents: 
45839diff
changeset | 175 | else | 
| 
426ed18eba43
discontinued old-style Term.list_abs in favour of plain Term.abs;
 wenzelm parents: 
45839diff
changeset | 176 | (case strip_type T of | 
| 37678 
0040bafffdef
"prod" and "sum" replace "*" and "+" respectively
 haftmann parents: 
37236diff
changeset | 177 |                   (Ts, Type (@{type_name Product_Type.prod}, [T1, T2])) =>
 | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 178 | let | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 179 | val fx = Free (x, Ts ---> T1); | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 180 | val fr = Free (r, Ts ---> T2); | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 181 | val bs = map Bound (length Ts - 1 downto 0); | 
| 46219 
426ed18eba43
discontinued old-style Term.list_abs in favour of plain Term.abs;
 wenzelm parents: 
45839diff
changeset | 182 | val t = | 
| 
426ed18eba43
discontinued old-style Term.list_abs in favour of plain Term.abs;
 wenzelm parents: 
45839diff
changeset | 183 | fold_rev (Term.abs o pair "z") Ts | 
| 
426ed18eba43
discontinued old-style Term.list_abs in favour of plain Term.abs;
 wenzelm parents: 
45839diff
changeset | 184 | (HOLogic.mk_prod (list_comb (fx, bs), list_comb (fr, bs))); | 
| 
426ed18eba43
discontinued old-style Term.list_abs in favour of plain Term.abs;
 wenzelm parents: 
45839diff
changeset | 185 | in fun_of (fx :: ts) (fr :: rts) (t::args) (x :: r :: used) prems end | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 186 | | (Ts, U) => fun_of (Free (x, T) :: ts) | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 187 | (Free (r, binder_types T ---> HOLogic.unitT) :: rts) | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 188 | (Free (x, T) :: args) (x :: r :: used) prems) | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 189 | 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 | 190 | (x :: used) prems | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 191 | end | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 192 | | fun_of ts rts args used [] = | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 193 | let val xs = rev (rts @ ts) | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 194 | in if conclT = Extraction.nullT | 
| 44241 | 195 | then fold_rev (absfree o dest_Free) xs HOLogic.unit | 
| 196 | else fold_rev (absfree o dest_Free) xs | |
| 197 | (list_comb | |
| 198 |                 (Free ("r" ^ Long_Name.base_name (name_of_thm intr),
 | |
| 199 | map fastype_of (rev args) ---> conclT), rev args)) | |
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 200 | end | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 201 | |
| 13921 
69c627b6b28d
Fixed problem in add_elim_realizer which caused bound variables to
 berghofe parents: 
13725diff
changeset | 202 | 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 | 203 | |
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 204 | 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 | 205 | let | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 206 | val concls = HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of raw_induct)); | 
| 31986 | 207 | val premss = map_filter (fn (s, rs) => if member (op =) rsets s then | 
| 208 | SOME (rs, map (fn (_, r) => nth (prems_of raw_induct) | |
| 209 | (find_index (fn prp => prp = prop_of r) (map prop_of intrs))) rs) else NONE) rss; | |
| 22271 | 210 | val fs = maps (fn ((intrs, prems), dummy) => | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 211 | let | 
| 22271 | 212 | val fs = map (fn (rule, (ivs, intr)) => | 
| 213 | fun_of_prem thy rsets vs params rule ivs intr) (prems ~~ intrs) | |
| 35364 | 214 | in | 
| 215 |         if dummy then Const (@{const_name default},
 | |
| 216 | HOLogic.unitT --> body_type (fastype_of (hd fs))) :: fs | |
| 22271 | 217 | else fs | 
| 218 | end) (premss ~~ dummies); | |
| 16861 | 219 | val frees = fold Term.add_frees fs []; | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 220 | val Ts = map fastype_of fs; | 
| 30364 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 wenzelm parents: 
30345diff
changeset | 221 | fun name_of_fn intr = "r" ^ Long_Name.base_name (name_of_thm intr) | 
| 22271 | 222 | in | 
| 223 | fst (fold_map (fn concl => fn names => | |
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 224 | let val T = Extraction.etype_of thy vs [] concl | 
| 22271 | 225 | in if T = Extraction.nullT then (Extraction.nullt, names) else | 
| 226 | let | |
| 227 |           val Type ("fun", [U, _]) = T;
 | |
| 228 | val a :: names' = names | |
| 44241 | 229 | in | 
| 230 |           (fold_rev absfree (("x", U) :: map_filter (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')
 | |
| 22271 | 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 | |
| 45839 
43a5b86bc102
'datatype' specifications allow explicit sort constraints;
 wenzelm parents: 
45701diff
changeset | 238 | fun add_dummy name dname (x as (_, ((s, vs, mx), cs))) = | 
| 
43a5b86bc102
'datatype' specifications allow explicit sort constraints;
 wenzelm parents: 
45701diff
changeset | 239 | if Binding.eq_name (name, s) | 
| 
43a5b86bc102
'datatype' specifications allow explicit sort constraints;
 wenzelm parents: 
45701diff
changeset | 240 | then (true, ((s, vs, mx), (dname, [HOLogic.unitT], NoSyn) :: cs)) | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 241 | else x; | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 242 | |
| 18314 | 243 | fun add_dummies f [] _ thy = | 
| 244 | (([], NONE), thy) | |
| 245 | | add_dummies f dts used thy = | |
| 246 | thy | |
| 247 | |> f (map snd dts) | |
| 30345 | 248 | |-> (fn dtinfo => pair (map fst dts, SOME dtinfo)) | 
| 58112 
8081087096ad
renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
 blanchet parents: 
58111diff
changeset | 249 | handle Old_Datatype_Aux.Datatype_Empty name' => | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 250 | let | 
| 30364 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 wenzelm parents: 
30345diff
changeset | 251 | val name = Long_Name.base_name name'; | 
| 43324 
2b47822868e4
discontinued Name.variant to emphasize that this is old-style / indirect;
 wenzelm parents: 
42375diff
changeset | 252 | val dname = singleton (Name.variant_list used) "Dummy"; | 
| 18314 | 253 | in | 
| 254 | thy | |
| 30345 | 255 | |> add_dummies f (map (add_dummy (Binding.name name) (Binding.name dname)) dts) (dname :: used) | 
| 14888 
99ac3eb0f84e
add_dummies no longer uses transform_error but handles specific
 berghofe parents: 
13928diff
changeset | 256 | end; | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 257 | |
| 22271 | 258 | 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 | 259 | let | 
| 13725 
12404b452034
Changed format of realizers / correctness proofs.
 berghofe parents: 
13710diff
changeset | 260 | val rvs = map fst (relevant_vars (prop_of rule)); | 
| 16861 | 261 | val xs = rev (Term.add_vars (prop_of rule) []); | 
| 36692 
54b64d4ad524
farewell to old-style mem infixes -- type inference in situations with mem_int and mem_string should provide enough information to resolve the type of (op =)
 haftmann parents: 
36610diff
changeset | 262 | val vs1 = map Var (filter_out (fn ((a, _), _) => member (op =) rvs a) xs); | 
| 16861 | 263 | val rlzvs = rev (Term.add_vars (prop_of rrule) []); | 
| 17485 | 264 | val vs2 = map (fn (ixn, _) => Var (ixn, (the o AList.lookup (op =) rlzvs) ixn)) xs; | 
| 59058 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 wenzelm parents: 
58963diff
changeset | 265 | val rs = map Var (subtract (op = o apply2 fst) xs rlzvs); | 
| 37233 
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
 berghofe parents: 
36945diff
changeset | 266 | val rlz' = fold_rev Logic.all rs (prop_of rrule) | 
| 22271 | 267 | in (name, (vs, | 
| 33338 | 268 | if rt = Extraction.nullt then rt else fold_rev lambda vs1 rt, | 
| 37233 
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
 berghofe parents: 
36945diff
changeset | 269 | Extraction.abs_corr_shyps thy rule vs vs2 | 
| 
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
 berghofe parents: 
36945diff
changeset | 270 | (ProofRewriteRules.un_hhf_proof rlz' (attach_typeS rlz) | 
| 
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
 berghofe parents: 
36945diff
changeset | 271 | (fold_rev Proofterm.forall_intr_proof' rs (prf_of rrule))))) | 
| 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 | |
| 24157 
409cd6eaa7ea
Added renaming function to prevent correctness proof for realizer
 berghofe parents: 
23590diff
changeset | 274 | fun rename tab = map (fn x => the_default x (AList.lookup op = tab x)); | 
| 
409cd6eaa7ea
Added renaming function to prevent correctness proof for realizer
 berghofe parents: 
23590diff
changeset | 275 | |
| 33244 | 276 | fun add_ind_realizer rsets intrs induct raw_induct elims vs thy = | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 277 | let | 
| 30364 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 wenzelm parents: 
30345diff
changeset | 278 | val qualifier = Long_Name.qualifier (name_of_thm induct); | 
| 39557 
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
 wenzelm parents: 
37678diff
changeset | 279 | val inducts = Global_Theory.get_thms thy (Long_Name.qualify qualifier "inducts"); | 
| 37233 
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
 berghofe parents: 
36945diff
changeset | 280 | val iTs = rev (Term.add_tvars (prop_of (hd intrs)) []); | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 281 | val ar = length vs + length iTs; | 
| 31723 
f5cafe803b55
discontinued ancient tradition to suffix certain ML module names with "_package"
 haftmann parents: 
31668diff
changeset | 282 | val params = Inductive.params_of raw_induct; | 
| 
f5cafe803b55
discontinued ancient tradition to suffix certain ML module names with "_package"
 haftmann parents: 
31668diff
changeset | 283 | val arities = Inductive.arities_of raw_induct; | 
| 22271 | 284 | val nparms = length params; | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 285 | val params' = map dest_Var params; | 
| 31723 
f5cafe803b55
discontinued ancient tradition to suffix certain ML module names with "_package"
 haftmann parents: 
31668diff
changeset | 286 | val rss = Inductive.partition_rules raw_induct intrs; | 
| 22271 | 287 | val rss' = map (fn (((s, rs), (_, arity)), elim) => | 
| 31723 
f5cafe803b55
discontinued ancient tradition to suffix certain ML module names with "_package"
 haftmann parents: 
31668diff
changeset | 288 | (s, (Inductive.infer_intro_vars elim arity rs ~~ rs))) | 
| 22790 
e1cff9268177
Moved functions infer_intro_vars, arities_of, params_of, and
 berghofe parents: 
22606diff
changeset | 289 | (rss ~~ arities ~~ elims); | 
| 30364 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 wenzelm parents: 
30345diff
changeset | 290 | val (prfx, _) = split_last (Long_Name.explode (fst (hd rss))); | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 291 | val tnames = map (fn s => space_implode "_" (s ^ "T" :: vs)) rsets; | 
| 16123 | 292 | |
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 293 | val thy1 = thy |> | 
| 24712 
64ed05609568
proper Sign operations instead of Theory aliases;
 wenzelm parents: 
24157diff
changeset | 294 | Sign.root_path |> | 
| 30364 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 wenzelm parents: 
30345diff
changeset | 295 | Sign.add_path (Long_Name.implode prfx); | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 296 | val (ty_eqs, rlz_eqs) = split_list | 
| 36692 
54b64d4ad524
farewell to old-style mem infixes -- type inference in situations with mem_int and mem_string should provide enough information to resolve the type of (op =)
 haftmann parents: 
36610diff
changeset | 297 | (map (fn (s, rs) => mk_realizes_eqn (not (member (op =) rsets s)) vs nparms rs) rss); | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 298 | |
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 299 | val thy1' = thy1 |> | 
| 42375 
774df7c59508
report Name_Space.declare/define, relatively to context;
 wenzelm parents: 
42361diff
changeset | 300 | Sign.add_types_global | 
| 
774df7c59508
report Name_Space.declare/define, relatively to context;
 wenzelm parents: 
42361diff
changeset | 301 | (map (fn s => (Binding.name (Long_Name.base_name s), ar, NoSyn)) tnames) |> | 
| 
774df7c59508
report Name_Space.declare/define, relatively to context;
 wenzelm parents: 
42361diff
changeset | 302 | Extraction.add_typeof_eqns_i ty_eqs; | 
| 36692 
54b64d4ad524
farewell to old-style mem infixes -- type inference in situations with mem_int and mem_string should provide enough information to resolve the type of (op =)
 haftmann parents: 
36610diff
changeset | 303 | val dts = map_filter (fn (s, rs) => if member (op =) rsets s then | 
| 22271 | 304 | 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 | 305 | |
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 306 | (** datatype representing computational content of inductive set **) | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 307 | |
| 31783 
cfbe9609ceb1
add_datatypes does not yield particular rules any longer
 haftmann parents: 
31781diff
changeset | 308 | val ((dummies, some_dt_names), thy2) = | 
| 18008 | 309 | thy1 | 
| 58112 
8081087096ad
renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
 blanchet parents: 
58111diff
changeset | 310 |       |> add_dummies (Old_Datatype.add_datatype {strict = false, quiet = false})
 | 
| 45701 | 311 | (map (pair false) dts) [] | 
| 18314 | 312 | ||> Extraction.add_typeof_eqns_i ty_eqs | 
| 313 | ||> Extraction.add_realizes_eqns_i rlz_eqs; | |
| 31783 
cfbe9609ceb1
add_datatypes does not yield particular rules any longer
 haftmann parents: 
31781diff
changeset | 314 | val dt_names = these some_dt_names; | 
| 58112 
8081087096ad
renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
 blanchet parents: 
58111diff
changeset | 315 | val case_thms = map (#case_rewrites o Old_Datatype_Data.the_info thy2) dt_names; | 
| 45701 | 316 | val rec_thms = | 
| 317 | if null dt_names then [] | |
| 58112 
8081087096ad
renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
 blanchet parents: 
58111diff
changeset | 318 | else #rec_rewrites (Old_Datatype_Data.the_info thy2 (hd dt_names)); | 
| 19046 
bc5c6c9b114e
removed distinct, renamed gen_distinct to distinct;
 wenzelm parents: 
18929diff
changeset | 319 | val rec_names = distinct (op =) (map (fst o dest_Const o head_of o fst o | 
| 31781 
861e675f01e6
add_datatype interface yields type names and less rules
 haftmann parents: 
31723diff
changeset | 320 | HOLogic.dest_eq o HOLogic.dest_Trueprop o prop_of) rec_thms); | 
| 31458 | 321 | val (constrss, _) = fold_map (fn (s, rs) => fn (recs, dummies) => | 
| 322 | if member (op =) rsets s then | |
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 323 | let | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 324 | val (d :: dummies') = dummies; | 
| 19473 | 325 | val (recs1, recs2) = chop (length rs) (if d then tl recs else recs) | 
| 31458 | 326 | in (map (head_of o hd o rev o snd o strip_comb o fst o | 
| 327 | HOLogic.dest_eq o HOLogic.dest_Trueprop o prop_of) recs1, (recs2, dummies')) | |
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 328 | end | 
| 31458 | 329 | else (replicate (length rs) Extraction.nullt, (recs, dummies))) | 
| 31781 
861e675f01e6
add_datatype interface yields type names and less rules
 haftmann parents: 
31723diff
changeset | 330 | rss (rec_thms, dummies); | 
| 37233 
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
 berghofe parents: 
36945diff
changeset | 331 | val rintrs = map (fn (intr, c) => attach_typeS (Envir.eta_contract | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 332 | (Extraction.realizes_of thy2 vs | 
| 22271 | 333 | (if c = Extraction.nullt then c else list_comb (c, map Var (rev | 
| 37233 
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
 berghofe parents: 
36945diff
changeset | 334 | (subtract (op =) params' (Term.add_vars (prop_of intr) []))))) (prop_of intr)))) | 
| 32952 | 335 | (maps snd rss ~~ flat constrss); | 
| 30345 | 336 | val (rlzpreds, rlzpreds') = | 
| 337 | rintrs |> map (fn rintr => | |
| 22271 | 338 | let | 
| 30345 | 339 | val Const (s, T) = head_of (HOLogic.dest_Trueprop (Logic.strip_assums_concl rintr)); | 
| 30364 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 wenzelm parents: 
30345diff
changeset | 340 | val s' = Long_Name.base_name s; | 
| 35845 
e5980f0ad025
renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
 wenzelm parents: 
35625diff
changeset | 341 | val T' = Logic.unvarifyT_global T; | 
| 30345 | 342 | in (((s', T'), NoSyn), (Const (s, T'), Free (s', T'))) end) | 
| 59058 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 wenzelm parents: 
58963diff
changeset | 343 | |> distinct (op = o apply2 (#1 o #1)) | 
| 30345 | 344 | |> map (apfst (apfst (apfst Binding.name))) | 
| 345 | |> split_list; | |
| 346 | ||
| 35845 
e5980f0ad025
renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
 wenzelm parents: 
35625diff
changeset | 347 | val rlzparams = map (fn Var ((s, _), T) => (s, Logic.unvarifyT_global T)) | 
| 22271 | 348 | (List.take (snd (strip_comb | 
| 349 | (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 | 350 | |
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 351 | (** realizability predicate **) | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 352 | |
| 22271 | 353 | val (ind_info, thy3') = thy2 |> | 
| 33726 
0878aecbf119
eliminated slightly odd name space grouping -- now managed by Isar toplevel;
 wenzelm parents: 
33671diff
changeset | 354 | Inductive.add_inductive_global | 
| 33669 | 355 |         {quiet_mode = false, verbose = false, alt_name = Binding.empty, coind = false,
 | 
| 49170 
03bee3a6a1b7
discontinued obsolete fork_mono to loosen some brakes -- NB: TTY interaction has Goal.future_proofs disabled due to missing Future.worker_task;
 wenzelm parents: 
47060diff
changeset | 356 | no_elim = false, no_ind = false, skip_mono = false} | 
| 22271 | 357 | rlzpreds rlzparams (map (fn (rintr, intr) => | 
| 30364 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 wenzelm parents: 
30345diff
changeset | 358 | ((Binding.name (Long_Name.base_name (name_of_thm intr)), []), | 
| 35845 
e5980f0ad025
renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
 wenzelm parents: 
35625diff
changeset | 359 | subst_atomic rlzpreds' (Logic.unvarify_global rintr))) | 
| 22271 | 360 | (rintrs ~~ maps snd rss)) [] ||> | 
| 30435 
e62d6ecab6ad
explicit Binding.qualified_name -- prevents implicitly qualified bstring;
 wenzelm parents: 
30364diff
changeset | 361 | Sign.root_path; | 
| 39557 
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
 wenzelm parents: 
37678diff
changeset | 362 | val thy3 = fold (Global_Theory.hide_fact false o name_of_thm) (#intrs ind_info) thy3'; | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 363 | |
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 364 | (** realizer for induction rule **) | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 365 | |
| 36692 
54b64d4ad524
farewell to old-style mem infixes -- type inference in situations with mem_int and mem_string should provide enough information to resolve the type of (op =)
 haftmann parents: 
36610diff
changeset | 366 | val Ps = map_filter (fn _ $ M $ P => if member (op =) rsets (pred_of M) then | 
| 15531 | 367 | 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 | 368 | (HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of raw_induct))); | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 369 | |
| 33244 | 370 | fun add_ind_realizer Ps thy = | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 371 | let | 
| 59058 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 wenzelm parents: 
58963diff
changeset | 372 | val vs' = rename (map (apply2 (fst o fst o dest_Var)) | 
| 24157 
409cd6eaa7ea
Added renaming function to prevent correctness proof for realizer
 berghofe parents: 
23590diff
changeset | 373 | (params ~~ List.take (snd (strip_comb (HOLogic.dest_Trueprop | 
| 
409cd6eaa7ea
Added renaming function to prevent correctness proof for realizer
 berghofe parents: 
23590diff
changeset | 374 | (hd (prems_of (hd inducts))))), nparms))) vs; | 
| 22271 | 375 | val rs = indrule_realizer thy induct raw_induct rsets params' | 
| 24157 
409cd6eaa7ea
Added renaming function to prevent correctness proof for realizer
 berghofe parents: 
23590diff
changeset | 376 | (vs' @ Ps) rec_names rss' intrs dummies; | 
| 
409cd6eaa7ea
Added renaming function to prevent correctness proof for realizer
 berghofe parents: 
23590diff
changeset | 377 | val rlzs = map (fn (r, ind) => Extraction.realizes_of thy (vs' @ Ps) r | 
| 22271 | 378 | (prop_of ind)) (rs ~~ inducts); | 
| 29281 | 379 | val used = fold Term.add_free_names rlzs []; | 
| 22271 | 380 | val rnames = Name.variant_list used (replicate (length inducts) "r"); | 
| 381 | val rnames' = Name.variant_list | |
| 382 | (used @ rnames) (replicate (length intrs) "s"); | |
| 383 | val rlzs' as (prems, _, _) :: _ = map (fn (rlz, name) => | |
| 384 | let | |
| 35845 
e5980f0ad025
renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
 wenzelm parents: 
35625diff
changeset | 385 | val (P, Q) = strip_one name (Logic.unvarify_global rlz); | 
| 22271 | 386 | val Q' = strip_all' [] rnames' Q | 
| 387 | in | |
| 388 | (Logic.strip_imp_prems Q', P, Logic.strip_imp_concl Q') | |
| 389 | end) (rlzs ~~ rnames); | |
| 390 | val concl = HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj (map | |
| 391 | (fn (_, _ $ P, _ $ Q) => HOLogic.mk_imp (P, Q)) rlzs')); | |
| 37136 | 392 |         val rews = map mk_meta_eq (@{thm fst_conv} :: @{thm snd_conv} :: rec_thms);
 | 
| 37233 
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
 berghofe parents: 
36945diff
changeset | 393 | val thm = Goal.prove_global thy [] | 
| 
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
 berghofe parents: 
36945diff
changeset | 394 | (map attach_typeS prems) (attach_typeS concl) | 
| 54742 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
52788diff
changeset | 395 |           (fn {context = ctxt, prems} => EVERY
 | 
| 22271 | 396 | [rtac (#raw_induct ind_info) 1, | 
| 54742 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
52788diff
changeset | 397 | rewrite_goals_tac ctxt rews, | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 398 | REPEAT ((resolve_tac prems THEN_ALL_NEW EVERY' | 
| 54742 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
52788diff
changeset | 399 | [K (rewrite_goals_tac ctxt rews), Object_Logic.atomize_prems_tac ctxt, | 
| 58963 
26bf09b95dda
proper context for assume_tac (atac remains as fall-back without context);
 wenzelm parents: 
58372diff
changeset | 400 | DEPTH_SOLVE_1 o FIRST' [assume_tac ctxt, etac allE, etac impE]]) 1)]); | 
| 39557 
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
 wenzelm parents: 
37678diff
changeset | 401 | val (thm', thy') = Global_Theory.store_thm (Binding.qualified_name (space_implode "_" | 
| 30364 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 wenzelm parents: 
30345diff
changeset | 402 | (Long_Name.qualify qualifier "induct" :: vs' @ Ps @ ["correctness"])), thm) thy; | 
| 22271 | 403 | val thms = map (fn th => zero_var_indexes (rotate_prems ~1 (th RS mp))) | 
| 58112 
8081087096ad
renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
 blanchet parents: 
58111diff
changeset | 404 | (Old_Datatype_Aux.split_conj_thm thm'); | 
| 39557 
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
 wenzelm parents: 
37678diff
changeset | 405 | val ([thms'], thy'') = Global_Theory.add_thmss | 
| 30435 
e62d6ecab6ad
explicit Binding.qualified_name -- prevents implicitly qualified bstring;
 wenzelm parents: 
30364diff
changeset | 406 | [((Binding.qualified_name (space_implode "_" | 
| 30364 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 wenzelm parents: 
30345diff
changeset | 407 | (Long_Name.qualify qualifier "inducts" :: vs' @ Ps @ | 
| 29579 | 408 | ["correctness"])), thms), [])] thy'; | 
| 22271 | 409 | val realizers = inducts ~~ thms' ~~ rlzs ~~ rs; | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 410 | in | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 411 | Extraction.add_realizers_i | 
| 22271 | 412 | (map (fn (((ind, corr), rlz), r) => | 
| 37233 
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
 berghofe parents: 
36945diff
changeset | 413 | mk_realizer thy'' (vs' @ Ps) (Thm.derivation_name ind, ind, corr, rlz, r)) | 
| 22271 | 414 | realizers @ (case realizers of | 
| 415 | [(((ind, corr), rlz), r)] => | |
| 37233 
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
 berghofe parents: 
36945diff
changeset | 416 | [mk_realizer thy'' (vs' @ Ps) (Long_Name.qualify qualifier "induct", | 
| 22271 | 417 | ind, corr, rlz, r)] | 
| 418 | | _ => [])) thy'' | |
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 419 | end; | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 420 | |
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 421 | (** realizer for elimination rules **) | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 422 | |
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 423 | val case_names = map (fst o dest_Const o head_of o fst o HOLogic.dest_eq o | 
| 31781 
861e675f01e6
add_datatype interface yields type names and less rules
 haftmann parents: 
31723diff
changeset | 424 | HOLogic.dest_Trueprop o prop_of o hd) case_thms; | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 425 | |
| 13921 
69c627b6b28d
Fixed problem in add_elim_realizer which caused bound variables to
 berghofe parents: 
13725diff
changeset | 426 | fun add_elim_realizer Ps | 
| 
69c627b6b28d
Fixed problem in add_elim_realizer which caused bound variables to
 berghofe parents: 
13725diff
changeset | 427 | (((((elim, elimR), intrs), case_thms), case_name), dummy) thy = | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 428 | let | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 429 | val (prem :: prems) = prems_of elim; | 
| 22271 | 430 | fun reorder1 (p, (_, intr)) = | 
| 33244 | 431 | fold (fn ((s, _), T) => Logic.all (Free (s, T))) | 
| 432 | (subtract (op =) params' (Term.add_vars (prop_of intr) [])) | |
| 433 | (strip_all p); | |
| 22271 | 434 | fun reorder2 ((ivs, intr), i) = | 
| 33040 | 435 | let val fs = subtract (op =) params' (Term.add_vars (prop_of intr) []) | 
| 33244 | 436 | in fold (lambda o Var) fs (list_comb (Bound (i + length ivs), ivs)) end; | 
| 13921 
69c627b6b28d
Fixed problem in add_elim_realizer which caused bound variables to
 berghofe parents: 
13725diff
changeset | 437 | val p = Logic.list_implies | 
| 
69c627b6b28d
Fixed problem in add_elim_realizer which caused bound variables to
 berghofe parents: 
13725diff
changeset | 438 | (map reorder1 (prems ~~ intrs) @ [prem], concl_of elim); | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 439 | val T' = Extraction.etype_of thy (vs @ Ps) [] p; | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 440 | 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 | 441 | val Ts = map (Extraction.etype_of thy (vs @ Ps) []) (prems_of elim); | 
| 46219 
426ed18eba43
discontinued old-style Term.list_abs in favour of plain Term.abs;
 wenzelm parents: 
45839diff
changeset | 442 | val r = | 
| 
426ed18eba43
discontinued old-style Term.list_abs in favour of plain Term.abs;
 wenzelm parents: 
45839diff
changeset | 443 | if null Ps then Extraction.nullt | 
| 
426ed18eba43
discontinued old-style Term.list_abs in favour of plain Term.abs;
 wenzelm parents: 
45839diff
changeset | 444 | else | 
| 
426ed18eba43
discontinued old-style Term.list_abs in favour of plain Term.abs;
 wenzelm parents: 
45839diff
changeset | 445 | fold_rev (Term.abs o pair "x") Ts | 
| 
426ed18eba43
discontinued old-style Term.list_abs in favour of plain Term.abs;
 wenzelm parents: 
45839diff
changeset | 446 | (list_comb (Const (case_name, T), | 
| 
426ed18eba43
discontinued old-style Term.list_abs in favour of plain Term.abs;
 wenzelm parents: 
45839diff
changeset | 447 | (if dummy then | 
| 
426ed18eba43
discontinued old-style Term.list_abs in favour of plain Term.abs;
 wenzelm parents: 
45839diff
changeset | 448 |                    [Abs ("x", HOLogic.unitT, Const (@{const_name default}, body_type T))]
 | 
| 
426ed18eba43
discontinued old-style Term.list_abs in favour of plain Term.abs;
 wenzelm parents: 
45839diff
changeset | 449 | else []) @ | 
| 
426ed18eba43
discontinued old-style Term.list_abs in favour of plain Term.abs;
 wenzelm parents: 
45839diff
changeset | 450 | map reorder2 (intrs ~~ (length prems - 1 downto 0)) @ | 
| 
426ed18eba43
discontinued old-style Term.list_abs in favour of plain Term.abs;
 wenzelm parents: 
45839diff
changeset | 451 | [Bound (length prems)])); | 
| 22271 | 452 | val rlz = Extraction.realizes_of thy (vs @ Ps) r (prop_of elim); | 
| 37233 
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
 berghofe parents: 
36945diff
changeset | 453 | val rlz' = attach_typeS (strip_all (Logic.unvarify_global rlz)); | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 454 | val rews = map mk_meta_eq case_thms; | 
| 22271 | 455 | val thm = Goal.prove_global thy [] | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49170diff
changeset | 456 | (Logic.strip_imp_prems rlz') (Logic.strip_imp_concl rlz') | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49170diff
changeset | 457 |           (fn {context = ctxt, prems, ...} => EVERY
 | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49170diff
changeset | 458 | [cut_tac (hd prems) 1, | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49170diff
changeset | 459 | etac elimR 1, | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49170diff
changeset | 460 | ALLGOALS (asm_simp_tac (put_simpset HOL_basic_ss ctxt)), | 
| 54742 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
52788diff
changeset | 461 | rewrite_goals_tac ctxt rews, | 
| 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
52788diff
changeset | 462 | REPEAT ((resolve_tac prems THEN_ALL_NEW (Object_Logic.atomize_prems_tac ctxt THEN' | 
| 58963 
26bf09b95dda
proper context for assume_tac (atac remains as fall-back without context);
 wenzelm parents: 
58372diff
changeset | 463 | DEPTH_SOLVE_1 o FIRST' [assume_tac ctxt, etac allE, etac impE])) 1)]); | 
| 39557 
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
 wenzelm parents: 
37678diff
changeset | 464 | val (thm', thy') = Global_Theory.store_thm (Binding.qualified_name (space_implode "_" | 
| 29579 | 465 | (name_of_thm elim :: vs @ Ps @ ["correctness"])), thm) thy | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 466 | in | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 467 | Extraction.add_realizers_i | 
| 22271 | 468 | [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 | 469 | end; | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 470 | |
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 471 | (** add realizers to theory **) | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 472 | |
| 33244 | 473 | val thy4 = fold add_ind_realizer (subsets Ps) thy3; | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 474 | val thy5 = Extraction.add_realizers_i | 
| 22271 | 475 | (map (mk_realizer thy4 vs) (map (fn (((rule, rrule), rlz), c) => | 
| 476 | (name_of_thm rule, rule, rrule, rlz, | |
| 33040 | 477 | list_comb (c, map Var (subtract (op =) params' (rev (Term.add_vars (prop_of rule) [])))))) | 
| 32952 | 478 | (maps snd rss ~~ #intrs ind_info ~~ rintrs ~~ flat constrss))) thy4; | 
| 479 | val elimps = map_filter (fn ((s, intrs), p) => | |
| 36692 
54b64d4ad524
farewell to old-style mem infixes -- type inference in situations with mem_int and mem_string should provide enough information to resolve the type of (op =)
 haftmann parents: 
36610diff
changeset | 480 | if member (op =) rsets s then SOME (p, intrs) else NONE) | 
| 22271 | 481 | (rss' ~~ (elims ~~ #elims ind_info)); | 
| 33244 | 482 | val thy6 = | 
| 483 | fold (fn p as (((((elim, _), _), _), _), _) => | |
| 484 | add_elim_realizer [] p #> | |
| 485 | add_elim_realizer [fst (fst (dest_Var (HOLogic.dest_Trueprop (concl_of elim))))] p) | |
| 486 | (elimps ~~ case_thms ~~ case_names ~~ dummies) thy5; | |
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 487 | |
| 24712 
64ed05609568
proper Sign operations instead of Theory aliases;
 wenzelm parents: 
24157diff
changeset | 488 | in Sign.restore_naming thy thy6 end; | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 489 | |
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 490 | fun add_ind_realizers name rsets thy = | 
| 
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 |     val (_, {intrs, induct, raw_induct, elims, ...}) =
 | 
| 42361 | 493 | Inductive.the_inductive (Proof_Context.init_global thy) name; | 
| 59058 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 wenzelm parents: 
58963diff
changeset | 494 | val vss = sort (int_ord o apply2 length) | 
| 22271 | 495 | (subsets (map fst (relevant_vars (concl_of (hd intrs))))) | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 496 | in | 
| 37233 
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
 berghofe parents: 
36945diff
changeset | 497 | fold_rev (add_ind_realizer rsets intrs induct raw_induct elims) vss thy | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 498 | end | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 499 | |
| 20897 | 500 | 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 | 501 | let | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 502 | fun err () = error "ind_realizer: bad rule"; | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 503 | val sets = | 
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 504 | (case HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of thm)) of | 
| 22271 | 505 | [_] => [pred_of (HOLogic.dest_Trueprop (hd (prems_of thm)))] | 
| 506 | | xs => map (pred_of o fst o HOLogic.dest_imp) xs) | |
| 47060 
e2741ec9ae36
prefer explicitly qualified exception List.Empty;
 wenzelm parents: 
46708diff
changeset | 507 | handle TERM _ => err () | List.Empty => err (); | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 508 | in | 
| 18728 | 509 | add_ind_realizers (hd sets) | 
| 510 | (case arg of | |
| 15531 | 511 | NONE => sets | SOME NONE => [] | 
| 33040 | 512 | | SOME (SOME sets') => subtract (op =) sets' sets) | 
| 20897 | 513 | end I); | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 514 | |
| 58372 
bfd497f2f4c2
moved 'old_datatype' out of 'Main' (but put it in 'HOL-Proofs' because of the inductive realizer)
 blanchet parents: 
58274diff
changeset | 515 | val _ = Theory.setup (Attrib.setup @{binding ind_realizer}
 | 
| 
bfd497f2f4c2
moved 'old_datatype' out of 'Main' (but put it in 'HOL-Proofs' because of the inductive realizer)
 blanchet parents: 
58274diff
changeset | 516 | ((Scan.option (Scan.lift (Args.$$$ "irrelevant") |-- | 
| 
bfd497f2f4c2
moved 'old_datatype' out of 'Main' (but put it in 'HOL-Proofs' because of the inductive realizer)
 blanchet parents: 
58274diff
changeset | 517 | Scan.option (Scan.lift (Args.colon) |-- | 
| 
bfd497f2f4c2
moved 'old_datatype' out of 'Main' (but put it in 'HOL-Proofs' because of the inductive realizer)
 blanchet parents: 
58274diff
changeset | 518 |       Scan.repeat1 (Args.const {proper = true, strict = true})))) >> rlz_attrib)
 | 
| 
bfd497f2f4c2
moved 'old_datatype' out of 'Main' (but put it in 'HOL-Proofs' because of the inductive realizer)
 blanchet parents: 
58274diff
changeset | 519 | "add realizers for inductive set"); | 
| 13710 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 520 | |
| 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 berghofe parents: diff
changeset | 521 | end; |