| author | wenzelm | 
| Wed, 16 Jan 2019 18:54:18 +0100 | |
| changeset 69673 | cc47e7e06f38 | 
| parent 69593 | 3dda49e08b9d | 
| child 69829 | 3bfa28b3a5b2 | 
| 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: 
28965 
diff
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: 
22596 
diff
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  | 
|
| 62922 | 22  | 
fun prf_of ctxt thm =  | 
23  | 
Reconstruct.proof_of ctxt thm  | 
|
24  | 
  |> Reconstruct.expand_proof ctxt [("", 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  | 
|
| 69593 | 33  | 
fun strip_all' used names (Const (\<^const_name>\<open>Pure.all\<close>, _) $ 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: 
42375 
diff
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  | 
|
| 69593 | 38  | 
| strip_all' used names ((t as Const (\<^const_name>\<open>Pure.imp\<close>, _) $ 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  | 
| 69593 | 45  | 
(Const (\<^const_name>\<open>Pure.all\<close>, _) $ Abs (s, T, Const (\<^const_name>\<open>Pure.imp\<close>, _) $ P $ Q)) =  | 
| 22271 | 46  | 
(subst_bound (Free (name, T), P), subst_bound (Free (name, T), Q))  | 
| 69593 | 47  | 
| strip_one _ (Const (\<^const_name>\<open>Pure.imp\<close>, _) $ 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: 
36945 
diff
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: 
36945 
diff
changeset
 | 
50  | 
(case strip_type T of  | 
| 69593 | 51  | 
(_, Type (s, _)) => if s = \<^type_name>\<open>bool\<close> then (a, T) :: vs else vs  | 
| 
37233
 
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
 
berghofe 
parents: 
36945 
diff
changeset
 | 
52  | 
| _ => vs)) (Term.add_vars prop []) [];  | 
| 
 
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
 
berghofe 
parents: 
36945 
diff
changeset
 | 
53  | 
|
| 
 
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
 
berghofe 
parents: 
36945 
diff
changeset
 | 
54  | 
val attach_typeS = map_types (map_atyps  | 
| 69593 | 55  | 
(fn TFree (s, []) => TFree (s, \<^sort>\<open>type\<close>)  | 
56  | 
| TVar (ixn, []) => TVar (ixn, \<^sort>\<open>type\<close>)  | 
|
| 
37233
 
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
 
berghofe 
parents: 
36945 
diff
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  | 
| 59582 | 61  | 
val iTs = rev (Term.add_tvars (Thm.prop_of (hd intrs)) []);  | 
| 22271 | 62  | 
val (Const (s, _), ts) = strip_comb (HOLogic.dest_Trueprop  | 
| 59582 | 63  | 
(Logic.strip_imp_concl (Thm.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: 
30345 
diff
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: 
30345 
diff
changeset
 | 
66  | 
fun constr_of_intr intr = (Binding.name (Long_Name.base_name (name_of_thm intr)),  | 
| 59582 | 67  | 
map (Logic.unvarifyT_global o snd)  | 
68  | 
(subtract (op =) params (rev (Term.add_vars (Thm.prop_of intr) []))) @  | 
|
69  | 
filter_out (equal Extraction.nullT)  | 
|
70  | 
(map (Logic.unvarifyT_global o Extraction.etype_of thy vs []) (Thm.prems_of intr)), NoSyn);  | 
|
| 
45839
 
43a5b86bc102
'datatype' specifications allow explicit sort constraints;
 
wenzelm 
parents: 
45701 
diff
changeset
 | 
71  | 
in  | 
| 
 
43a5b86bc102
'datatype' specifications allow explicit sort constraints;
 
wenzelm 
parents: 
45701 
diff
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: 
45701 
diff
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: 
36945 
diff
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: 
36610 
diff
changeset
 | 
89  | 
if member (op =) vs a then  | 
| 
46219
 
426ed18eba43
discontinued old-style Term.list_abs in favour of plain Term.abs;
 
wenzelm 
parents: 
45839 
diff
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: 
45839 
diff
changeset
 | 
91  | 
else  | 
| 
 
426ed18eba43
discontinued old-style Term.list_abs in favour of plain Term.abs;
 
wenzelm 
parents: 
45839 
diff
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  | 
| 59582 | 98  | 
val intr = map_types Type.strip_sorts (Thm.prop_of (hd intrs));  | 
| 
37233
 
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
 
berghofe 
parents: 
36945 
diff
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: 
36945 
diff
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: 
36945 
diff
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: 
30345 
diff
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: 
36945 
diff
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)  | 
| 59582 | 141  | 
(subtract (op =) params (Term.add_vars (Thm.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: 
29265 
diff
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  | 
|
| 69593 | 148  | 
fun is_meta (Const (\<^const_name>\<open>Pure.all\<close>, _) $ Abs (s, _, P)) = is_meta P  | 
149  | 
| is_meta (Const (\<^const_name>\<open>Pure.imp\<close>, _) $ _ $ Q) = is_meta Q  | 
|
150  | 
| is_meta (Const (\<^const_name>\<open>Trueprop\<close>, _) $ t) =  | 
|
| 35364 | 151  | 
(case head_of t of  | 
| 
65411
 
3c628937899d
use Item_Net to store inductive info
 
Lars Hupel <lars.hupel@mytum.de> 
parents: 
62922 
diff
changeset
 | 
152  | 
Const (s, _) => can (Inductive.the_inductive_global ctxt) s  | 
| 35364 | 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: 
19806 
diff
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: 
45839 
diff
changeset
 | 
167  | 
in  | 
| 
 
426ed18eba43
discontinued old-style Term.list_abs in favour of plain Term.abs;
 
wenzelm 
parents: 
45839 
diff
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: 
45839 
diff
changeset
 | 
175  | 
else  | 
| 
 
426ed18eba43
discontinued old-style Term.list_abs in favour of plain Term.abs;
 
wenzelm 
parents: 
45839 
diff
changeset
 | 
176  | 
(case strip_type T of  | 
| 69593 | 177  | 
(Ts, Type (\<^type_name>\<open>Product_Type.prod\<close>, [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: 
45839 
diff
changeset
 | 
182  | 
val t =  | 
| 
 
426ed18eba43
discontinued old-style Term.list_abs in favour of plain Term.abs;
 
wenzelm 
parents: 
45839 
diff
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: 
45839 
diff
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: 
45839 
diff
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: 
13725 
diff
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  | 
| 59582 | 206  | 
val concls = HOLogic.dest_conj (HOLogic.dest_Trueprop (Thm.concl_of raw_induct));  | 
| 31986 | 207  | 
val premss = map_filter (fn (s, rs) => if member (op =) rsets s then  | 
| 59582 | 208  | 
SOME (rs, map (fn (_, r) => nth (Thm.prems_of raw_induct)  | 
209  | 
(find_index (fn prp => prp = Thm.prop_of r) (map Thm.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  | 
| 69593 | 215  | 
if dummy then Const (\<^const_name>\<open>default\<close>,  | 
| 35364 | 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: 
30345 
diff
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: 
45701 
diff
changeset
 | 
238  | 
fun add_dummy name dname (x as (_, ((s, vs, mx), cs))) =  | 
| 
 
43a5b86bc102
'datatype' specifications allow explicit sort constraints;
 
wenzelm 
parents: 
45701 
diff
changeset
 | 
239  | 
if Binding.eq_name (name, s)  | 
| 
 
43a5b86bc102
'datatype' specifications allow explicit sort constraints;
 
wenzelm 
parents: 
45701 
diff
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))  | 
| 
67316
 
adaf279ce67b
ported inductive realizer to new datatype package
 
blanchet 
parents: 
67149 
diff
changeset
 | 
249  | 
handle BNF_FP_Util.EMPTY_DATATYPE 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: 
30345 
diff
changeset
 | 
251  | 
val name = Long_Name.base_name name';  | 
| 
43324
 
2b47822868e4
discontinued Name.variant to emphasize that this is old-style / indirect;
 
wenzelm 
parents: 
42375 
diff
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: 
13928 
diff
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  | 
| 62922 | 260  | 
val ctxt = Proof_Context.init_global thy;  | 
| 59582 | 261  | 
val rvs = map fst (relevant_vars (Thm.prop_of rule));  | 
262  | 
val xs = rev (Term.add_vars (Thm.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: 
36610 
diff
changeset
 | 
263  | 
val vs1 = map Var (filter_out (fn ((a, _), _) => member (op =) rvs a) xs);  | 
| 59582 | 264  | 
val rlzvs = rev (Term.add_vars (Thm.prop_of rrule) []);  | 
| 17485 | 265  | 
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: 
58963 
diff
changeset
 | 
266  | 
val rs = map Var (subtract (op = o apply2 fst) xs rlzvs);  | 
| 59582 | 267  | 
val rlz' = fold_rev Logic.all rs (Thm.prop_of rrule)  | 
| 22271 | 268  | 
in (name, (vs,  | 
| 33338 | 269  | 
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: 
36945 
diff
changeset
 | 
270  | 
Extraction.abs_corr_shyps thy rule vs vs2  | 
| 
 
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
 
berghofe 
parents: 
36945 
diff
changeset
 | 
271  | 
(ProofRewriteRules.un_hhf_proof rlz' (attach_typeS rlz)  | 
| 62922 | 272  | 
(fold_rev Proofterm.forall_intr_proof' rs (prf_of ctxt rrule)))))  | 
| 
13710
 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 
berghofe 
parents:  
diff
changeset
 | 
273  | 
end;  | 
| 
 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 
berghofe 
parents:  
diff
changeset
 | 
274  | 
|
| 
24157
 
409cd6eaa7ea
Added renaming function to prevent correctness proof for realizer
 
berghofe 
parents: 
23590 
diff
changeset
 | 
275  | 
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: 
23590 
diff
changeset
 | 
276  | 
|
| 33244 | 277  | 
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
 | 
278  | 
let  | 
| 
30364
 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 
wenzelm 
parents: 
30345 
diff
changeset
 | 
279  | 
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: 
37678 
diff
changeset
 | 
280  | 
val inducts = Global_Theory.get_thms thy (Long_Name.qualify qualifier "inducts");  | 
| 59582 | 281  | 
val iTs = rev (Term.add_tvars (Thm.prop_of (hd intrs)) []);  | 
| 
13710
 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 
berghofe 
parents:  
diff
changeset
 | 
282  | 
val ar = length vs + length iTs;  | 
| 
31723
 
f5cafe803b55
discontinued ancient tradition to suffix certain ML module names with "_package"
 
haftmann 
parents: 
31668 
diff
changeset
 | 
283  | 
val params = Inductive.params_of raw_induct;  | 
| 
 
f5cafe803b55
discontinued ancient tradition to suffix certain ML module names with "_package"
 
haftmann 
parents: 
31668 
diff
changeset
 | 
284  | 
val arities = Inductive.arities_of raw_induct;  | 
| 22271 | 285  | 
val nparms = length params;  | 
| 
13710
 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 
berghofe 
parents:  
diff
changeset
 | 
286  | 
val params' = map dest_Var params;  | 
| 
31723
 
f5cafe803b55
discontinued ancient tradition to suffix certain ML module names with "_package"
 
haftmann 
parents: 
31668 
diff
changeset
 | 
287  | 
val rss = Inductive.partition_rules raw_induct intrs;  | 
| 22271 | 288  | 
val rss' = map (fn (((s, rs), (_, arity)), elim) =>  | 
| 60362 | 289  | 
(s, (Inductive.infer_intro_vars thy elim arity rs ~~ rs)))  | 
| 
22790
 
e1cff9268177
Moved functions infer_intro_vars, arities_of, params_of, and
 
berghofe 
parents: 
22606 
diff
changeset
 | 
290  | 
(rss ~~ arities ~~ elims);  | 
| 
30364
 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 
wenzelm 
parents: 
30345 
diff
changeset
 | 
291  | 
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
 | 
292  | 
val tnames = map (fn s => space_implode "_" (s ^ "T" :: vs)) rsets;  | 
| 16123 | 293  | 
|
| 
13710
 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 
berghofe 
parents:  
diff
changeset
 | 
294  | 
val thy1 = thy |>  | 
| 
24712
 
64ed05609568
proper Sign operations instead of Theory aliases;
 
wenzelm 
parents: 
24157 
diff
changeset
 | 
295  | 
Sign.root_path |>  | 
| 
30364
 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 
wenzelm 
parents: 
30345 
diff
changeset
 | 
296  | 
Sign.add_path (Long_Name.implode prfx);  | 
| 
13710
 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 
berghofe 
parents:  
diff
changeset
 | 
297  | 
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: 
36610 
diff
changeset
 | 
298  | 
(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
 | 
299  | 
|
| 
 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 
berghofe 
parents:  
diff
changeset
 | 
300  | 
val thy1' = thy1 |>  | 
| 
42375
 
774df7c59508
report Name_Space.declare/define, relatively to context;
 
wenzelm 
parents: 
42361 
diff
changeset
 | 
301  | 
Sign.add_types_global  | 
| 
 
774df7c59508
report Name_Space.declare/define, relatively to context;
 
wenzelm 
parents: 
42361 
diff
changeset
 | 
302  | 
(map (fn s => (Binding.name (Long_Name.base_name s), ar, NoSyn)) tnames) |>  | 
| 
 
774df7c59508
report Name_Space.declare/define, relatively to context;
 
wenzelm 
parents: 
42361 
diff
changeset
 | 
303  | 
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: 
36610 
diff
changeset
 | 
304  | 
val dts = map_filter (fn (s, rs) => if member (op =) rsets s 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  | 
|
| 
31783
 
cfbe9609ceb1
add_datatypes does not yield particular rules any longer
 
haftmann 
parents: 
31781 
diff
changeset
 | 
309  | 
val ((dummies, some_dt_names), thy2) =  | 
| 18008 | 310  | 
thy1  | 
| 
67316
 
adaf279ce67b
ported inductive realizer to new datatype package
 
blanchet 
parents: 
67149 
diff
changeset
 | 
311  | 
|> add_dummies (BNF_LFP_Compat.add_datatype [BNF_LFP_Compat.Kill_Type_Args])  | 
| 45701 | 312  | 
(map (pair false) dts) []  | 
| 18314 | 313  | 
||> Extraction.add_typeof_eqns_i ty_eqs  | 
314  | 
||> Extraction.add_realizes_eqns_i rlz_eqs;  | 
|
| 
31783
 
cfbe9609ceb1
add_datatypes does not yield particular rules any longer
 
haftmann 
parents: 
31781 
diff
changeset
 | 
315  | 
val dt_names = these some_dt_names;  | 
| 
67316
 
adaf279ce67b
ported inductive realizer to new datatype package
 
blanchet 
parents: 
67149 
diff
changeset
 | 
316  | 
val case_thms = map (#case_rewrites o BNF_LFP_Compat.the_info thy2 []) dt_names;  | 
| 45701 | 317  | 
val rec_thms =  | 
318  | 
if null dt_names then []  | 
|
| 
67316
 
adaf279ce67b
ported inductive realizer to new datatype package
 
blanchet 
parents: 
67149 
diff
changeset
 | 
319  | 
else #rec_rewrites (BNF_LFP_Compat.the_info thy2 [] (hd dt_names));  | 
| 
19046
 
bc5c6c9b114e
removed distinct, renamed gen_distinct to distinct;
 
wenzelm 
parents: 
18929 
diff
changeset
 | 
320  | 
val rec_names = distinct (op =) (map (fst o dest_Const o head_of o fst o  | 
| 59582 | 321  | 
HOLogic.dest_eq o HOLogic.dest_Trueprop o Thm.prop_of) rec_thms);  | 
| 31458 | 322  | 
val (constrss, _) = fold_map (fn (s, rs) => fn (recs, dummies) =>  | 
323  | 
if member (op =) rsets s then  | 
|
| 
13710
 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 
berghofe 
parents:  
diff
changeset
 | 
324  | 
let  | 
| 
 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 
berghofe 
parents:  
diff
changeset
 | 
325  | 
val (d :: dummies') = dummies;  | 
| 19473 | 326  | 
val (recs1, recs2) = chop (length rs) (if d then tl recs else recs)  | 
| 31458 | 327  | 
in (map (head_of o hd o rev o snd o strip_comb o fst o  | 
| 59582 | 328  | 
HOLogic.dest_eq o HOLogic.dest_Trueprop o Thm.prop_of) recs1, (recs2, dummies'))  | 
| 
13710
 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 
berghofe 
parents:  
diff
changeset
 | 
329  | 
end  | 
| 31458 | 330  | 
else (replicate (length rs) Extraction.nullt, (recs, dummies)))  | 
| 
31781
 
861e675f01e6
add_datatype interface yields type names and less rules
 
haftmann 
parents: 
31723 
diff
changeset
 | 
331  | 
rss (rec_thms, dummies);  | 
| 
37233
 
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
 
berghofe 
parents: 
36945 
diff
changeset
 | 
332  | 
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
 | 
333  | 
(Extraction.realizes_of thy2 vs  | 
| 22271 | 334  | 
(if c = Extraction.nullt then c else list_comb (c, map Var (rev  | 
| 59582 | 335  | 
(subtract (op =) params' (Term.add_vars (Thm.prop_of intr) []))))) (Thm.prop_of intr))))  | 
| 32952 | 336  | 
(maps snd rss ~~ flat constrss);  | 
| 30345 | 337  | 
val (rlzpreds, rlzpreds') =  | 
338  | 
rintrs |> map (fn rintr =>  | 
|
| 22271 | 339  | 
let  | 
| 30345 | 340  | 
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: 
30345 
diff
changeset
 | 
341  | 
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: 
35625 
diff
changeset
 | 
342  | 
val T' = Logic.unvarifyT_global T;  | 
| 30345 | 343  | 
in (((s', T'), NoSyn), (Const (s, T'), Free (s', T'))) end)  | 
| 
59058
 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 
wenzelm 
parents: 
58963 
diff
changeset
 | 
344  | 
|> distinct (op = o apply2 (#1 o #1))  | 
| 30345 | 345  | 
|> map (apfst (apfst (apfst Binding.name)))  | 
346  | 
|> split_list;  | 
|
347  | 
||
| 
35845
 
e5980f0ad025
renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
 
wenzelm 
parents: 
35625 
diff
changeset
 | 
348  | 
val rlzparams = map (fn Var ((s, _), T) => (s, Logic.unvarifyT_global T))  | 
| 22271 | 349  | 
(List.take (snd (strip_comb  | 
350  | 
(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
 | 
351  | 
|
| 
 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 
berghofe 
parents:  
diff
changeset
 | 
352  | 
(** realizability predicate **)  | 
| 
 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 
berghofe 
parents:  
diff
changeset
 | 
353  | 
|
| 22271 | 354  | 
val (ind_info, thy3') = thy2 |>  | 
| 
33726
 
0878aecbf119
eliminated slightly odd name space grouping -- now managed by Isar toplevel;
 
wenzelm 
parents: 
33671 
diff
changeset
 | 
355  | 
Inductive.add_inductive_global  | 
| 33669 | 356  | 
        {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: 
47060 
diff
changeset
 | 
357  | 
no_elim = false, no_ind = false, skip_mono = false}  | 
| 22271 | 358  | 
rlzpreds rlzparams (map (fn (rintr, intr) =>  | 
| 
30364
 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 
wenzelm 
parents: 
30345 
diff
changeset
 | 
359  | 
((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: 
35625 
diff
changeset
 | 
360  | 
subst_atomic rlzpreds' (Logic.unvarify_global rintr)))  | 
| 22271 | 361  | 
(rintrs ~~ maps snd rss)) [] ||>  | 
| 
30435
 
e62d6ecab6ad
explicit Binding.qualified_name -- prevents implicitly qualified bstring;
 
wenzelm 
parents: 
30364 
diff
changeset
 | 
362  | 
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: 
37678 
diff
changeset
 | 
363  | 
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
 | 
364  | 
|
| 
 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 
berghofe 
parents:  
diff
changeset
 | 
365  | 
(** realizer for induction rule **)  | 
| 
 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 
berghofe 
parents:  
diff
changeset
 | 
366  | 
|
| 
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: 
36610 
diff
changeset
 | 
367  | 
val Ps = map_filter (fn _ $ M $ P => if member (op =) rsets (pred_of M) then  | 
| 15531 | 368  | 
SOME (fst (fst (dest_Var (head_of P)))) else NONE)  | 
| 59582 | 369  | 
(HOLogic.dest_conj (HOLogic.dest_Trueprop (Thm.concl_of raw_induct)));  | 
| 
13710
 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 
berghofe 
parents:  
diff
changeset
 | 
370  | 
|
| 33244 | 371  | 
fun add_ind_realizer Ps thy =  | 
| 
13710
 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 
berghofe 
parents:  
diff
changeset
 | 
372  | 
let  | 
| 
59058
 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 
wenzelm 
parents: 
58963 
diff
changeset
 | 
373  | 
val vs' = rename (map (apply2 (fst o fst o dest_Var))  | 
| 
24157
 
409cd6eaa7ea
Added renaming function to prevent correctness proof for realizer
 
berghofe 
parents: 
23590 
diff
changeset
 | 
374  | 
(params ~~ List.take (snd (strip_comb (HOLogic.dest_Trueprop  | 
| 59582 | 375  | 
(hd (Thm.prems_of (hd inducts))))), nparms))) vs;  | 
| 22271 | 376  | 
val rs = indrule_realizer thy induct raw_induct rsets params'  | 
| 
24157
 
409cd6eaa7ea
Added renaming function to prevent correctness proof for realizer
 
berghofe 
parents: 
23590 
diff
changeset
 | 
377  | 
(vs' @ Ps) rec_names rss' intrs dummies;  | 
| 
 
409cd6eaa7ea
Added renaming function to prevent correctness proof for realizer
 
berghofe 
parents: 
23590 
diff
changeset
 | 
378  | 
val rlzs = map (fn (r, ind) => Extraction.realizes_of thy (vs' @ Ps) r  | 
| 59582 | 379  | 
(Thm.prop_of ind)) (rs ~~ inducts);  | 
| 29281 | 380  | 
val used = fold Term.add_free_names rlzs [];  | 
| 22271 | 381  | 
val rnames = Name.variant_list used (replicate (length inducts) "r");  | 
382  | 
val rnames' = Name.variant_list  | 
|
383  | 
(used @ rnames) (replicate (length intrs) "s");  | 
|
384  | 
val rlzs' as (prems, _, _) :: _ = map (fn (rlz, name) =>  | 
|
385  | 
let  | 
|
| 
35845
 
e5980f0ad025
renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
 
wenzelm 
parents: 
35625 
diff
changeset
 | 
386  | 
val (P, Q) = strip_one name (Logic.unvarify_global rlz);  | 
| 22271 | 387  | 
val Q' = strip_all' [] rnames' Q  | 
388  | 
in  | 
|
389  | 
(Logic.strip_imp_prems Q', P, Logic.strip_imp_concl Q')  | 
|
390  | 
end) (rlzs ~~ rnames);  | 
|
391  | 
val concl = HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj (map  | 
|
392  | 
(fn (_, _ $ P, _ $ Q) => HOLogic.mk_imp (P, Q)) rlzs'));  | 
|
| 37136 | 393  | 
        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: 
36945 
diff
changeset
 | 
394  | 
val thm = Goal.prove_global thy []  | 
| 
 
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
 
berghofe 
parents: 
36945 
diff
changeset
 | 
395  | 
(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: 
52788 
diff
changeset
 | 
396  | 
          (fn {context = ctxt, prems} => EVERY
 | 
| 60752 | 397  | 
[resolve_tac ctxt [#raw_induct ind_info] 1,  | 
| 
54742
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
52788 
diff
changeset
 | 
398  | 
rewrite_goals_tac ctxt rews,  | 
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59058 
diff
changeset
 | 
399  | 
REPEAT ((resolve_tac ctxt prems THEN_ALL_NEW EVERY'  | 
| 
54742
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
52788 
diff
changeset
 | 
400  | 
[K (rewrite_goals_tac ctxt rews), Object_Logic.atomize_prems_tac ctxt,  | 
| 60752 | 401  | 
DEPTH_SOLVE_1 o  | 
402  | 
FIRST' [assume_tac ctxt, eresolve_tac ctxt [allE], eresolve_tac ctxt [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: 
37678 
diff
changeset
 | 
403  | 
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: 
30345 
diff
changeset
 | 
404  | 
(Long_Name.qualify qualifier "induct" :: vs' @ Ps @ ["correctness"])), thm) thy;  | 
| 22271 | 405  | 
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: 
58111 
diff
changeset
 | 
406  | 
(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: 
37678 
diff
changeset
 | 
407  | 
val ([thms'], thy'') = Global_Theory.add_thmss  | 
| 
30435
 
e62d6ecab6ad
explicit Binding.qualified_name -- prevents implicitly qualified bstring;
 
wenzelm 
parents: 
30364 
diff
changeset
 | 
408  | 
[((Binding.qualified_name (space_implode "_"  | 
| 
30364
 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 
wenzelm 
parents: 
30345 
diff
changeset
 | 
409  | 
(Long_Name.qualify qualifier "inducts" :: vs' @ Ps @  | 
| 29579 | 410  | 
["correctness"])), thms), [])] thy';  | 
| 22271 | 411  | 
val realizers = inducts ~~ thms' ~~ rlzs ~~ rs;  | 
| 
13710
 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 
berghofe 
parents:  
diff
changeset
 | 
412  | 
in  | 
| 
 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 
berghofe 
parents:  
diff
changeset
 | 
413  | 
Extraction.add_realizers_i  | 
| 22271 | 414  | 
(map (fn (((ind, corr), rlz), r) =>  | 
| 
37233
 
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
 
berghofe 
parents: 
36945 
diff
changeset
 | 
415  | 
mk_realizer thy'' (vs' @ Ps) (Thm.derivation_name ind, ind, corr, rlz, r))  | 
| 22271 | 416  | 
realizers @ (case realizers of  | 
417  | 
[(((ind, corr), rlz), r)] =>  | 
|
| 
37233
 
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
 
berghofe 
parents: 
36945 
diff
changeset
 | 
418  | 
[mk_realizer thy'' (vs' @ Ps) (Long_Name.qualify qualifier "induct",  | 
| 22271 | 419  | 
ind, corr, rlz, r)]  | 
420  | 
| _ => [])) thy''  | 
|
| 
13710
 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 
berghofe 
parents:  
diff
changeset
 | 
421  | 
end;  | 
| 
 
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  | 
(** realizer for elimination rules **)  | 
| 
 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 
berghofe 
parents:  
diff
changeset
 | 
424  | 
|
| 
 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 
berghofe 
parents:  
diff
changeset
 | 
425  | 
val case_names = map (fst o dest_Const o head_of o fst o HOLogic.dest_eq o  | 
| 59582 | 426  | 
HOLogic.dest_Trueprop o Thm.prop_of o hd) case_thms;  | 
| 
13710
 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 
berghofe 
parents:  
diff
changeset
 | 
427  | 
|
| 
13921
 
69c627b6b28d
Fixed problem in add_elim_realizer which caused bound variables to
 
berghofe 
parents: 
13725 
diff
changeset
 | 
428  | 
fun add_elim_realizer Ps  | 
| 
 
69c627b6b28d
Fixed problem in add_elim_realizer which caused bound variables to
 
berghofe 
parents: 
13725 
diff
changeset
 | 
429  | 
(((((elim, elimR), intrs), case_thms), case_name), dummy) thy =  | 
| 
13710
 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 
berghofe 
parents:  
diff
changeset
 | 
430  | 
let  | 
| 59582 | 431  | 
val (prem :: prems) = Thm.prems_of elim;  | 
| 22271 | 432  | 
fun reorder1 (p, (_, intr)) =  | 
| 33244 | 433  | 
fold (fn ((s, _), T) => Logic.all (Free (s, T)))  | 
| 59582 | 434  | 
(subtract (op =) params' (Term.add_vars (Thm.prop_of intr) []))  | 
| 33244 | 435  | 
(strip_all p);  | 
| 22271 | 436  | 
fun reorder2 ((ivs, intr), i) =  | 
| 59582 | 437  | 
let val fs = subtract (op =) params' (Term.add_vars (Thm.prop_of intr) [])  | 
| 33244 | 438  | 
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: 
13725 
diff
changeset
 | 
439  | 
val p = Logic.list_implies  | 
| 59582 | 440  | 
(map reorder1 (prems ~~ intrs) @ [prem], Thm.concl_of elim);  | 
| 
13710
 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 
berghofe 
parents:  
diff
changeset
 | 
441  | 
val T' = Extraction.etype_of thy (vs @ Ps) [] p;  | 
| 
 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 
berghofe 
parents:  
diff
changeset
 | 
442  | 
val T = if dummy then (HOLogic.unitT --> body_type T') --> T' else T';  | 
| 59582 | 443  | 
val Ts = map (Extraction.etype_of thy (vs @ Ps) []) (Thm.prems_of elim);  | 
| 
46219
 
426ed18eba43
discontinued old-style Term.list_abs in favour of plain Term.abs;
 
wenzelm 
parents: 
45839 
diff
changeset
 | 
444  | 
val r =  | 
| 
 
426ed18eba43
discontinued old-style Term.list_abs in favour of plain Term.abs;
 
wenzelm 
parents: 
45839 
diff
changeset
 | 
445  | 
if null Ps then Extraction.nullt  | 
| 
 
426ed18eba43
discontinued old-style Term.list_abs in favour of plain Term.abs;
 
wenzelm 
parents: 
45839 
diff
changeset
 | 
446  | 
else  | 
| 
 
426ed18eba43
discontinued old-style Term.list_abs in favour of plain Term.abs;
 
wenzelm 
parents: 
45839 
diff
changeset
 | 
447  | 
fold_rev (Term.abs o pair "x") Ts  | 
| 
 
426ed18eba43
discontinued old-style Term.list_abs in favour of plain Term.abs;
 
wenzelm 
parents: 
45839 
diff
changeset
 | 
448  | 
(list_comb (Const (case_name, T),  | 
| 
 
426ed18eba43
discontinued old-style Term.list_abs in favour of plain Term.abs;
 
wenzelm 
parents: 
45839 
diff
changeset
 | 
449  | 
(if dummy then  | 
| 69593 | 450  | 
                   [Abs ("x", HOLogic.unitT, Const (\<^const_name>\<open>default\<close>, body_type T))]
 | 
| 
46219
 
426ed18eba43
discontinued old-style Term.list_abs in favour of plain Term.abs;
 
wenzelm 
parents: 
45839 
diff
changeset
 | 
451  | 
else []) @  | 
| 
 
426ed18eba43
discontinued old-style Term.list_abs in favour of plain Term.abs;
 
wenzelm 
parents: 
45839 
diff
changeset
 | 
452  | 
map reorder2 (intrs ~~ (length prems - 1 downto 0)) @  | 
| 
 
426ed18eba43
discontinued old-style Term.list_abs in favour of plain Term.abs;
 
wenzelm 
parents: 
45839 
diff
changeset
 | 
453  | 
[Bound (length prems)]));  | 
| 59582 | 454  | 
val rlz = Extraction.realizes_of thy (vs @ Ps) r (Thm.prop_of elim);  | 
| 
37233
 
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
 
berghofe 
parents: 
36945 
diff
changeset
 | 
455  | 
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
 | 
456  | 
val rews = map mk_meta_eq case_thms;  | 
| 22271 | 457  | 
val thm = Goal.prove_global thy []  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
49170 
diff
changeset
 | 
458  | 
(Logic.strip_imp_prems rlz') (Logic.strip_imp_concl rlz')  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
49170 
diff
changeset
 | 
459  | 
          (fn {context = ctxt, prems, ...} => EVERY
 | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
49170 
diff
changeset
 | 
460  | 
[cut_tac (hd prems) 1,  | 
| 60752 | 461  | 
eresolve_tac ctxt [elimR] 1,  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
49170 
diff
changeset
 | 
462  | 
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: 
52788 
diff
changeset
 | 
463  | 
rewrite_goals_tac ctxt rews,  | 
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59058 
diff
changeset
 | 
464  | 
REPEAT ((resolve_tac ctxt prems THEN_ALL_NEW (Object_Logic.atomize_prems_tac ctxt THEN'  | 
| 60752 | 465  | 
DEPTH_SOLVE_1 o  | 
466  | 
FIRST' [assume_tac ctxt, eresolve_tac ctxt [allE], eresolve_tac ctxt [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: 
37678 
diff
changeset
 | 
467  | 
val (thm', thy') = Global_Theory.store_thm (Binding.qualified_name (space_implode "_"  | 
| 29579 | 468  | 
(name_of_thm elim :: vs @ Ps @ ["correctness"])), thm) thy  | 
| 
13710
 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 
berghofe 
parents:  
diff
changeset
 | 
469  | 
in  | 
| 
 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 
berghofe 
parents:  
diff
changeset
 | 
470  | 
Extraction.add_realizers_i  | 
| 22271 | 471  | 
[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
 | 
472  | 
end;  | 
| 
 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 
berghofe 
parents:  
diff
changeset
 | 
473  | 
|
| 
 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 
berghofe 
parents:  
diff
changeset
 | 
474  | 
(** add realizers to theory **)  | 
| 
 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 
berghofe 
parents:  
diff
changeset
 | 
475  | 
|
| 33244 | 476  | 
val thy4 = fold add_ind_realizer (subsets Ps) thy3;  | 
| 
13710
 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 
berghofe 
parents:  
diff
changeset
 | 
477  | 
val thy5 = Extraction.add_realizers_i  | 
| 22271 | 478  | 
(map (mk_realizer thy4 vs) (map (fn (((rule, rrule), rlz), c) =>  | 
479  | 
(name_of_thm rule, rule, rrule, rlz,  | 
|
| 59582 | 480  | 
list_comb (c, map Var (subtract (op =) params' (rev (Term.add_vars (Thm.prop_of rule) []))))))  | 
| 32952 | 481  | 
(maps snd rss ~~ #intrs ind_info ~~ rintrs ~~ flat constrss))) thy4;  | 
482  | 
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: 
36610 
diff
changeset
 | 
483  | 
if member (op =) rsets s then SOME (p, intrs) else NONE)  | 
| 22271 | 484  | 
(rss' ~~ (elims ~~ #elims ind_info));  | 
| 33244 | 485  | 
val thy6 =  | 
486  | 
fold (fn p as (((((elim, _), _), _), _), _) =>  | 
|
487  | 
add_elim_realizer [] p #>  | 
|
| 59582 | 488  | 
add_elim_realizer [fst (fst (dest_Var (HOLogic.dest_Trueprop (Thm.concl_of elim))))] p)  | 
| 33244 | 489  | 
(elimps ~~ case_thms ~~ case_names ~~ dummies) thy5;  | 
| 
13710
 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 
berghofe 
parents:  
diff
changeset
 | 
490  | 
|
| 
24712
 
64ed05609568
proper Sign operations instead of Theory aliases;
 
wenzelm 
parents: 
24157 
diff
changeset
 | 
491  | 
in Sign.restore_naming thy thy6 end;  | 
| 
13710
 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 
berghofe 
parents:  
diff
changeset
 | 
492  | 
|
| 
 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 
berghofe 
parents:  
diff
changeset
 | 
493  | 
fun add_ind_realizers name rsets thy =  | 
| 
 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 
berghofe 
parents:  
diff
changeset
 | 
494  | 
let  | 
| 
 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 
berghofe 
parents:  
diff
changeset
 | 
495  | 
    val (_, {intrs, induct, raw_induct, elims, ...}) =
 | 
| 
65411
 
3c628937899d
use Item_Net to store inductive info
 
Lars Hupel <lars.hupel@mytum.de> 
parents: 
62922 
diff
changeset
 | 
496  | 
Inductive.the_inductive_global (Proof_Context.init_global thy) name;  | 
| 
59058
 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 
wenzelm 
parents: 
58963 
diff
changeset
 | 
497  | 
val vss = sort (int_ord o apply2 length)  | 
| 59582 | 498  | 
(subsets (map fst (relevant_vars (Thm.concl_of (hd intrs)))))  | 
| 
13710
 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 
berghofe 
parents:  
diff
changeset
 | 
499  | 
in  | 
| 
37233
 
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
 
berghofe 
parents: 
36945 
diff
changeset
 | 
500  | 
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
 | 
501  | 
end  | 
| 
 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 
berghofe 
parents:  
diff
changeset
 | 
502  | 
|
| 20897 | 503  | 
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
 | 
504  | 
let  | 
| 
 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 
berghofe 
parents:  
diff
changeset
 | 
505  | 
fun err () = error "ind_realizer: bad rule";  | 
| 
 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 
berghofe 
parents:  
diff
changeset
 | 
506  | 
val sets =  | 
| 59582 | 507  | 
(case HOLogic.dest_conj (HOLogic.dest_Trueprop (Thm.concl_of thm)) of  | 
508  | 
[_] => [pred_of (HOLogic.dest_Trueprop (hd (Thm.prems_of thm)))]  | 
|
| 22271 | 509  | 
| xs => map (pred_of o fst o HOLogic.dest_imp) xs)  | 
| 
47060
 
e2741ec9ae36
prefer explicitly qualified exception List.Empty;
 
wenzelm 
parents: 
46708 
diff
changeset
 | 
510  | 
handle TERM _ => err () | List.Empty => err ();  | 
| 
13710
 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 
berghofe 
parents:  
diff
changeset
 | 
511  | 
in  | 
| 18728 | 512  | 
add_ind_realizers (hd sets)  | 
513  | 
(case arg of  | 
|
| 15531 | 514  | 
NONE => sets | SOME NONE => []  | 
| 33040 | 515  | 
| SOME (SOME sets') => subtract (op =) sets' sets)  | 
| 20897 | 516  | 
end I);  | 
| 
13710
 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 
berghofe 
parents:  
diff
changeset
 | 
517  | 
|
| 67149 | 518  | 
val _ = Theory.setup (Attrib.setup \<^binding>\<open>ind_realizer\<close>  | 
| 
58372
 
bfd497f2f4c2
moved 'old_datatype' out of 'Main' (but put it in 'HOL-Proofs' because of the inductive realizer)
 
blanchet 
parents: 
58274 
diff
changeset
 | 
519  | 
((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: 
58274 
diff
changeset
 | 
520  | 
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: 
58274 
diff
changeset
 | 
521  | 
      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: 
58274 
diff
changeset
 | 
522  | 
"add realizers for inductive set");  | 
| 
13710
 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 
berghofe 
parents:  
diff
changeset
 | 
523  | 
|
| 
 
75bec2c1bfd5
New package for constructing realizers for introduction and elimination
 
berghofe 
parents:  
diff
changeset
 | 
524  | 
end;  |