author | Christian Urban <christian dot urban at kcl dot ac dot uk> |
Sat, 19 Mar 2016 21:06:48 +0000 | |
branch | Nominal2-Isabelle2016 |
changeset 3243 | c4f31f1564b7 |
parent 3239 | 67370521c09c |
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 |
|
3197
25d11b449e92
definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents:
3045
diff
changeset
|
7 |
Joachim Breitner helped with the auxiliary graph |
25d11b449e92
definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents:
3045
diff
changeset
|
8 |
definitions (7 August 2012) |
2665
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 |
Mutual recursive nominal function definitions. |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
11 |
*) |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
12 |
|
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
|
13 |
|
2665
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
14 |
signature NOMINAL_FUNCTION_MUTUAL = |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
15 |
sig |
2819
4bd584ff4fab
added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
2781
diff
changeset
|
16 |
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
|
17 |
-> string (* defname *) |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
18 |
-> ((string * typ) * mixfix) list |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
19 |
-> term list |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
20 |
-> local_theory |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
21 |
-> ((thm (* goalstate *) |
3243
c4f31f1564b7
updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3239
diff
changeset
|
22 |
* (Proof.context -> 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
|
23 |
) * local_theory) |
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 |
structure Nominal_Function_Mutual: NOMINAL_FUNCTION_MUTUAL = |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
27 |
struct |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
28 |
|
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
29 |
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
|
30 |
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
|
31 |
open Nominal_Function_Common |
2665
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
32 |
|
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
33 |
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
|
34 |
|
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
35 |
datatype mutual_part = MutualPart of |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
36 |
{i : int, |
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 |
fvar : string * typ, |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
39 |
cargTs: typ list, |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
40 |
f_def: term, |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
41 |
f: term option, |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
42 |
f_defthm : thm option} |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
43 |
|
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
44 |
datatype mutual_info = Mutual of |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
45 |
{n : int, |
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 |
fsum_var : string * typ, |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
48 |
|
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
49 |
ST: typ, |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
50 |
RST: typ, |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
51 |
|
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
52 |
parts: mutual_part list, |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
53 |
fqgars: qgar list, |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
54 |
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
|
55 |
|
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
56 |
fsum : term option} |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
57 |
|
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
58 |
fun mutual_induct_Pnames n = |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
59 |
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
|
60 |
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
|
61 |
|
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
62 |
fun get_part fname = |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
63 |
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
|
64 |
|
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
65 |
(* FIXME *) |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
66 |
fun mk_prod_abs e (t1, t2) = |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
67 |
let |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
68 |
val bTs = rev (map snd e) |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
69 |
val T1 = fastype_of1 (bTs, t1) |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
70 |
val T2 = fastype_of1 (bTs, t2) |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
71 |
in |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
72 |
HOLogic.pair_const T1 T2 $ t1 $ t2 |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
73 |
end |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
74 |
|
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
75 |
fun analyze_eqs ctxt defname fs eqs = |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
76 |
let |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
77 |
val num = length fs |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
78 |
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
|
79 |
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
|
80 |
|> AList.lookup (op =) #> the |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
81 |
|
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
82 |
fun curried_types (fname, fT) = |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
83 |
let |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
84 |
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
|
85 |
in |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
86 |
(caTs, uaTs ---> body_type fT) |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
87 |
end |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
88 |
|
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
89 |
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
|
90 |
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
|
91 |
|
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
92 |
val dresultTs = distinct (op =) resultTs |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
93 |
val n' = length dresultTs |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
94 |
|
3229
b52e8651591f
updated to Isabelle changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3227
diff
changeset
|
95 |
val RST = Balanced_Tree.make (uncurry Sum_Tree.mk_sumT) dresultTs |
b52e8651591f
updated to Isabelle changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3227
diff
changeset
|
96 |
val ST = Balanced_Tree.make (uncurry Sum_Tree.mk_sumT) argTs |
2665
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
97 |
|
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
98 |
val fsum_type = ST --> RST |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
99 |
|
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
100 |
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
|
101 |
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
|
102 |
|
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
103 |
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
|
104 |
let |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
105 |
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
|
106 |
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
|
107 |
|
3229
b52e8651591f
updated to Isabelle changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3227
diff
changeset
|
108 |
val f_exp = Sum_Tree.mk_proj RST n' i' (Free fsum_var $ Sum_Tree.mk_inj ST num i (foldr1 HOLogic.mk_prod vars)) |
2665
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
109 |
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
|
110 |
|
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
111 |
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
|
112 |
in |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
113 |
(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
|
114 |
end |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
115 |
|
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
116 |
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
|
117 |
|
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
118 |
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
|
119 |
let |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
120 |
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
|
121 |
val rhs' = rhs |
542ff50555f5
updated to new Isabelle (> 9 May)
Christian Urban <urbanc@in.tum.de>
parents:
2745
diff
changeset
|
122 |
|> 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
|
123 |
in |
3229
b52e8651591f
updated to Isabelle changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3227
diff
changeset
|
124 |
(qs, gs, Sum_Tree.mk_inj ST num i (foldr1 (mk_prod_abs qs) args), |
b52e8651591f
updated to Isabelle changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3227
diff
changeset
|
125 |
Envir.beta_norm (Sum_Tree.mk_inj RST n' i' rhs')) |
2665
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
126 |
end |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
127 |
|
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
128 |
val qglrs = map convert_eqs fqgars |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
129 |
in |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
130 |
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
|
131 |
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
|
132 |
end |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
133 |
|
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
134 |
fun define_projections fixes mutual fsum lthy = |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
135 |
let |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
136 |
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
|
137 |
let |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
138 |
val ((f, (_, f_defthm)), lthy') = |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
139 |
Local_Theory.define |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
140 |
((Binding.name fname, mixfix), |
3239
67370521c09c
updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3231
diff
changeset
|
141 |
((Binding.concealed (Binding.name (fname ^ "_def")), []), |
2665
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
142 |
Term.subst_bound (fsum, f_def))) lthy |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
143 |
in |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
144 |
(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
|
145 |
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
|
146 |
lthy') |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
147 |
end |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
148 |
|
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
149 |
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
|
150 |
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
|
151 |
in |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
152 |
(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
|
153 |
fqgars=fqgars, qglrs=qglrs, fsum=SOME fsum }, |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
154 |
lthy') |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
155 |
end |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
156 |
|
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
157 |
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
|
158 |
let |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
159 |
val oqnames = map fst pre_qs |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
160 |
val (qs, _) = Variable.variant_fixes oqnames ctxt |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
161 |
|>> 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
|
162 |
|
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
163 |
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
|
164 |
val gs = map inst pre_gs |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
165 |
val args = map inst pre_args |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
166 |
val rhs = inst pre_rhs |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
167 |
|
3239
67370521c09c
updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3231
diff
changeset
|
168 |
val cqs = map (Thm.cterm_of ctxt) qs |
67370521c09c
updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3231
diff
changeset
|
169 |
val (ags, ctxt') = fold_map Thm.assume_hyps (map (Thm.cterm_of ctxt) gs) ctxt |
2665
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 import = fold Thm.forall_elim cqs |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
172 |
#> fold Thm.elim_implies ags |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
173 |
|
3239
67370521c09c
updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3231
diff
changeset
|
174 |
val export = fold_rev (Thm.implies_intr o Thm.cprop_of) ags |
2665
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
175 |
#> fold_rev forall_intr_rename (oqnames ~~ cqs) |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
176 |
in |
3227
35bb5b013f0e
updated with current AFP version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3226
diff
changeset
|
177 |
F ctxt' (f, qs, gs, args, rhs) import export |
2665
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
178 |
end |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
179 |
|
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
180 |
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
|
181 |
import (export : thm -> thm) sum_psimp_eq = |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
182 |
let |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
183 |
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
|
184 |
|
2665
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
185 |
val psimp = import sum_psimp_eq |
3227
35bb5b013f0e
updated with current AFP version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3226
diff
changeset
|
186 |
val ((simp, restore_cond), ctxt') = |
2665
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
187 |
case cprems_of psimp of |
3227
35bb5b013f0e
updated with current AFP version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3226
diff
changeset
|
188 |
[] => ((psimp, I), ctxt) |
35bb5b013f0e
updated with current AFP version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3226
diff
changeset
|
189 |
| [cond] => |
35bb5b013f0e
updated with current AFP version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3226
diff
changeset
|
190 |
let val (asm, ctxt') = Thm.assume_hyps cond ctxt |
35bb5b013f0e
updated with current AFP version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3226
diff
changeset
|
191 |
in ((Thm.implies_elim psimp asm, Thm.implies_intr cond), ctxt') end |
2665
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
192 |
| _ => raise General.Fail "Too many conditions" |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
193 |
in |
3227
35bb5b013f0e
updated with current AFP version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3226
diff
changeset
|
194 |
Goal.prove ctxt' [] [] |
2665
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
195 |
(HOLogic.Trueprop $ HOLogic.mk_eq (list_comb (f, args), rhs)) |
3227
35bb5b013f0e
updated with current AFP version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3226
diff
changeset
|
196 |
(fn _ => (Local_Defs.unfold_tac ctxt' all_orig_fdefs) |
35bb5b013f0e
updated with current AFP version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3226
diff
changeset
|
197 |
THEN EqSubst.eqsubst_tac ctxt' [0] [simp] 1 |
35bb5b013f0e
updated with current AFP version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3226
diff
changeset
|
198 |
THEN (simp_tac ctxt') 1) |
2981 | 199 |
|> restore_cond |
200 |
|> export |
|
201 |
end |
|
202 |
||
3229
b52e8651591f
updated to Isabelle changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3227
diff
changeset
|
203 |
val inl_perm = @{lemma "x = Inl y ==> projl (permute p x) = permute p (projl x)" by simp} |
b52e8651591f
updated to Isabelle changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3227
diff
changeset
|
204 |
val inr_perm = @{lemma "x = Inr y ==> projr (permute p x) = permute p (projr x)" by simp} |
2981 | 205 |
|
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
|
206 |
fun recover_mutual_eqvt eqvt_thm all_orig_fdefs parts ctxt (fname, _, _, args, _) |
2981 | 207 |
import (export : thm -> thm) sum_psimp_eq = |
208 |
let |
|
209 |
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
|
210 |
|
2981 | 211 |
val psimp = import sum_psimp_eq |
3227
35bb5b013f0e
updated with current AFP version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3226
diff
changeset
|
212 |
val ((cond, simp, restore_cond), ctxt') = |
2981 | 213 |
case cprems_of psimp of |
3227
35bb5b013f0e
updated with current AFP version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3226
diff
changeset
|
214 |
[] => (([], psimp, I), ctxt) |
35bb5b013f0e
updated with current AFP version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3226
diff
changeset
|
215 |
| [cond] => |
35bb5b013f0e
updated with current AFP version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3226
diff
changeset
|
216 |
let val (asm, ctxt') = Thm.assume_hyps cond ctxt |
35bb5b013f0e
updated with current AFP version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3226
diff
changeset
|
217 |
in (([asm], Thm.implies_elim psimp asm, Thm.implies_intr cond), ctxt') end |
2981 | 218 |
| _ => raise General.Fail "Too many conditions" |
219 |
||
3227
35bb5b013f0e
updated with current AFP version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3226
diff
changeset
|
220 |
val ([p], ctxt'') = ctxt' |
3204
b69c8660de14
fixed problem with not fresh enough permutation name in nominal_primrec
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3197
diff
changeset
|
221 |
|> fold Variable.declare_term args |
3229
b52e8651591f
updated to Isabelle changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3227
diff
changeset
|
222 |
|> Variable.variant_fixes ["p"] |
2981 | 223 |
val p = Free (p, @{typ perm}) |
3204
b69c8660de14
fixed problem with not fresh enough permutation name in nominal_primrec
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3197
diff
changeset
|
224 |
|
3218
89158f401b07
updated to simplifier changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3204
diff
changeset
|
225 |
val simpset = |
3227
35bb5b013f0e
updated with current AFP version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3226
diff
changeset
|
226 |
put_simpset HOL_basic_ss ctxt'' addsimps |
3229
b52e8651591f
updated to Isabelle changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3227
diff
changeset
|
227 |
@{thms permute_sum.simps[symmetric] Pair_eqvt[symmetric] sum.sel} @ |
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
|
228 |
[(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
|
229 |
[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
|
230 |
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
|
231 |
val goal_rhs = list_comb (f, map (mk_perm p) args) |
2981 | 232 |
in |
3227
35bb5b013f0e
updated with current AFP version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3226
diff
changeset
|
233 |
Goal.prove ctxt'' [] [] (HOLogic.Trueprop $ HOLogic.mk_eq (goal_lhs, goal_rhs)) |
35bb5b013f0e
updated with current AFP version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3226
diff
changeset
|
234 |
(fn _ => Local_Defs.unfold_tac ctxt'' all_orig_fdefs |
3218
89158f401b07
updated to simplifier changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3204
diff
changeset
|
235 |
THEN (asm_full_simp_tac simpset 1)) |
3227
35bb5b013f0e
updated with current AFP version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3226
diff
changeset
|
236 |
|> singleton (Proof_Context.export ctxt'' ctxt) |
2665
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
237 |
|> restore_cond |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
238 |
|> export |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
239 |
end |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
240 |
|
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
241 |
fun mk_applied_form ctxt caTs thm = |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
242 |
let |
3239
67370521c09c
updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3231
diff
changeset
|
243 |
val xs = map_index (fn (i,T) => Thm.cterm_of ctxt (Free ("x" ^ string_of_int i, T))) caTs (* FIXME: Bind xs properly *) |
2665
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
244 |
in |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
245 |
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
|
246 |
|> Conv.fconv_rule (Thm.beta_conversion true) |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
247 |
|> fold_rev Thm.forall_intr xs |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
248 |
|> Thm.forall_elim_vars 0 |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
249 |
end |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
250 |
|
3218
89158f401b07
updated to simplifier changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3204
diff
changeset
|
251 |
fun mutual_induct_rules ctxt induct all_f_defs (Mutual {n, ST, parts, ...}) = |
2665
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
252 |
let |
3239
67370521c09c
updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3231
diff
changeset
|
253 |
val cert = Thm.cterm_of ctxt |
2665
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
254 |
val newPs = |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
255 |
map2 (fn Pname => fn MutualPart {cargTs, ...} => |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
256 |
Free (Pname, cargTs ---> HOLogic.boolT)) |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
257 |
(mutual_induct_Pnames (length parts)) parts |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
258 |
|
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
259 |
fun mk_P (MutualPart {cargTs, ...}) P = |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
260 |
let |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
261 |
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
|
262 |
val atup = foldr1 HOLogic.mk_prod avars |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
263 |
in |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
264 |
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
|
265 |
end |
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 Ps = map2 mk_P parts newPs |
3229
b52e8651591f
updated to Isabelle changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3227
diff
changeset
|
268 |
val case_exp = Sum_Tree.mk_sumcases HOLogic.boolT Ps |
2665
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
269 |
|
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
270 |
val induct_inst = |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
271 |
Thm.forall_elim (cert case_exp) induct |
3229
b52e8651591f
updated to Isabelle changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3227
diff
changeset
|
272 |
|> full_simplify (put_simpset Sum_Tree.sumcase_split_ss ctxt) |
3218
89158f401b07
updated to simplifier changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3204
diff
changeset
|
273 |
|> full_simplify (put_simpset HOL_basic_ss ctxt addsimps all_f_defs) |
2665
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
274 |
|
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
275 |
fun project rule (MutualPart {cargTs, i, ...}) k = |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
276 |
let |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
277 |
val afs = map_index (fn (j,T) => Free ("a" ^ string_of_int (j + k), T)) cargTs (* FIXME! *) |
3229
b52e8651591f
updated to Isabelle changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3227
diff
changeset
|
278 |
val inj = Sum_Tree.mk_inj ST n i (foldr1 HOLogic.mk_prod afs) |
2665
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
279 |
in |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
280 |
(rule |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
281 |
|> Thm.forall_elim (cert inj) |
3229
b52e8651591f
updated to Isabelle changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3227
diff
changeset
|
282 |
|> full_simplify (put_simpset Sum_Tree.sumcase_split_ss ctxt) |
2665
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
283 |
|> 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
|
284 |
k + length cargTs) |
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 |
in |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
287 |
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
|
288 |
end |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
289 |
|
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
|
290 |
|
3231
188826f1ccdb
updated to massive changes in Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3229
diff
changeset
|
291 |
fun forall_elim s (Const (@{const_name Pure.all}, _) $ Abs (_, _, t)) = subst_bound (s, t) |
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
|
292 |
| 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
|
293 |
|
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 |
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
|
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 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
|
297 |
(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
|
298 |
|
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 |
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
|
300 |
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
|
301 |
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
|
302 |
|> 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
|
303 |
|> 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
|
304 |
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
|
305 |
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
|
306 |
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
|
307 |
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
|
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 ([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
|
310 |
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
|
311 |
|
2983 | 312 |
(* extracting the acc-premises from the induction theorems *) |
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
|
313 |
val acc_prems = |
3239
67370521c09c
updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3231
diff
changeset
|
314 |
map Thm.prop_of induct_thms |
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
|
315 |
|> map2 forall_elim_list argss |
3231
188826f1ccdb
updated to massive changes in Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3229
diff
changeset
|
316 |
|> map (strip_qnt_body @{const_name Pure.all}) |
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
|
317 |
|> 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
|
318 |
|> 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
|
319 |
|
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 |
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
|
321 |
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
|
322 |
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
|
323 |
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
|
324 |
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
|
325 |
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
|
326 |
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
|
327 |
|
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 |
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
|
329 |
|> 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
|
330 |
|
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 |
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
|
332 |
[thm] => thm |
3243
c4f31f1564b7
updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3239
diff
changeset
|
333 |
|> Variable.gen_all ctxt' |
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
|
334 |
|> Thm.permute_prems 0 1 |
3239
67370521c09c
updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3231
diff
changeset
|
335 |
|> (fn thm => atomize_rule ctxt' (length (Thm.prems_of thm) - 1) thm) |
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
|
336 |
| thms => thms |
3243
c4f31f1564b7
updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3239
diff
changeset
|
337 |
|> map (Variable.gen_all ctxt') |
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
|
338 |
|> 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
|
339 |
|> snd o Rule_Cases.strict_mutual_rule ctxt' |
3226
780b7a2c50b6
updated to changes in Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3219
diff
changeset
|
340 |
|> atomize_concl ctxt' |
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
|
341 |
|
3239
67370521c09c
updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3231
diff
changeset
|
342 |
fun tac ctxt thm = |
3243
c4f31f1564b7
updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3239
diff
changeset
|
343 |
resolve_tac ctxt [Variable.gen_all ctxt thm] |
c4f31f1564b7
updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3239
diff
changeset
|
344 |
THEN_ALL_NEW assume_tac ctxt |
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
|
345 |
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
|
346 |
Goal.prove ctxt' (flat arg_namess) [] goal |
3239
67370521c09c
updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3231
diff
changeset
|
347 |
(fn {context = ctxt'', ...} => |
3243
c4f31f1564b7
updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3239
diff
changeset
|
348 |
HEADGOAL (DETERM o (resolve_tac ctxt'' [induct_thm]) THEN' |
c4f31f1564b7
updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3239
diff
changeset
|
349 |
RANGE (map (tac ctxt'') eqvts_thms))) |
3045
d0ad264f8c4f
updated to Isabelle 3 Nov; it includes a hack to work around a bug in the localised version of the quotient package
Christian Urban <urbanc@in.tum.de>
parents:
2983
diff
changeset
|
350 |
|> singleton (Proof_Context.export ctxt' ctxt) |
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
|
351 |
|> 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
|
352 |
|> 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
|
353 |
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
|
354 |
|
3218
89158f401b07
updated to simplifier changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3204
diff
changeset
|
355 |
fun mk_partial_rules_mutual ctxt inner_cont (m as Mutual {parts, fqgars, ...}) proof = |
2665
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
356 |
let |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
357 |
val result = inner_cont proof |
3204
b69c8660de14
fixed problem with not fresh enough permutation name in nominal_primrec
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3197
diff
changeset
|
358 |
|
2973
d1038e67923a
added a flag (eqvt) to termination proofs arising fron nominal_primrecs
Christian Urban <urbanc@in.tum.de>
parents:
2821
diff
changeset
|
359 |
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
|
360 |
termination, domintros, eqvts=[eqvt],...} = result |
2665
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
361 |
|
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
362 |
val (all_f_defs, fs) = |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
363 |
map (fn MutualPart {f_defthm = SOME f_def, f = SOME f, cargTs, ...} => |
3218
89158f401b07
updated to simplifier changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3204
diff
changeset
|
364 |
(mk_applied_form ctxt cargTs (Thm.symmetric f_def), f)) |
2665
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
365 |
parts |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
366 |
|> split_list |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
367 |
|
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
368 |
val all_orig_fdefs = |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
369 |
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
|
370 |
|
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
|
371 |
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
|
372 |
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
|
373 |
|
2665
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
374 |
fun mk_mpsimp fqgar sum_psimp = |
3218
89158f401b07
updated to simplifier changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3204
diff
changeset
|
375 |
in_context ctxt fqgar (recover_mutual_psimp all_orig_fdefs parts) sum_psimp |
2665
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
376 |
|
2981 | 377 |
fun mk_meqvts fqgar sum_psimp = |
3218
89158f401b07
updated to simplifier changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3204
diff
changeset
|
378 |
in_context ctxt fqgar (recover_mutual_eqvt eqvt all_orig_fdefs parts) sum_psimp |
2981 | 379 |
|
3218
89158f401b07
updated to simplifier changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3204
diff
changeset
|
380 |
val rew_simpset = put_simpset HOL_basic_ss ctxt addsimps all_f_defs |
2665
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
381 |
val mpsimps = map2 mk_mpsimp fqgars psimps |
3218
89158f401b07
updated to simplifier changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3204
diff
changeset
|
382 |
val minducts = mutual_induct_rules ctxt simple_pinduct all_f_defs m |
89158f401b07
updated to simplifier changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3204
diff
changeset
|
383 |
val mtermination = full_simplify rew_simpset termination |
89158f401b07
updated to simplifier changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3204
diff
changeset
|
384 |
val mdomintros = Option.map (map (full_simplify rew_simpset)) domintros |
2981 | 385 |
val meqvts = map2 mk_meqvts fqgars psimps |
3218
89158f401b07
updated to simplifier changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3204
diff
changeset
|
386 |
val meqvt_funs = prove_eqvt ctxt fs cargTss meqvts minducts |
2974
b95a2065aa10
generated the partial eqvt-theorem for functions
Christian Urban <urbanc@in.tum.de>
parents:
2973
diff
changeset
|
387 |
in |
2973
d1038e67923a
added a flag (eqvt) to termination proofs arising fron nominal_primrecs
Christian Urban <urbanc@in.tum.de>
parents:
2821
diff
changeset
|
388 |
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
|
389 |
psimps=mpsimps, simple_pinducts=minducts, |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
390 |
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
|
391 |
domintros=mdomintros, eqvts=meqvt_funs } |
2665
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
392 |
end |
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 |
(* nominal *) |
3197
25d11b449e92
definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents:
3045
diff
changeset
|
395 |
fun subst_all s (Q $ Abs(_, _, t)) = |
25d11b449e92
definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents:
3045
diff
changeset
|
396 |
let |
25d11b449e92
definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents:
3045
diff
changeset
|
397 |
val vs = map Free (Term.add_frees s []) |
25d11b449e92
definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents:
3045
diff
changeset
|
398 |
in |
25d11b449e92
definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents:
3045
diff
changeset
|
399 |
fold Logic.all vs (subst_bound (s, t)) |
25d11b449e92
definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents:
3045
diff
changeset
|
400 |
end |
25d11b449e92
definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents:
3045
diff
changeset
|
401 |
|
25d11b449e92
definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents:
3045
diff
changeset
|
402 |
fun mk_comp_dummy t s = Const (@{const_name comp}, dummyT) $ t $ s |
25d11b449e92
definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents:
3045
diff
changeset
|
403 |
|
25d11b449e92
definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents:
3045
diff
changeset
|
404 |
fun all v t = |
25d11b449e92
definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents:
3045
diff
changeset
|
405 |
let |
25d11b449e92
definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents:
3045
diff
changeset
|
406 |
val T = Term.fastype_of v |
25d11b449e92
definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents:
3045
diff
changeset
|
407 |
in |
25d11b449e92
definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents:
3045
diff
changeset
|
408 |
Logic.all_const T $ absdummy T (abstract_over (v, t)) |
25d11b449e92
definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents:
3045
diff
changeset
|
409 |
end |
25d11b449e92
definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents:
3045
diff
changeset
|
410 |
|
25d11b449e92
definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents:
3045
diff
changeset
|
411 |
(* nominal *) |
2665
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
412 |
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
|
413 |
let |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
414 |
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
|
415 |
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
|
416 |
|
3197
25d11b449e92
definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents:
3045
diff
changeset
|
417 |
val ((fsum, G, GIntro_thms, G_induct, goalstate, cont), lthy') = |
2665
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
418 |
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
|
419 |
|
3197
25d11b449e92
definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents:
3045
diff
changeset
|
420 |
val (mutual' as Mutual {n', parts, ST, RST, ...}, lthy'') = define_projections fixes mutual fsum lthy' |
2665
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
421 |
|
3197
25d11b449e92
definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents:
3045
diff
changeset
|
422 |
(* defining the auxiliary graph *) |
25d11b449e92
definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents:
3045
diff
changeset
|
423 |
fun mk_cases (MutualPart {i', fvar as (n, T), ...}) = |
25d11b449e92
definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents:
3045
diff
changeset
|
424 |
let |
25d11b449e92
definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents:
3045
diff
changeset
|
425 |
val (tys, ty) = strip_type T |
25d11b449e92
definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents:
3045
diff
changeset
|
426 |
val fun_var = Free (n ^ "_aux", HOLogic.mk_tupleT tys --> ty) |
3229
b52e8651591f
updated to Isabelle changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3227
diff
changeset
|
427 |
val inj_fun = absdummy dummyT (Sum_Tree.mk_inj RST n' i' (Bound 0)) |
3197
25d11b449e92
definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents:
3045
diff
changeset
|
428 |
in |
25d11b449e92
definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents:
3045
diff
changeset
|
429 |
Syntax.check_term lthy'' (mk_comp_dummy inj_fun fun_var) |
25d11b449e92
definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents:
3045
diff
changeset
|
430 |
end |
25d11b449e92
definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents:
3045
diff
changeset
|
431 |
|
3229
b52e8651591f
updated to Isabelle changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3227
diff
changeset
|
432 |
val case_sum_exp = map mk_cases parts |
b52e8651591f
updated to Isabelle changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3227
diff
changeset
|
433 |
|> Sum_Tree.mk_sumcases RST |
3197
25d11b449e92
definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents:
3045
diff
changeset
|
434 |
|
25d11b449e92
definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents:
3045
diff
changeset
|
435 |
val (G_name, G_type) = dest_Free G |
25d11b449e92
definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents:
3045
diff
changeset
|
436 |
val G_name_aux = G_name ^ "_aux" |
25d11b449e92
definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents:
3045
diff
changeset
|
437 |
val subst = [(G, Free (G_name_aux, G_type))] |
25d11b449e92
definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents:
3045
diff
changeset
|
438 |
val GIntros_aux = GIntro_thms |
3239
67370521c09c
updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3231
diff
changeset
|
439 |
|> map Thm.prop_of |
3197
25d11b449e92
definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents:
3045
diff
changeset
|
440 |
|> map (Term.subst_free subst) |
3229
b52e8651591f
updated to Isabelle changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3227
diff
changeset
|
441 |
|> map (subst_all case_sum_exp) |
3197
25d11b449e92
definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents:
3045
diff
changeset
|
442 |
|
25d11b449e92
definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents:
3045
diff
changeset
|
443 |
val ((G_aux, GIntro_aux_thms, _, G_aux_induct), lthy''') = |
25d11b449e92
definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents:
3045
diff
changeset
|
444 |
Nominal_Function_Core.inductive_def ((Binding.name G_name_aux, G_type), NoSyn) GIntros_aux lthy'' |
25d11b449e92
definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents:
3045
diff
changeset
|
445 |
|
3243
c4f31f1564b7
updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3239
diff
changeset
|
446 |
fun mutual_cont ctxt = mk_partial_rules_mutual lthy''' (cont ctxt) mutual' |
3219
e5d9b6bca88c
updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3218
diff
changeset
|
447 |
|
3197
25d11b449e92
definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents:
3045
diff
changeset
|
448 |
(* proof of equivalence between graph and auxiliary graph *) |
25d11b449e92
definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents:
3045
diff
changeset
|
449 |
val x = Var(("x", 0), ST) |
25d11b449e92
definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents:
3045
diff
changeset
|
450 |
val y = Var(("y", 1), RST) |
25d11b449e92
definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents:
3045
diff
changeset
|
451 |
val G_aux_prem = HOLogic.mk_Trueprop (G_aux $ x $ y) |
25d11b449e92
definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents:
3045
diff
changeset
|
452 |
val G_prem = HOLogic.mk_Trueprop (G $ x $ y) |
25d11b449e92
definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents:
3045
diff
changeset
|
453 |
|
25d11b449e92
definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents:
3045
diff
changeset
|
454 |
fun mk_inj_goal (MutualPart {i', ...}) = |
25d11b449e92
definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents:
3045
diff
changeset
|
455 |
let |
3229
b52e8651591f
updated to Isabelle changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3227
diff
changeset
|
456 |
val injs = Sum_Tree.mk_inj ST n' i' (Bound 0) |
3197
25d11b449e92
definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents:
3045
diff
changeset
|
457 |
val projs = y |
3229
b52e8651591f
updated to Isabelle changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3227
diff
changeset
|
458 |
|> Sum_Tree.mk_proj RST n' i' |
b52e8651591f
updated to Isabelle changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3227
diff
changeset
|
459 |
|> Sum_Tree.mk_inj RST n' i' |
3197
25d11b449e92
definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents:
3045
diff
changeset
|
460 |
in |
25d11b449e92
definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents:
3045
diff
changeset
|
461 |
Const (@{const_name "All"}, dummyT) $ absdummy dummyT |
25d11b449e92
definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents:
3045
diff
changeset
|
462 |
(HOLogic.mk_imp (HOLogic.mk_eq(x, injs), HOLogic.mk_eq(projs, y))) |
25d11b449e92
definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents:
3045
diff
changeset
|
463 |
end |
25d11b449e92
definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents:
3045
diff
changeset
|
464 |
|
25d11b449e92
definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents:
3045
diff
changeset
|
465 |
val goal_inj = Logic.mk_implies (G_aux_prem, |
25d11b449e92
definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents:
3045
diff
changeset
|
466 |
HOLogic.mk_Trueprop (fold_conj (map mk_inj_goal parts))) |
25d11b449e92
definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents:
3045
diff
changeset
|
467 |
|> all x |> all y |
25d11b449e92
definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents:
3045
diff
changeset
|
468 |
|> Syntax.check_term lthy''' |
25d11b449e92
definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents:
3045
diff
changeset
|
469 |
val goal_iff1 = Logic.mk_implies (G_aux_prem, G_prem) |
25d11b449e92
definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents:
3045
diff
changeset
|
470 |
|> all x |> all y |
25d11b449e92
definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents:
3045
diff
changeset
|
471 |
val goal_iff2 = Logic.mk_implies (G_prem, G_aux_prem) |
25d11b449e92
definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents:
3045
diff
changeset
|
472 |
|> all x |> all y |
25d11b449e92
definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents:
3045
diff
changeset
|
473 |
|
3229
b52e8651591f
updated to Isabelle changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3227
diff
changeset
|
474 |
val simp_thms = @{thms sum.sel sum.inject sum.case sum.distinct o_apply} |
3218
89158f401b07
updated to simplifier changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3204
diff
changeset
|
475 |
val simpset0 = put_simpset HOL_basic_ss lthy''' addsimps simp_thms |
89158f401b07
updated to simplifier changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3204
diff
changeset
|
476 |
val simpset1 = put_simpset HOL_ss lthy''' addsimps simp_thms |
3197
25d11b449e92
definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents:
3045
diff
changeset
|
477 |
|
25d11b449e92
definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents:
3045
diff
changeset
|
478 |
val inj_thm = Goal.prove lthy''' [] [] goal_inj |
3243
c4f31f1564b7
updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3239
diff
changeset
|
479 |
(K (HEADGOAL (DETERM o eresolve_tac lthy''' [G_aux_induct] |
c4f31f1564b7
updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3239
diff
changeset
|
480 |
THEN_ALL_NEW asm_simp_tac simpset1))) |
3197
25d11b449e92
definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents:
3045
diff
changeset
|
481 |
|
25d11b449e92
definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents:
3045
diff
changeset
|
482 |
fun aux_tac thm = |
3243
c4f31f1564b7
updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3239
diff
changeset
|
483 |
resolve_tac lthy''' [Variable.gen_all lthy''' thm] THEN_ALL_NEW |
3239
67370521c09c
updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3231
diff
changeset
|
484 |
asm_full_simp_tac (simpset1 addsimps [inj_thm]) |
3197
25d11b449e92
definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents:
3045
diff
changeset
|
485 |
|
25d11b449e92
definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents:
3045
diff
changeset
|
486 |
val iff1_thm = Goal.prove lthy''' [] [] goal_iff1 |
3243
c4f31f1564b7
updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3239
diff
changeset
|
487 |
(K (HEADGOAL (DETERM o eresolve_tac lthy''' [G_aux_induct] |
c4f31f1564b7
updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3239
diff
changeset
|
488 |
THEN' RANGE (map aux_tac GIntro_thms)))) |
c4f31f1564b7
updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3239
diff
changeset
|
489 |
|> Variable.gen_all lthy''' |
3197
25d11b449e92
definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents:
3045
diff
changeset
|
490 |
val iff2_thm = Goal.prove lthy''' [] [] goal_iff2 |
3243
c4f31f1564b7
updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3239
diff
changeset
|
491 |
(K (HEADGOAL (DETERM o eresolve_tac lthy''' [G_induct] THEN' RANGE |
3218
89158f401b07
updated to simplifier changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3204
diff
changeset
|
492 |
(map (aux_tac o simplify simpset0) GIntro_aux_thms)))) |
3243
c4f31f1564b7
updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3239
diff
changeset
|
493 |
|> Variable.gen_all lthy''' |
3197
25d11b449e92
definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents:
3045
diff
changeset
|
494 |
|
25d11b449e92
definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents:
3045
diff
changeset
|
495 |
val iff_thm = Goal.prove lthy''' [] [] (HOLogic.mk_Trueprop (HOLogic.mk_eq (G, G_aux))) |
3243
c4f31f1564b7
updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3239
diff
changeset
|
496 |
(K (HEADGOAL (EVERY' ((map (resolve_tac lthy''' o single) @{thms ext ext iffI}) @ |
c4f31f1564b7
updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3239
diff
changeset
|
497 |
[eresolve_tac lthy''' [iff2_thm], eresolve_tac lthy''' [iff1_thm]])))) |
3197
25d11b449e92
definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents:
3045
diff
changeset
|
498 |
|
3218
89158f401b07
updated to simplifier changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3204
diff
changeset
|
499 |
val tac = HEADGOAL (simp_tac (put_simpset HOL_basic_ss lthy''' addsimps [iff_thm])) |
3197
25d11b449e92
definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents:
3045
diff
changeset
|
500 |
val goalstate' = |
25d11b449e92
definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents:
3045
diff
changeset
|
501 |
case (SINGLE tac) goalstate of |
25d11b449e92
definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents:
3045
diff
changeset
|
502 |
NONE => error "auxiliary equivalence proof failed" |
25d11b449e92
definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents:
3045
diff
changeset
|
503 |
| SOME st => st |
2665
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
504 |
in |
3197
25d11b449e92
definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents:
3045
diff
changeset
|
505 |
((goalstate', mutual_cont), lthy''') |
2665
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
506 |
end |
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
507 |
|
16b5a67ee279
exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
508 |
end |