author | Christian Urban <urbanc@in.tum.de> |
Fri, 22 Jul 2011 11:37:16 +0100 | |
changeset 2982 | 4a00077c008f |
parent 2981 | c8acaded1777 |
child 2983 | 4436039cc5e1 |
permissions | -rw-r--r-- |
2665
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1 |
(* Nominal Mutual Functions |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
2 |
Author: Christian Urban |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
3 |
|
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
4 |
heavily based on the code of Alexander Krauss |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
5 |
(code forked on 14 January 2011) |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
6 |
|
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
7 |
|
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
8 |
Mutual recursive nominal function definitions. |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
9 |
*) |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
10 |
|
2982
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
11 |
|
2665
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
12 |
signature NOMINAL_FUNCTION_MUTUAL = |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
13 |
sig |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
14 |
|
2819
4bd584ff4fab
added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
2781
diff
changeset
|
15 |
val prepare_nominal_function_mutual : Nominal_Function_Common.nominal_function_config |
2665
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
16 |
-> string (* defname *) |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
17 |
-> ((string * typ) * mixfix) list |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
18 |
-> term list |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
19 |
-> local_theory |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
20 |
-> ((thm (* goalstate *) |
2973
d1038e67923a
added a flag (eqvt) to termination proofs arising fron nominal_primrecs
Christian Urban <urbanc@in.tum.de>
parents:
2821
diff
changeset
|
21 |
* (thm -> Nominal_Function_Common.nominal_function_result) (* proof continuation *) |
2665
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
22 |
) * local_theory) |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
23 |
|
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
24 |
end |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
25 |
|
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
26 |
|
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
27 |
structure Nominal_Function_Mutual: NOMINAL_FUNCTION_MUTUAL = |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
28 |
struct |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
29 |
|
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
30 |
open Function_Lib |
2821
c7d4bd9e89e0
fixed problem with earlier commit about nominal_function_common; added facility for specifying an invariant - added a definition of frees_set which need a finiteness invariant
Christian Urban <urbanc@in.tum.de>
parents:
2819
diff
changeset
|
31 |
open Function_Common |
2819
4bd584ff4fab
added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
2781
diff
changeset
|
32 |
open Nominal_Function_Common |
2665
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
33 |
|
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
34 |
type qgar = string * (string * typ) list * term list * term list * term |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
35 |
|
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
36 |
datatype mutual_part = MutualPart of |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
37 |
{i : int, |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
38 |
i' : int, |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
39 |
fvar : string * typ, |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
40 |
cargTs: typ list, |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
41 |
f_def: term, |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
42 |
f: term option, |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
43 |
f_defthm : thm option} |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
44 |
|
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
45 |
datatype mutual_info = Mutual of |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
46 |
{n : int, |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
47 |
n' : int, |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
48 |
fsum_var : string * typ, |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
49 |
|
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
50 |
ST: typ, |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
51 |
RST: typ, |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
52 |
|
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
53 |
parts: mutual_part list, |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
54 |
fqgars: qgar list, |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
55 |
qglrs: ((string * typ) list * term list * term * term) list, |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
56 |
|
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
57 |
fsum : term option} |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
58 |
|
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
59 |
fun mutual_induct_Pnames n = |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
60 |
if n < 5 then fst (chop n ["P","Q","R","S"]) |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
61 |
else map (fn i => "P" ^ string_of_int i) (1 upto n) |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
62 |
|
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
63 |
fun get_part fname = |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
64 |
the o find_first (fn (MutualPart {fvar=(n,_), ...}) => n = fname) |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
65 |
|
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
66 |
(* FIXME *) |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
67 |
fun mk_prod_abs e (t1, t2) = |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
68 |
let |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
69 |
val bTs = rev (map snd e) |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
70 |
val T1 = fastype_of1 (bTs, t1) |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
71 |
val T2 = fastype_of1 (bTs, t2) |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
72 |
in |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
73 |
HOLogic.pair_const T1 T2 $ t1 $ t2 |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
74 |
end |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
75 |
|
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
76 |
fun analyze_eqs ctxt defname fs eqs = |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
77 |
let |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
78 |
val num = length fs |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
79 |
val fqgars = map (split_def ctxt (K true)) eqs |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
80 |
val arity_of = map (fn (fname,_,_,args,_) => (fname, length args)) fqgars |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
81 |
|> AList.lookup (op =) #> the |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
82 |
|
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
83 |
fun curried_types (fname, fT) = |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
84 |
let |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
85 |
val (caTs, uaTs) = chop (arity_of fname) (binder_types fT) |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
86 |
in |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
87 |
(caTs, uaTs ---> body_type fT) |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
88 |
end |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
89 |
|
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
90 |
val (caTss, resultTs) = split_list (map curried_types fs) |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
91 |
val argTs = map (foldr1 HOLogic.mk_prodT) caTss |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
92 |
|
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
93 |
val dresultTs = distinct (op =) resultTs |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
94 |
val n' = length dresultTs |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
95 |
|
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
96 |
val RST = Balanced_Tree.make (uncurry SumTree.mk_sumT) dresultTs |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
97 |
val ST = Balanced_Tree.make (uncurry SumTree.mk_sumT) argTs |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
98 |
|
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
99 |
val fsum_type = ST --> RST |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
100 |
|
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
101 |
val ([fsum_var_name], _) = Variable.add_fixes [ defname ^ "_sum" ] ctxt |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
102 |
val fsum_var = (fsum_var_name, fsum_type) |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
103 |
|
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
104 |
fun define (fvar as (n, _)) caTs resultT i = |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
105 |
let |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
106 |
val vars = map_index (fn (j,T) => Free ("x" ^ string_of_int j, T)) caTs (* FIXME: Bind xs properly *) |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
107 |
val i' = find_index (fn Ta => Ta = resultT) dresultTs + 1 |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
108 |
|
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
109 |
val f_exp = SumTree.mk_proj RST n' i' (Free fsum_var $ SumTree.mk_inj ST num i (foldr1 HOLogic.mk_prod vars)) |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
110 |
val def = Term.abstract_over (Free fsum_var, fold_rev lambda vars f_exp) |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
111 |
|
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
112 |
val rew = (n, fold_rev lambda vars f_exp) |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
113 |
in |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
114 |
(MutualPart {i=i, i'=i', fvar=fvar,cargTs=caTs,f_def=def,f=NONE,f_defthm=NONE}, rew) |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
115 |
end |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
116 |
|
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
117 |
val (parts, rews) = split_list (map4 define fs caTss resultTs (1 upto num)) |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
118 |
|
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
119 |
fun convert_eqs (f, qs, gs, args, rhs) = |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
120 |
let |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
121 |
val MutualPart {i, i', ...} = get_part f parts |
2781
542ff50555f5
updated to new Isabelle (> 9 May)
Christian Urban <urbanc@in.tum.de>
parents:
2745
diff
changeset
|
122 |
val rhs' = rhs |
542ff50555f5
updated to new Isabelle (> 9 May)
Christian Urban <urbanc@in.tum.de>
parents:
2745
diff
changeset
|
123 |
|> map_aterms (fn t as Free (n, _) => the_default t (AList.lookup (op =) rews n) | t => t) |
2665
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
124 |
in |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
125 |
(qs, gs, SumTree.mk_inj ST num i (foldr1 (mk_prod_abs qs) args), |
2781
542ff50555f5
updated to new Isabelle (> 9 May)
Christian Urban <urbanc@in.tum.de>
parents:
2745
diff
changeset
|
126 |
Envir.beta_norm (SumTree.mk_inj RST n' i' rhs')) |
2665
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
127 |
end |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
128 |
|
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
129 |
val qglrs = map convert_eqs fqgars |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
130 |
in |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
131 |
Mutual {n=num, n'=n', fsum_var=fsum_var, ST=ST, RST=RST, |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
132 |
parts=parts, fqgars=fqgars, qglrs=qglrs, fsum=NONE} |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
133 |
end |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
134 |
|
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
135 |
fun define_projections fixes mutual fsum lthy = |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
136 |
let |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
137 |
fun def ((MutualPart {i=i, i'=i', fvar=(fname, fT), cargTs, f_def, ...}), (_, mixfix)) lthy = |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
138 |
let |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
139 |
val ((f, (_, f_defthm)), lthy') = |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
140 |
Local_Theory.define |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
141 |
((Binding.name fname, mixfix), |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
142 |
((Binding.conceal (Binding.name (fname ^ "_def")), []), |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
143 |
Term.subst_bound (fsum, f_def))) lthy |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
144 |
in |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
145 |
(MutualPart {i=i, i'=i', fvar=(fname, fT), cargTs=cargTs, f_def=f_def, |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
146 |
f=SOME f, f_defthm=SOME f_defthm }, |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
147 |
lthy') |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
148 |
end |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
149 |
|
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
150 |
val Mutual { n, n', fsum_var, ST, RST, parts, fqgars, qglrs, ... } = mutual |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
151 |
val (parts', lthy') = fold_map def (parts ~~ fixes) lthy |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
152 |
in |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
153 |
(Mutual { n=n, n'=n', fsum_var=fsum_var, ST=ST, RST=RST, parts=parts', |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
154 |
fqgars=fqgars, qglrs=qglrs, fsum=SOME fsum }, |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
155 |
lthy') |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
156 |
end |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
157 |
|
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
158 |
fun in_context ctxt (f, pre_qs, pre_gs, pre_args, pre_rhs) F = |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
159 |
let |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
160 |
val thy = ProofContext.theory_of ctxt |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
161 |
|
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
162 |
val oqnames = map fst pre_qs |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
163 |
val (qs, _) = Variable.variant_fixes oqnames ctxt |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
164 |
|>> map2 (fn (_, T) => fn n => Free (n, T)) pre_qs |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
165 |
|
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
166 |
fun inst t = subst_bounds (rev qs, t) |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
167 |
val gs = map inst pre_gs |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
168 |
val args = map inst pre_args |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
169 |
val rhs = inst pre_rhs |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
170 |
|
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
171 |
val cqs = map (cterm_of thy) qs |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
172 |
val ags = map (Thm.assume o cterm_of thy) gs |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
173 |
|
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
174 |
val import = fold Thm.forall_elim cqs |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
175 |
#> fold Thm.elim_implies ags |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
176 |
|
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
177 |
val export = fold_rev (Thm.implies_intr o cprop_of) ags |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
178 |
#> fold_rev forall_intr_rename (oqnames ~~ cqs) |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
179 |
in |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
180 |
F ctxt (f, qs, gs, args, rhs) import export |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
181 |
end |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
182 |
|
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
183 |
fun recover_mutual_psimp all_orig_fdefs parts ctxt (fname, _, _, args, rhs) |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
184 |
import (export : thm -> thm) sum_psimp_eq = |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
185 |
let |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
186 |
val (MutualPart {f=SOME f, ...}) = get_part fname parts |
2974
b95a2065aa10
generated the partial eqvt-theorem for functions
Christian Urban <urbanc@in.tum.de>
parents:
2973
diff
changeset
|
187 |
|
2665
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
188 |
val psimp = import sum_psimp_eq |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
189 |
val (simp, restore_cond) = |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
190 |
case cprems_of psimp of |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
191 |
[] => (psimp, I) |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
192 |
| [cond] => (Thm.implies_elim psimp (Thm.assume cond), Thm.implies_intr cond) |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
193 |
| _ => raise General.Fail "Too many conditions" |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
194 |
in |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
195 |
Goal.prove ctxt [] [] |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
196 |
(HOLogic.Trueprop $ HOLogic.mk_eq (list_comb (f, args), rhs)) |
2982
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
197 |
(fn _ => (Local_Defs.unfold_tac ctxt all_orig_fdefs) |
2665
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
198 |
THEN EqSubst.eqsubst_tac ctxt [0] [simp] 1 |
2982
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
199 |
THEN (simp_tac (simpset_of ctxt)) 1) (* FIXME: global simpset?!! *) |
2981 | 200 |
|> restore_cond |
201 |
|> export |
|
202 |
end |
|
203 |
||
2982
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
204 |
val inl_perm = @{lemma "x = Inl y ==> Sum_Type.Projl (permute p x) = permute p (Sum_Type.Projl x)" by simp} |
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
205 |
val inr_perm = @{lemma "x = Inr y ==> Sum_Type.Projr (permute p x) = permute p (Sum_Type.Projr x)" by simp} |
2981 | 206 |
|
2982
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
207 |
fun recover_mutual_eqvt eqvt_thm all_orig_fdefs parts ctxt (fname, _, _, args, _) |
2981 | 208 |
import (export : thm -> thm) sum_psimp_eq = |
209 |
let |
|
210 |
val (MutualPart {f=SOME f, ...}) = get_part fname parts |
|
2982
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
211 |
|
2981 | 212 |
val psimp = import sum_psimp_eq |
2982
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
213 |
val (cond, simp, restore_cond) = |
2981 | 214 |
case cprems_of psimp of |
2982
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
215 |
[] => ([], psimp, I) |
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
216 |
| [cond] => ([Thm.assume cond], Thm.implies_elim psimp (Thm.assume cond), Thm.implies_intr cond) |
2981 | 217 |
| _ => raise General.Fail "Too many conditions" |
218 |
||
219 |
val ([p], ctxt') = Variable.variant_fixes ["p"] ctxt |
|
220 |
val p = Free (p, @{typ perm}) |
|
2982
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
221 |
val ss = HOL_basic_ss addsimps |
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
222 |
@{thms permute_sum.simps[symmetric] Pair_eqvt[symmetric]} @ |
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
223 |
@{thms Projr.simps Projl.simps} @ |
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
224 |
[(cond MRS eqvt_thm) RS @{thm sym}] @ |
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
225 |
[inl_perm, inr_perm, simp] |
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
226 |
val goal_lhs = mk_perm p (list_comb (f, args)) |
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
227 |
val goal_rhs = list_comb (f, map (mk_perm p) args) |
2981 | 228 |
in |
2982
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
229 |
Goal.prove ctxt' [] [] (HOLogic.Trueprop $ HOLogic.mk_eq (goal_lhs, goal_rhs)) |
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
230 |
(fn _ => (Local_Defs.unfold_tac ctxt all_orig_fdefs) |
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
231 |
THEN (asm_full_simp_tac ss 1)) |
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
232 |
|> singleton (ProofContext.export ctxt' ctxt) |
2665
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
233 |
|> restore_cond |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
234 |
|> export |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
235 |
end |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
236 |
|
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
237 |
fun mk_applied_form ctxt caTs thm = |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
238 |
let |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
239 |
val thy = ProofContext.theory_of ctxt |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
240 |
val xs = map_index (fn (i,T) => cterm_of thy (Free ("x" ^ string_of_int i, T))) caTs (* FIXME: Bind xs properly *) |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
241 |
in |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
242 |
fold (fn x => fn thm => Thm.combination thm (Thm.reflexive x)) xs thm |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
243 |
|> Conv.fconv_rule (Thm.beta_conversion true) |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
244 |
|> fold_rev Thm.forall_intr xs |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
245 |
|> Thm.forall_elim_vars 0 |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
246 |
end |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
247 |
|
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
248 |
fun mutual_induct_rules lthy induct all_f_defs (Mutual {n, ST, parts, ...}) = |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
249 |
let |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
250 |
val cert = cterm_of (ProofContext.theory_of lthy) |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
251 |
val newPs = |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
252 |
map2 (fn Pname => fn MutualPart {cargTs, ...} => |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
253 |
Free (Pname, cargTs ---> HOLogic.boolT)) |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
254 |
(mutual_induct_Pnames (length parts)) parts |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
255 |
|
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
256 |
fun mk_P (MutualPart {cargTs, ...}) P = |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
257 |
let |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
258 |
val avars = map_index (fn (i,T) => Var (("a", i), T)) cargTs |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
259 |
val atup = foldr1 HOLogic.mk_prod avars |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
260 |
in |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
261 |
HOLogic.tupled_lambda atup (list_comb (P, avars)) |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
262 |
end |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
263 |
|
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
264 |
val Ps = map2 mk_P parts newPs |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
265 |
val case_exp = SumTree.mk_sumcases HOLogic.boolT Ps |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
266 |
|
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
267 |
val induct_inst = |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
268 |
Thm.forall_elim (cert case_exp) induct |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
269 |
|> full_simplify SumTree.sumcase_split_ss |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
270 |
|> full_simplify (HOL_basic_ss addsimps all_f_defs) |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
271 |
|
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
272 |
fun project rule (MutualPart {cargTs, i, ...}) k = |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
273 |
let |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
274 |
val afs = map_index (fn (j,T) => Free ("a" ^ string_of_int (j + k), T)) cargTs (* FIXME! *) |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
275 |
val inj = SumTree.mk_inj ST n i (foldr1 HOLogic.mk_prod afs) |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
276 |
in |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
277 |
(rule |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
278 |
|> Thm.forall_elim (cert inj) |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
279 |
|> full_simplify SumTree.sumcase_split_ss |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
280 |
|> fold_rev (Thm.forall_intr o cert) (afs @ newPs), |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
281 |
k + length cargTs) |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
282 |
end |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
283 |
in |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
284 |
fst (fold_map (project induct_inst) parts 0) |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
285 |
end |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
286 |
|
2982
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
287 |
|
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
288 |
fun forall_elim s (Const ("all", _) $ Abs (_, _, t)) = subst_bound (s, t) |
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
289 |
| forall_elim _ t = t |
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
290 |
|
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
291 |
val forall_elim_list = fold forall_elim |
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
292 |
|
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
293 |
fun split_conj_thm th = |
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
294 |
(split_conj_thm (th RS conjunct1)) @ (split_conj_thm (th RS conjunct2)) handle THM _ => [th]; |
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
295 |
|
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
296 |
fun prove_eqvt ctxt fs argTss eqvts_thms induct_thms = |
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
297 |
let |
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
298 |
fun aux argTs s = argTs |
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
299 |
|> map (pair s) |
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
300 |
|> Variable.variant_frees ctxt fs |
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
301 |
val argss' = map2 aux argTss (Name.invent (Variable.names_of ctxt) "" (length fs)) |
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
302 |
val argss = (map o map) Free argss' |
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
303 |
val arg_namess = (map o map) fst argss' |
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
304 |
val insts = (map o map) SOME arg_namess |
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
305 |
|
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
306 |
val ([p_name], ctxt') = Variable.variant_fixes ["p"] ctxt |
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
307 |
val p = Free (p_name, @{typ perm}) |
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
308 |
|
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
309 |
val acc_prems = |
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
310 |
map prop_of induct_thms |
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
311 |
|> map2 forall_elim_list argss |
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
312 |
|> map (strip_qnt_body "all") |
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
313 |
|> map (curry Logic.nth_prem 1) |
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
314 |
|> map HOLogic.dest_Trueprop |
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
315 |
|
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
316 |
fun mk_goal acc_prem (f, args) = |
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
317 |
let |
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
318 |
val goal_lhs = mk_perm p (list_comb (f, args)) |
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
319 |
val goal_rhs = list_comb (f, map (mk_perm p) args) |
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
320 |
in |
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
321 |
HOLogic.mk_imp (acc_prem, HOLogic.mk_eq (goal_lhs, goal_rhs)) |
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
322 |
end |
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
323 |
|
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
324 |
val goal = fold_conj_balanced (map2 mk_goal acc_prems (fs ~~ argss)) |
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
325 |
|> HOLogic.mk_Trueprop |
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
326 |
|
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
327 |
val induct_thm = case induct_thms of |
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
328 |
[thm] => thm |
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
329 |
|> Drule.gen_all |
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
330 |
|> Thm.permute_prems 0 1 |
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
331 |
|> (fn thm => atomize_rule (length (prems_of thm) - 1) thm) |
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
332 |
| thms => thms |
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
333 |
|> map Drule.gen_all |
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
334 |
|> map (Rule_Cases.add_consumes 1) |
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
335 |
|> snd o Rule_Cases.strict_mutual_rule ctxt' |
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
336 |
|> atomize_concl |
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
337 |
|
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
338 |
fun tac thm = rtac (Drule.gen_all thm) THEN_ALL_NEW atac |
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
339 |
in |
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
340 |
Goal.prove ctxt' (flat arg_namess) [] goal |
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
341 |
(fn {context, ...} => HEADGOAL (DETERM o (rtac induct_thm) THEN' RANGE (map tac eqvts_thms))) |
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
342 |
|> singleton (ProofContext.export ctxt' ctxt) |
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
343 |
|> split_conj_thm |
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
344 |
|> map (fn th => th RS mp) |
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
345 |
end |
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
346 |
|
2665
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
347 |
fun mk_partial_rules_mutual lthy inner_cont (m as Mutual {parts, fqgars, ...}) proof = |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
348 |
let |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
349 |
val result = inner_cont proof |
2973
d1038e67923a
added a flag (eqvt) to termination proofs arising fron nominal_primrecs
Christian Urban <urbanc@in.tum.de>
parents:
2821
diff
changeset
|
350 |
val NominalFunctionResult {G, R, cases, psimps, simple_pinducts=[simple_pinduct], |
2974
b95a2065aa10
generated the partial eqvt-theorem for functions
Christian Urban <urbanc@in.tum.de>
parents:
2973
diff
changeset
|
351 |
termination, domintros, eqvts=[eqvt],...} = result |
2665
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
352 |
|
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
353 |
val (all_f_defs, fs) = |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
354 |
map (fn MutualPart {f_defthm = SOME f_def, f = SOME f, cargTs, ...} => |
2982
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
355 |
(mk_applied_form lthy cargTs (Thm.symmetric f_def), f)) |
2665
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
356 |
parts |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
357 |
|> split_list |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
358 |
|
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
359 |
val all_orig_fdefs = |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
360 |
map (fn MutualPart {f_defthm = SOME f_def, ...} => f_def) parts |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
361 |
|
2982
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
362 |
val cargTss = |
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
363 |
map (fn MutualPart {f = SOME f, cargTs, ...} => cargTs) parts |
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
364 |
|
2665
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
365 |
fun mk_mpsimp fqgar sum_psimp = |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
366 |
in_context lthy fqgar (recover_mutual_psimp all_orig_fdefs parts) sum_psimp |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
367 |
|
2981 | 368 |
fun mk_meqvts fqgar sum_psimp = |
369 |
in_context lthy fqgar (recover_mutual_eqvt eqvt all_orig_fdefs parts) sum_psimp |
|
370 |
||
2665
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
371 |
val rew_ss = HOL_basic_ss addsimps all_f_defs |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
372 |
val mpsimps = map2 mk_mpsimp fqgars psimps |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
373 |
val minducts = mutual_induct_rules lthy simple_pinduct all_f_defs m |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
374 |
val mtermination = full_simplify rew_ss termination |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
375 |
val mdomintros = Option.map (map (full_simplify rew_ss)) domintros |
2981 | 376 |
val meqvts = map2 mk_meqvts fqgars psimps |
2982
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
377 |
val meqvt_funs = prove_eqvt lthy fs cargTss meqvts minducts |
2974
b95a2065aa10
generated the partial eqvt-theorem for functions
Christian Urban <urbanc@in.tum.de>
parents:
2973
diff
changeset
|
378 |
in |
2973
d1038e67923a
added a flag (eqvt) to termination proofs arising fron nominal_primrecs
Christian Urban <urbanc@in.tum.de>
parents:
2821
diff
changeset
|
379 |
NominalFunctionResult { fs=fs, G=G, R=R, |
2665
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
380 |
psimps=mpsimps, simple_pinducts=minducts, |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
381 |
cases=cases, termination=mtermination, |
2982
4a00077c008f
completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents:
2981
diff
changeset
|
382 |
domintros=mdomintros, eqvts=meqvt_funs } |
2665
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
383 |
end |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
384 |
|
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
385 |
(* nominal *) |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
386 |
fun prepare_nominal_function_mutual config defname fixes eqss lthy = |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
387 |
let |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
388 |
val mutual as Mutual {fsum_var=(n, T), qglrs, ...} = |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
389 |
analyze_eqs lthy defname (map fst fixes) (map Envir.beta_eta_contract eqss) |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
390 |
|
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
391 |
val ((fsum, goalstate, cont), lthy') = |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
392 |
Nominal_Function_Core.prepare_nominal_function config defname [((n, T), NoSyn)] qglrs lthy |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
393 |
|
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
394 |
val (mutual', lthy'') = define_projections fixes mutual fsum lthy' |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
395 |
|
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
396 |
val mutual_cont = mk_partial_rules_mutual lthy'' cont mutual' |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
397 |
in |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
398 |
((goalstate, mutual_cont), lthy'') |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
399 |
end |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
400 |
|
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
401 |
end |