2665
+ − 1
(* Nominal Function Core
+ − 2
Author: Christian Urban
+ − 3
+ − 4
heavily based on the code of Alexander Krauss
+ − 5
(code forked on 14 January 2011)
+ − 6
+ − 7
Core of the nominal function package.
+ − 8
*)
+ − 9
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>
diff
changeset
+ − 10
2665
+ − 11
signature NOMINAL_FUNCTION_CORE =
+ − 12
sig
+ − 13
val trace: bool Unsynchronized.ref
+ − 14
2819
+ − 15
val prepare_nominal_function : Nominal_Function_Common.nominal_function_config
2665
+ − 16
-> string (* defname *)
+ − 17
-> ((bstring * typ) * mixfix) list (* defined symbol *)
+ − 18
-> ((bstring * typ) list * term list * term * term) list (* specification *)
+ − 19
-> local_theory
+ − 20
-> (term (* f *)
3197
25d11b449e92
definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 21
* term (* G(raph) *)
25d11b449e92
definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 22
* thm list (* GIntros *)
25d11b449e92
definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 23
* thm (* Ginduct *)
2665
+ − 24
* thm (* goalstate *)
2973
d1038e67923a
added a flag (eqvt) to termination proofs arising fron nominal_primrecs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 25
* (thm -> Nominal_Function_Common.nominal_function_result) (* continuation *)
2665
+ − 26
) * local_theory
+ − 27
3197
25d11b449e92
definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 28
val inductive_def : (binding * typ) * mixfix -> term list -> local_theory
25d11b449e92
definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 29
-> (term * thm list * thm * thm) * local_theory
2665
+ − 30
end
+ − 31
+ − 32
structure Nominal_Function_Core : NOMINAL_FUNCTION_CORE =
+ − 33
struct
+ − 34
+ − 35
val trace = Unsynchronized.ref false
+ − 36
fun trace_msg msg = if ! trace then tracing (msg ()) else ()
+ − 37
+ − 38
val boolT = HOLogic.boolT
+ − 39
val mk_eq = HOLogic.mk_eq
+ − 40
+ − 41
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>
diff
changeset
+ − 42
open Function_Common
2819
+ − 43
open Nominal_Function_Common
2665
+ − 44
+ − 45
datatype globals = Globals of
+ − 46
{fvar: term,
+ − 47
domT: typ,
+ − 48
ranT: typ,
+ − 49
h: term,
+ − 50
y: term,
+ − 51
x: term,
+ − 52
z: term,
+ − 53
a: term,
+ − 54
P: term,
+ − 55
D: term,
+ − 56
Pbool:term}
+ − 57
+ − 58
datatype rec_call_info = RCInfo of
+ − 59
{RIvs: (string * typ) list, (* Call context: fixes and assumes *)
+ − 60
CCas: thm list,
+ − 61
rcarg: term, (* The recursive argument *)
+ − 62
llRI: thm,
+ − 63
h_assum: term}
+ − 64
+ − 65
+ − 66
datatype clause_context = ClauseContext of
+ − 67
{ctxt : Proof.context,
+ − 68
qs : term list,
+ − 69
gs : term list,
+ − 70
lhs: term,
+ − 71
rhs: term,
+ − 72
cqs: cterm list,
+ − 73
ags: thm list,
+ − 74
case_hyp : thm}
+ − 75
+ − 76
+ − 77
fun transfer_clause_ctx thy (ClauseContext { ctxt, qs, gs, lhs, rhs, cqs, ags, case_hyp }) =
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>
diff
changeset
+ − 78
ClauseContext { ctxt = Proof_Context.transfer thy ctxt,
2665
+ − 79
qs = qs, gs = gs, lhs = lhs, rhs = rhs, cqs = cqs, ags = ags, case_hyp = case_hyp }
+ − 80
+ − 81
+ − 82
datatype clause_info = ClauseInfo of
+ − 83
{no: int,
+ − 84
qglr : ((string * typ) list * term list * term * term),
+ − 85
cdata : clause_context,
+ − 86
tree: Function_Ctx_Tree.ctx_tree,
+ − 87
lGI: thm,
+ − 88
RCs: rec_call_info list}
+ − 89
+ − 90
+ − 91
(* Theory dependencies. *)
+ − 92
val acc_induct_rule = @{thm accp_induct_rule}
+ − 93
3229
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 94
val ex1_implies_ex = @{thm Fun_Def.fundef_ex1_existence}
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 95
val ex1_implies_un = @{thm Fun_Def.fundef_ex1_uniqueness}
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 96
val ex1_implies_iff = @{thm Fun_Def.fundef_ex1_iff}
2665
+ − 97
+ − 98
val acc_downward = @{thm accp_downward}
+ − 99
val accI = @{thm accp.accI}
+ − 100
val case_split = @{thm HOL.case_split}
3229
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 101
val fundef_default_value = @{thm Fun_Def.fundef_default_value}
2665
+ − 102
val not_acc_down = @{thm not_accp_down}
+ − 103
+ − 104
+ − 105
+ − 106
fun find_calls tree =
+ − 107
let
+ − 108
fun add_Ri (fixes,assumes) (_ $ arg) _ (_, xs) =
+ − 109
([], (fixes, assumes, arg) :: xs)
+ − 110
| add_Ri _ _ _ _ = raise Match
+ − 111
in
+ − 112
rev (Function_Ctx_Tree.traverse_tree add_Ri tree [])
+ − 113
end
+ − 114
+ − 115
(* nominal *)
+ − 116
fun mk_eqvt_at (f_trm, arg_trm) =
+ − 117
let
+ − 118
val f_ty = fastype_of f_trm
+ − 119
val arg_ty = domain_type f_ty
+ − 120
in
+ − 121
Const (@{const_name eqvt_at}, [f_ty, arg_ty] ---> @{typ bool}) $ f_trm $ arg_trm
+ − 122
|> HOLogic.mk_Trueprop
+ − 123
end
+ − 124
2707
+ − 125
fun mk_eqvt trm =
+ − 126
let
+ − 127
val ty = fastype_of trm
+ − 128
in
+ − 129
Const (@{const_name eqvt}, ty --> @{typ bool}) $ trm
+ − 130
|> HOLogic.mk_Trueprop
+ − 131
end
+ − 132
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>
diff
changeset
+ − 133
fun mk_inv inv (f_trm, arg_trm) =
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>
diff
changeset
+ − 134
betapplys (inv, [arg_trm, (f_trm $ arg_trm)])
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>
diff
changeset
+ − 135
|> HOLogic.mk_Trueprop
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>
diff
changeset
+ − 136
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>
diff
changeset
+ − 137
fun mk_invariant (Globals {x, y, ...}) G invariant =
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>
diff
changeset
+ − 138
let
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>
diff
changeset
+ − 139
val prem = HOLogic.mk_Trueprop (G $ x $ y)
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>
diff
changeset
+ − 140
val concl = HOLogic.mk_Trueprop (betapplys (invariant, [x, y]))
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>
diff
changeset
+ − 141
in
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>
diff
changeset
+ − 142
Logic.mk_implies (prem, concl)
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>
diff
changeset
+ − 143
|> mk_forall_rename ("y", y)
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>
diff
changeset
+ − 144
|> mk_forall_rename ("x", x)
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>
diff
changeset
+ − 145
end
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>
diff
changeset
+ − 146
2790
+ − 147
(** building proof obligations *)
2802
+ − 148
fun mk_eqvt_proof_obligation qs fvar (vs, assms, arg) =
2789
32979078bfe9
functions involving if and case do not throw exceptions anymore; but eqvt_at assumption has now a precondition
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 149
mk_eqvt_at (fvar, arg)
2803
+ − 150
|> curry Logic.list_implies (map prop_of assms)
3108
+ − 151
|> fold_rev (Logic.all o Free) vs
2994
+ − 152
|> fold_rev absfree qs
2789
32979078bfe9
functions involving if and case do not throw exceptions anymore; but eqvt_at assumption has now a precondition
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 153
|> strip_abs_body
2665
+ − 154
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>
diff
changeset
+ − 155
fun mk_inv_proof_obligation inv qs fvar (vs, assms, arg) =
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>
diff
changeset
+ − 156
mk_inv inv (fvar, arg)
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>
diff
changeset
+ − 157
|> curry Logic.list_implies (map prop_of assms)
3108
+ − 158
|> fold_rev (Logic.all o Free) vs
2994
+ − 159
|> fold_rev absfree qs
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>
diff
changeset
+ − 160
|> strip_abs_body
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>
diff
changeset
+ − 161
2665
+ − 162
(** building proof obligations *)
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>
diff
changeset
+ − 163
fun mk_compat_proof_obligations domT ranT fvar f RCss inv glrs =
2665
+ − 164
let
2862
+ − 165
fun mk_impl (((qs, gs, lhs, rhs), RCs_lhs), ((qs', gs', lhs', rhs'), RCs_rhs)) =
2665
+ − 166
let
+ − 167
val shift = incr_boundvars (length qs')
2862
+ − 168
val eqvts_obligations_lhs = map (shift o mk_eqvt_proof_obligation qs fvar) RCs_lhs
+ − 169
val eqvts_obligations_rhs = map (mk_eqvt_proof_obligation qs' fvar) RCs_rhs
+ − 170
val invs_obligations_lhs = map (shift o mk_inv_proof_obligation inv qs fvar) RCs_lhs
+ − 171
val invs_obligations_rhs = map (mk_inv_proof_obligation inv qs' fvar) RCs_rhs
2665
+ − 172
in
+ − 173
Logic.mk_implies
+ − 174
(HOLogic.mk_Trueprop (HOLogic.eq_const domT $ shift lhs $ lhs'),
+ − 175
HOLogic.mk_Trueprop (HOLogic.eq_const ranT $ shift rhs $ rhs'))
+ − 176
|> fold_rev (curry Logic.mk_implies) (map shift gs @ gs')
2862
+ − 177
|> fold_rev (curry Logic.mk_implies) invs_obligations_rhs (* nominal *)
+ − 178
|> fold_rev (curry Logic.mk_implies) invs_obligations_lhs (* nominal *)
+ − 179
|> fold_rev (curry Logic.mk_implies) eqvts_obligations_rhs (* nominal *)
+ − 180
|> fold_rev (curry Logic.mk_implies) eqvts_obligations_lhs (* nominal *)
3108
+ − 181
|> fold_rev (fn (n,T) => fn b => Logic.all_const T $ Abs(n,T,b)) (qs @ qs')
2665
+ − 182
|> curry abstract_over fvar
+ − 183
|> curry subst_bound f
+ − 184
end
+ − 185
in
2789
32979078bfe9
functions involving if and case do not throw exceptions anymore; but eqvt_at assumption has now a precondition
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 186
map mk_impl (unordered_pairs (glrs ~~ RCss))
2665
+ − 187
end
+ − 188
+ − 189
fun mk_completeness (Globals {x, Pbool, ...}) clauses qglrs =
+ − 190
let
+ − 191
fun mk_case (ClauseContext {qs, gs, lhs, ...}, (oqs, _, _, _)) =
+ − 192
HOLogic.mk_Trueprop Pbool
+ − 193
|> curry Logic.mk_implies (HOLogic.mk_Trueprop (mk_eq (x, lhs)))
+ − 194
|> fold_rev (curry Logic.mk_implies) gs
+ − 195
|> fold_rev mk_forall_rename (map fst oqs ~~ qs)
+ − 196
in
+ − 197
HOLogic.mk_Trueprop Pbool
+ − 198
|> fold_rev (curry Logic.mk_implies o mk_case) (clauses ~~ qglrs)
+ − 199
|> mk_forall_rename ("x", x)
+ − 200
|> mk_forall_rename ("P", Pbool)
+ − 201
end
+ − 202
+ − 203
(** making a context with it's own local bindings **)
+ − 204
+ − 205
fun mk_clause_context x ctxt (pre_qs,pre_gs,pre_lhs,pre_rhs) =
+ − 206
let
+ − 207
val (qs, ctxt') = Variable.variant_fixes (map fst pre_qs) ctxt
+ − 208
|>> map2 (fn (_, T) => fn n => Free (n, T)) pre_qs
+ − 209
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>
diff
changeset
+ − 210
val thy = Proof_Context.theory_of ctxt'
2665
+ − 211
+ − 212
fun inst t = subst_bounds (rev qs, t)
+ − 213
val gs = map inst pre_gs
+ − 214
val lhs = inst pre_lhs
+ − 215
val rhs = inst pre_rhs
+ − 216
+ − 217
val cqs = map (cterm_of thy) qs
+ − 218
val ags = map (Thm.assume o cterm_of thy) gs
+ − 219
+ − 220
val case_hyp = Thm.assume (cterm_of thy (HOLogic.mk_Trueprop (mk_eq (x, lhs))))
+ − 221
in
+ − 222
ClauseContext { ctxt = ctxt', qs = qs, gs = gs, lhs = lhs, rhs = rhs,
+ − 223
cqs = cqs, ags = ags, case_hyp = case_hyp }
+ − 224
end
+ − 225
+ − 226
+ − 227
(* lowlevel term function. FIXME: remove *)
+ − 228
fun abstract_over_list vs body =
+ − 229
let
+ − 230
fun abs lev v tm =
+ − 231
if v aconv tm then Bound lev
+ − 232
else
+ − 233
(case tm of
+ − 234
Abs (a, T, t) => Abs (a, T, abs (lev + 1) v t)
+ − 235
| t $ u => abs lev v t $ abs lev v u
+ − 236
| t => t)
+ − 237
in
+ − 238
fold_index (fn (i, v) => fn t => abs i v t) vs body
+ − 239
end
+ − 240
+ − 241
+ − 242
+ − 243
fun mk_clause_info globals G f no cdata qglr tree RCs GIntro_thm RIntro_thms =
+ − 244
let
+ − 245
val Globals {h, ...} = globals
+ − 246
+ − 247
val ClauseContext { ctxt, qs, cqs, ags, ... } = cdata
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>
diff
changeset
+ − 248
val cert = Thm.cterm_of (Proof_Context.theory_of ctxt)
2665
+ − 249
+ − 250
(* Instantiate the GIntro thm with "f" and import into the clause context. *)
+ − 251
val lGI = GIntro_thm
+ − 252
|> Thm.forall_elim (cert f)
+ − 253
|> fold Thm.forall_elim cqs
+ − 254
|> fold Thm.elim_implies ags
+ − 255
+ − 256
fun mk_call_info (rcfix, rcassm, rcarg) RI =
+ − 257
let
+ − 258
val llRI = RI
+ − 259
|> fold Thm.forall_elim cqs
+ − 260
|> fold (Thm.forall_elim o cert o Free) rcfix
+ − 261
|> fold Thm.elim_implies ags
+ − 262
|> fold Thm.elim_implies rcassm
+ − 263
+ − 264
val h_assum =
+ − 265
HOLogic.mk_Trueprop (G $ rcarg $ (h $ rcarg))
+ − 266
|> fold_rev (curry Logic.mk_implies o prop_of) rcassm
+ − 267
|> fold_rev (Logic.all o Free) rcfix
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>
diff
changeset
+ − 268
|> Pattern.rewrite_term (Proof_Context.theory_of ctxt) [(f, h)] []
2665
+ − 269
|> abstract_over_list (rev qs)
+ − 270
in
+ − 271
RCInfo {RIvs=rcfix, rcarg=rcarg, CCas=rcassm, llRI=llRI, h_assum=h_assum}
+ − 272
end
+ − 273
+ − 274
val RC_infos = map2 mk_call_info RCs RIntro_thms
+ − 275
in
+ − 276
ClauseInfo {no=no, cdata=cdata, qglr=qglr, lGI=lGI, RCs=RC_infos,
+ − 277
tree=tree}
+ − 278
end
+ − 279
+ − 280
+ − 281
fun store_compat_thms 0 thms = []
+ − 282
| store_compat_thms n thms =
+ − 283
let
+ − 284
val (thms1, thms2) = chop n thms
+ − 285
in
+ − 286
(thms1 :: store_compat_thms (n - 1) thms2)
+ − 287
end
+ − 288
+ − 289
(* expects i <= j *)
+ − 290
fun lookup_compat_thm i j cts =
+ − 291
nth (nth cts (i - 1)) (j - i)
+ − 292
+ − 293
(* nominal *)
+ − 294
(* Returns "Gsi, Gsj, lhs_i = lhs_j |-- rhs_j_f = rhs_i_f" *)
+ − 295
(* if j < i, then turn around *)
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>
diff
changeset
+ − 296
fun get_compat_thm thy cts eqvtsi eqvtsj invsi invsj i j ctxi ctxj =
2665
+ − 297
let
2789
32979078bfe9
functions involving if and case do not throw exceptions anymore; but eqvt_at assumption has now a precondition
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 298
val ClauseContext {cqs=cqsi,ags=agsi,lhs=lhsi,case_hyp=case_hypi,...} = ctxi
32979078bfe9
functions involving if and case do not throw exceptions anymore; but eqvt_at assumption has now a precondition
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 299
val ClauseContext {cqs=cqsj,ags=agsj,lhs=lhsj,case_hyp=case_hypj,...} = ctxj
2665
+ − 300
+ − 301
val lhsi_eq_lhsj = cterm_of thy (HOLogic.mk_Trueprop (mk_eq (lhsi, lhsj)))
2848
da7e6655cd4c
fixed the problem when giving a complex default-term; the fundef lemmas in Nominal_Base were not general enough
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 302
2665
+ − 303
in if j < i then
+ − 304
let
+ − 305
val compat = lookup_compat_thm j i cts
+ − 306
in
+ − 307
compat (* "!!qj qi. Gsj => Gsi => lhsj = lhsi ==> rhsj = rhsi" *)
+ − 308
|> fold Thm.forall_elim (cqsj @ cqsi) (* "Gsj => Gsi => lhsj = lhsi ==> rhsj = rhsi" *)
2789
32979078bfe9
functions involving if and case do not throw exceptions anymore; but eqvt_at assumption has now a precondition
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 309
|> fold Thm.elim_implies eqvtsj (* nominal *)
2862
+ − 310
|> fold Thm.elim_implies eqvtsi (* nominal *)
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>
diff
changeset
+ − 311
|> fold Thm.elim_implies invsj (* nominal *)
2862
+ − 312
|> fold Thm.elim_implies invsi (* nominal *)
2665
+ − 313
|> fold Thm.elim_implies agsj
+ − 314
|> fold Thm.elim_implies agsi
+ − 315
|> Thm.elim_implies ((Thm.assume lhsi_eq_lhsj) RS sym) (* "Gsj, Gsi, lhsi = lhsj |-- rhsj = rhsi" *)
+ − 316
end
+ − 317
else
+ − 318
let
+ − 319
val compat = lookup_compat_thm i j cts
+ − 320
in
+ − 321
compat (* "!!qi qj. Gsi => Gsj => lhsi = lhsj ==> rhsi = rhsj" *)
+ − 322
|> fold Thm.forall_elim (cqsi @ cqsj) (* "Gsi => Gsj => lhsi = lhsj ==> rhsi = rhsj" *)
2789
32979078bfe9
functions involving if and case do not throw exceptions anymore; but eqvt_at assumption has now a precondition
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 323
|> fold Thm.elim_implies eqvtsi (* nominal *)
2862
+ − 324
|> fold Thm.elim_implies eqvtsj (* nominal *)
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>
diff
changeset
+ − 325
|> fold Thm.elim_implies invsi (* nominal *)
2862
+ − 326
|> fold Thm.elim_implies invsj (* nominal *)
2665
+ − 327
|> fold Thm.elim_implies agsi
+ − 328
|> fold Thm.elim_implies agsj
+ − 329
|> Thm.elim_implies (Thm.assume lhsi_eq_lhsj)
+ − 330
|> (fn thm => thm RS sym) (* "Gsi, Gsj, lhsi = lhsj |-- rhsj = rhsi" *)
+ − 331
end
+ − 332
end
+ − 333
+ − 334
+ − 335
(* Generates the replacement lemma in fully quantified form. *)
+ − 336
fun mk_replacement_lemma thy h ih_elim clause =
+ − 337
let
+ − 338
val ClauseInfo {cdata=ClauseContext {qs, lhs, cqs, ags, case_hyp, ...},
+ − 339
RCs, tree, ...} = clause
+ − 340
local open Conv in
+ − 341
val ih_conv = arg1_conv o arg_conv o arg_conv
+ − 342
end
+ − 343
+ − 344
val ih_elim_case =
+ − 345
Conv.fconv_rule (ih_conv (K (case_hyp RS eq_reflection))) ih_elim
+ − 346
+ − 347
val Ris = map (fn RCInfo {llRI, ...} => llRI) RCs
+ − 348
val h_assums = map (fn RCInfo {h_assum, ...} =>
+ − 349
Thm.assume (cterm_of thy (subst_bounds (rev qs, h_assum)))) RCs
+ − 350
+ − 351
val (eql, _) =
3218
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 352
Function_Ctx_Tree.rewrite_by_tree (Proof_Context.init_global thy)
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 353
h ih_elim_case (Ris ~~ h_assums) tree
2665
+ − 354
+ − 355
val replace_lemma = (eql RS meta_eq_to_obj_eq)
+ − 356
|> Thm.implies_intr (cprop_of case_hyp)
+ − 357
|> fold_rev (Thm.implies_intr o cprop_of) h_assums
+ − 358
|> fold_rev (Thm.implies_intr o cprop_of) ags
+ − 359
|> fold_rev Thm.forall_intr cqs
+ − 360
|> Thm.close_derivation
+ − 361
in
+ − 362
replace_lemma
+ − 363
end
+ − 364
+ − 365
(* nominal *)
+ − 366
(* Generates the eqvt lemmas for each clause *)
+ − 367
fun mk_eqvt_lemma thy ih_eqvt clause =
+ − 368
let
+ − 369
val ClauseInfo {cdata=ClauseContext {cqs, ags, case_hyp, ...}, RCs, ...} = clause
2789
32979078bfe9
functions involving if and case do not throw exceptions anymore; but eqvt_at assumption has now a precondition
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 370
2665
+ − 371
local open Conv in
+ − 372
val ih_conv = arg1_conv o arg_conv o arg_conv
+ − 373
end
+ − 374
+ − 375
val ih_eqvt_case =
+ − 376
Conv.fconv_rule (ih_conv (K (case_hyp RS eq_reflection))) ih_eqvt
+ − 377
+ − 378
fun prep_eqvt (RCInfo {llRI, RIvs, CCas, ...}) =
+ − 379
(llRI RS ih_eqvt_case)
+ − 380
|> fold_rev (Thm.implies_intr o cprop_of) CCas
2802
+ − 381
|> fold_rev (Thm.forall_intr o cterm_of thy o Free) RIvs
2665
+ − 382
in
+ − 383
map prep_eqvt RCs
+ − 384
|> map (fold_rev (Thm.implies_intr o cprop_of) ags)
+ − 385
|> map (Thm.implies_intr (cprop_of case_hyp))
+ − 386
|> map (fold_rev Thm.forall_intr cqs)
+ − 387
|> map (Thm.close_derivation)
+ − 388
end
+ − 389
+ − 390
(* nominal *)
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>
diff
changeset
+ − 391
fun mk_invariant_lemma thy ih_inv clause =
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>
diff
changeset
+ − 392
let
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>
diff
changeset
+ − 393
val ClauseInfo {cdata=ClauseContext {cqs, ags, case_hyp, ...}, RCs, ...} = clause
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>
diff
changeset
+ − 394
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>
diff
changeset
+ − 395
local open Conv in
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>
diff
changeset
+ − 396
val ih_conv = arg1_conv o arg_conv o arg_conv
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>
diff
changeset
+ − 397
end
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>
diff
changeset
+ − 398
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>
diff
changeset
+ − 399
val ih_inv_case =
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>
diff
changeset
+ − 400
Conv.fconv_rule (ih_conv (K (case_hyp RS eq_reflection))) ih_inv
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>
diff
changeset
+ − 401
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>
diff
changeset
+ − 402
fun prep_inv (RCInfo {llRI, RIvs, CCas, ...}) =
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>
diff
changeset
+ − 403
(llRI RS ih_inv_case)
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>
diff
changeset
+ − 404
|> fold_rev (Thm.implies_intr o cprop_of) CCas
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>
diff
changeset
+ − 405
|> fold_rev (Thm.forall_intr o cterm_of thy o Free) RIvs
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>
diff
changeset
+ − 406
in
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>
diff
changeset
+ − 407
map prep_inv RCs
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>
diff
changeset
+ − 408
|> map (fold_rev (Thm.implies_intr o cprop_of) ags)
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>
diff
changeset
+ − 409
|> map (Thm.implies_intr (cprop_of case_hyp))
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>
diff
changeset
+ − 410
|> map (fold_rev Thm.forall_intr cqs)
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>
diff
changeset
+ − 411
|> map (Thm.close_derivation)
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>
diff
changeset
+ − 412
end
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>
diff
changeset
+ − 413
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>
diff
changeset
+ − 414
(* nominal *)
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>
diff
changeset
+ − 415
fun mk_uniqueness_clause thy globals compat_store eqvts invs clausei clausej RLj =
2665
+ − 416
let
+ − 417
val Globals {h, y, x, fvar, ...} = globals
+ − 418
val ClauseInfo {no=i, cdata=cctxi as ClauseContext {ctxt=ctxti, lhs=lhsi, case_hyp, cqs = cqsi,
+ − 419
ags = agsi, ...}, ...} = clausei
+ − 420
val ClauseInfo {no=j, qglr=cdescj, RCs=RCsj, ...} = clausej
+ − 421
+ − 422
val cctxj as ClauseContext {ags = agsj', lhs = lhsj', rhs = rhsj', qs = qsj', cqs = cqsj', ...} =
+ − 423
mk_clause_context x ctxti cdescj
+ − 424
+ − 425
val rhsj'h = Pattern.rewrite_term thy [(fvar,h)] [] rhsj'
+ − 426
+ − 427
val Ghsj' = map
+ − 428
(fn RCInfo {h_assum, ...} => Thm.assume (cterm_of thy (subst_bounds (rev qsj', h_assum)))) RCsj
+ − 429
+ − 430
val y_eq_rhsj'h = Thm.assume (cterm_of thy (HOLogic.mk_Trueprop (mk_eq (y, rhsj'h))))
+ − 431
val lhsi_eq_lhsj' = Thm.assume (cterm_of thy (HOLogic.mk_Trueprop (mk_eq (lhsi, lhsj'))))
+ − 432
(* lhs_i = lhs_j' |-- lhs_i = lhs_j' *)
+ − 433
+ − 434
val case_hypj' = trans OF [case_hyp, lhsi_eq_lhsj']
+ − 435
+ − 436
val RLj_import = RLj
+ − 437
|> fold Thm.forall_elim cqsj'
+ − 438
|> fold Thm.elim_implies agsj'
+ − 439
|> fold Thm.elim_implies Ghsj'
+ − 440
+ − 441
val eqvtsi = nth eqvts (i - 1)
+ − 442
|> map (fold Thm.forall_elim cqsi)
+ − 443
|> map (fold Thm.elim_implies [case_hyp])
+ − 444
|> map (fold Thm.elim_implies agsi)
+ − 445
+ − 446
val eqvtsj = nth eqvts (j - 1)
+ − 447
|> map (fold Thm.forall_elim cqsj')
+ − 448
|> map (fold Thm.elim_implies [case_hypj'])
+ − 449
|> map (fold Thm.elim_implies agsj')
+ − 450
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>
diff
changeset
+ − 451
val invsi = nth invs (i - 1)
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>
diff
changeset
+ − 452
|> map (fold Thm.forall_elim cqsi)
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>
diff
changeset
+ − 453
|> map (fold Thm.elim_implies [case_hyp])
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>
diff
changeset
+ − 454
|> map (fold Thm.elim_implies agsi)
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>
diff
changeset
+ − 455
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>
diff
changeset
+ − 456
val invsj = nth invs (j - 1)
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>
diff
changeset
+ − 457
|> map (fold Thm.forall_elim cqsj')
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>
diff
changeset
+ − 458
|> map (fold Thm.elim_implies [case_hypj'])
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>
diff
changeset
+ − 459
|> map (fold Thm.elim_implies agsj')
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>
diff
changeset
+ − 460
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>
diff
changeset
+ − 461
val compat = get_compat_thm thy compat_store eqvtsi eqvtsj invsi invsj i j cctxi cctxj
2665
+ − 462
+ − 463
in
+ − 464
(trans OF [case_hyp, lhsi_eq_lhsj']) (* lhs_i = lhs_j' |-- x = lhs_j' *)
+ − 465
|> Thm.implies_elim RLj_import
+ − 466
(* Rj1' ... Rjk', lhs_i = lhs_j' |-- rhs_j'_h = rhs_j'_f *)
+ − 467
|> (fn it => trans OF [it, compat])
+ − 468
(* lhs_i = lhs_j', Gj', Rj1' ... Rjk' |-- rhs_j'_h = rhs_i_f *)
+ − 469
|> (fn it => trans OF [y_eq_rhsj'h, it])
+ − 470
(* lhs_i = lhs_j', Gj', Rj1' ... Rjk', y = rhs_j_h' |-- y = rhs_i_f *)
+ − 471
|> fold_rev (Thm.implies_intr o cprop_of) Ghsj'
+ − 472
|> fold_rev (Thm.implies_intr o cprop_of) agsj'
+ − 473
(* lhs_i = lhs_j' , y = rhs_j_h' |-- Gj', Rj1'...Rjk' ==> y = rhs_i_f *)
+ − 474
|> Thm.implies_intr (cprop_of y_eq_rhsj'h)
+ − 475
|> Thm.implies_intr (cprop_of lhsi_eq_lhsj')
+ − 476
|> fold_rev Thm.forall_intr (cterm_of thy h :: cqsj')
+ − 477
end
+ − 478
+ − 479
(* nominal *)
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>
diff
changeset
+ − 480
fun mk_uniqueness_case thy globals G f ihyp ih_intro G_cases compat_store clauses replems eqvtlems invlems
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>
diff
changeset
+ − 481
clausei =
2665
+ − 482
let
+ − 483
val Globals {x, y, ranT, fvar, ...} = globals
+ − 484
val ClauseInfo {cdata = ClauseContext {lhs, rhs, cqs, ags, case_hyp, ...}, lGI, RCs, ...} = clausei
+ − 485
val rhsC = Pattern.rewrite_term thy [(fvar, f)] [] rhs
+ − 486
3218
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 487
val ih_intro_case =
3227
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 488
full_simplify (put_simpset HOL_basic_ss (Proof_Context.init_global thy) addsimps [case_hyp])
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 489
ih_intro
2665
+ − 490
+ − 491
fun prep_RC (RCInfo {llRI, RIvs, CCas, ...}) =
+ − 492
(llRI RS ih_intro_case)
+ − 493
|> fold_rev (Thm.implies_intr o cprop_of) CCas
+ − 494
|> fold_rev (Thm.forall_intr o cterm_of thy o Free) RIvs
2789
32979078bfe9
functions involving if and case do not throw exceptions anymore; but eqvt_at assumption has now a precondition
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 495
2665
+ − 496
val existence = fold (curry op COMP o prep_RC) RCs lGI
+ − 497
+ − 498
val P = cterm_of thy (mk_eq (y, rhsC))
+ − 499
val G_lhs_y = Thm.assume (cterm_of thy (HOLogic.mk_Trueprop (G $ lhs $ y)))
+ − 500
+ − 501
val unique_clauses =
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>
diff
changeset
+ − 502
map2 (mk_uniqueness_clause thy globals compat_store eqvtlems invlems clausei) clauses replems
2665
+ − 503
+ − 504
fun elim_implies_eta A AB =
3219
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 505
Thm.bicompose {flatten = false, match = true, incremented = false} (false, A, 0) 1 AB
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 506
|> Seq.list_of |> the_single
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 507
2665
+ − 508
val uniqueness = G_cases
+ − 509
|> Thm.forall_elim (cterm_of thy lhs)
+ − 510
|> Thm.forall_elim (cterm_of thy y)
+ − 511
|> Thm.forall_elim P
+ − 512
|> Thm.elim_implies G_lhs_y
+ − 513
|> fold elim_implies_eta unique_clauses
+ − 514
|> Thm.implies_intr (cprop_of G_lhs_y)
+ − 515
|> Thm.forall_intr (cterm_of thy y)
+ − 516
+ − 517
val P2 = cterm_of thy (lambda y (G $ lhs $ y)) (* P2 y := (lhs, y): G *)
+ − 518
+ − 519
val exactly_one =
+ − 520
ex1I |> instantiate' [SOME (ctyp_of thy ranT)] [SOME P2, SOME (cterm_of thy rhsC)]
+ − 521
|> curry (op COMP) existence
+ − 522
|> curry (op COMP) uniqueness
3227
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 523
|> simplify
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 524
(put_simpset HOL_basic_ss (Proof_Context.init_global thy) addsimps [case_hyp RS sym])
2665
+ − 525
|> Thm.implies_intr (cprop_of case_hyp)
+ − 526
|> fold_rev (Thm.implies_intr o cprop_of) ags
+ − 527
|> fold_rev Thm.forall_intr cqs
+ − 528
+ − 529
val function_value =
+ − 530
existence
+ − 531
|> Thm.implies_intr ihyp
+ − 532
|> Thm.implies_intr (cprop_of case_hyp)
+ − 533
|> Thm.forall_intr (cterm_of thy x)
+ − 534
|> Thm.forall_elim (cterm_of thy lhs)
+ − 535
|> curry (op RS) refl
+ − 536
in
+ − 537
(exactly_one, function_value)
+ − 538
end
+ − 539
+ − 540
+ − 541
(* nominal *)
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>
diff
changeset
+ − 542
fun prove_stuff ctxt globals G f R clauses complete compat compat_store G_elim G_eqvt invariant f_def =
2665
+ − 543
let
+ − 544
val Globals {h, domT, ranT, x, ...} = globals
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>
diff
changeset
+ − 545
val thy = Proof_Context.theory_of ctxt
2665
+ − 546
+ − 547
(* Inductive Hypothesis: !!z. (z,x):R ==> EX!y. (z,y):G *)
3108
+ − 548
val ihyp = Logic.all_const domT $ Abs ("z", domT,
2665
+ − 549
Logic.mk_implies (HOLogic.mk_Trueprop (R $ Bound 0 $ x),
+ − 550
HOLogic.mk_Trueprop (Const (@{const_name Ex1}, (ranT --> boolT) --> boolT) $
+ − 551
Abs ("y", ranT, G $ Bound 1 $ Bound 0))))
+ − 552
|> cterm_of thy
+ − 553
+ − 554
val ihyp_thm = Thm.assume ihyp |> Thm.forall_elim_vars 0
+ − 555
val ih_intro = ihyp_thm RS (f_def RS ex1_implies_ex)
+ − 556
val ih_elim = ihyp_thm RS (f_def RS ex1_implies_un)
+ − 557
|> instantiate' [] [NONE, SOME (cterm_of thy h)]
+ − 558
val ih_eqvt = ihyp_thm RS (G_eqvt RS (f_def RS @{thm fundef_ex1_eqvt_at}))
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>
diff
changeset
+ − 559
val ih_inv = ihyp_thm RS (invariant COMP (f_def RS @{thm fundef_ex1_prop}))
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>
diff
changeset
+ − 560
2665
+ − 561
val _ = trace_msg (K "Proving Replacement lemmas...")
+ − 562
val repLemmas = map (mk_replacement_lemma thy h ih_elim) clauses
+ − 563
+ − 564
val _ = trace_msg (K "Proving Equivariance lemmas...")
+ − 565
val eqvtLemmas = map (mk_eqvt_lemma thy ih_eqvt) clauses
+ − 566
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>
diff
changeset
+ − 567
val _ = trace_msg (K "Proving Invariance lemmas...")
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>
diff
changeset
+ − 568
val invLemmas = map (mk_invariant_lemma thy ih_inv) clauses
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>
diff
changeset
+ − 569
2665
+ − 570
val _ = trace_msg (K "Proving cases for unique existence...")
+ − 571
val (ex1s, values) =
+ − 572
split_list (map (mk_uniqueness_case thy globals G f
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>
diff
changeset
+ − 573
ihyp ih_intro G_elim compat_store clauses repLemmas eqvtLemmas invLemmas) clauses)
2665
+ − 574
+ − 575
val _ = trace_msg (K "Proving: Graph is a function")
+ − 576
val graph_is_function = complete
+ − 577
|> Thm.forall_elim_vars 0
+ − 578
|> fold (curry op COMP) ex1s
+ − 579
|> Thm.implies_intr (ihyp)
+ − 580
|> Thm.implies_intr (cterm_of thy (HOLogic.mk_Trueprop (mk_acc domT R $ x)))
+ − 581
|> Thm.forall_intr (cterm_of thy x)
3220
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 582
|> (fn it => Drule.compose (it, 2, acc_induct_rule)) (* "EX! y. (?x,y):G" *)
2665
+ − 583
|> (fn it => fold (Thm.forall_intr o cterm_of thy o Var) (Term.add_vars (prop_of it) []) it)
+ − 584
2707
+ − 585
val goalstate =
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>
diff
changeset
+ − 586
Conjunction.intr (Conjunction.intr (Conjunction.intr graph_is_function complete) invariant) G_eqvt
2665
+ − 587
|> Thm.close_derivation
3220
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 588
|> Goal.protect 0
2665
+ − 589
|> fold_rev (Thm.implies_intr o cprop_of) compat
+ − 590
|> Thm.implies_intr (cprop_of complete)
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>
diff
changeset
+ − 591
|> Thm.implies_intr (cprop_of invariant)
2707
+ − 592
|> Thm.implies_intr (cprop_of G_eqvt)
2665
+ − 593
in
+ − 594
(goalstate, values)
+ − 595
end
+ − 596
+ − 597
(* wrapper -- restores quantifiers in rule specifications *)
2707
+ − 598
fun inductive_def (binding as ((R, T), _)) intrs lthy =
3197
25d11b449e92
definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 599
let
2665
+ − 600
val ({intrs = intrs_gen, elims = [elim_gen], preds = [ Rdef ], induct, raw_induct, ...}, lthy) =
+ − 601
lthy
2885
1264f2a21ea9
some rudimentary infrastructure for storing data about nominal datatypes
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 602
|> Local_Theory.conceal
2665
+ − 603
|> Inductive.add_inductive_i
+ − 604
{quiet_mode = true,
+ − 605
verbose = ! trace,
+ − 606
alt_name = Binding.empty,
+ − 607
coind = false,
+ − 608
no_elim = false,
+ − 609
no_ind = false,
3200
+ − 610
skip_mono = true}
2665
+ − 611
[binding] (* relation *)
+ − 612
[] (* no parameters *)
+ − 613
(map (fn t => (Attrib.empty_binding, t)) intrs) (* intro rules *)
+ − 614
[] (* no special monos *)
+ − 615
||> Local_Theory.restore_naming lthy
+ − 616
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>
diff
changeset
+ − 617
val cert = cterm_of (Proof_Context.theory_of lthy)
2665
+ − 618
fun requantify orig_intro thm =
+ − 619
let
+ − 620
val (qs, t) = dest_all_all orig_intro
2781
+ − 621
val frees = Variable.add_frees lthy t [] |> remove (op =) (Binding.name_of R, T)
+ − 622
val vars = Term.add_vars (prop_of thm) []
2665
+ − 623
val varmap = AList.lookup (op =) (frees ~~ map fst vars)
2781
+ − 624
#> the_default ("",0)
2665
+ − 625
in
+ − 626
fold_rev (fn Free (n, T) =>
+ − 627
forall_intr_rename (n, cert (Var (varmap (n, T), T)))) qs thm
+ − 628
end
+ − 629
in
2707
+ − 630
((Rdef, map2 requantify intrs intrs_gen, forall_intr_vars elim_gen, induct), lthy)
2665
+ − 631
end
+ − 632
+ − 633
(* nominal *)
+ − 634
fun define_graph Gname fvar domT ranT clauses RCss lthy =
+ − 635
let
+ − 636
val GT = domT --> ranT --> boolT
+ − 637
val (Gvar as (n, T)) = singleton (Variable.variant_frees lthy []) (Gname, GT)
+ − 638
+ − 639
fun mk_GIntro (ClauseContext {qs, gs, lhs, rhs, ...}) RCs =
+ − 640
let
+ − 641
fun mk_h_assm (rcfix, rcassm, rcarg) =
+ − 642
HOLogic.mk_Trueprop (Free Gvar $ rcarg $ (fvar $ rcarg))
+ − 643
|> fold_rev (curry Logic.mk_implies o prop_of) rcassm
+ − 644
|> fold_rev (Logic.all o Free) rcfix
+ − 645
in
+ − 646
HOLogic.mk_Trueprop (Free Gvar $ lhs $ rhs)
+ − 647
|> fold_rev (curry Logic.mk_implies o mk_h_assm) RCs
+ − 648
|> fold_rev (curry Logic.mk_implies) gs
+ − 649
|> fold_rev Logic.all (fvar :: qs)
+ − 650
end
+ − 651
+ − 652
val G_intros = map2 mk_GIntro clauses RCss
+ − 653
in
2707
+ − 654
inductive_def ((Binding.name n, T), NoSyn) G_intros lthy
2665
+ − 655
end
+ − 656
+ − 657
fun define_function fdefname (fname, mixfix) domT ranT G default lthy =
+ − 658
let
+ − 659
val f_def =
3229
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 660
Abs ("x", domT, Const (@{const_name Fun_Def.THE_default}, ranT --> (ranT --> boolT) --> ranT)
2665
+ − 661
$ (default $ Bound 0) $ Abs ("y", ranT, G $ Bound 1 $ Bound 0))
+ − 662
|> Syntax.check_term lthy
+ − 663
in
+ − 664
Local_Theory.define
+ − 665
((Binding.name (function_name fname), mixfix),
+ − 666
((Binding.conceal (Binding.name fdefname), []), f_def)) lthy
+ − 667
end
+ − 668
+ − 669
(* nominal *)
+ − 670
fun define_recursion_relation Rname domT qglrs clauses RCss lthy =
+ − 671
let
+ − 672
val RT = domT --> domT --> boolT
+ − 673
val (Rvar as (n, T)) = singleton (Variable.variant_frees lthy []) (Rname, RT)
+ − 674
+ − 675
fun mk_RIntro (ClauseContext {qs, gs, lhs, ...}, (oqs, _, _, _)) (rcfix, rcassm, rcarg) =
+ − 676
HOLogic.mk_Trueprop (Free Rvar $ rcarg $ lhs)
+ − 677
|> fold_rev (curry Logic.mk_implies o prop_of) rcassm
+ − 678
|> fold_rev (curry Logic.mk_implies) gs
+ − 679
|> fold_rev (Logic.all o Free) rcfix
+ − 680
|> fold_rev mk_forall_rename (map fst oqs ~~ qs)
+ − 681
(* "!!qs xs. CS ==> G => (r, lhs) : R" *)
+ − 682
+ − 683
val R_intross = map2 (map o mk_RIntro) (clauses ~~ qglrs) RCss
+ − 684
2707
+ − 685
val ((R, RIntro_thms, R_elim, _), lthy) =
+ − 686
inductive_def ((Binding.name n, T), NoSyn) (flat R_intross) lthy
2665
+ − 687
in
+ − 688
((R, Library.unflat R_intross RIntro_thms, R_elim), lthy)
+ − 689
end
+ − 690
+ − 691
+ − 692
fun fix_globals domT ranT fvar ctxt =
+ − 693
let
+ − 694
val ([h, y, x, z, a, D, P, Pbool],ctxt') = Variable.variant_fixes
+ − 695
["h_fd", "y_fd", "x_fd", "z_fd", "a_fd", "D_fd", "P_fd", "Pb_fd"] ctxt
+ − 696
in
+ − 697
(Globals {h = Free (h, domT --> ranT),
+ − 698
y = Free (y, ranT),
+ − 699
x = Free (x, domT),
+ − 700
z = Free (z, domT),
+ − 701
a = Free (a, domT),
+ − 702
D = Free (D, domT --> boolT),
+ − 703
P = Free (P, domT --> boolT),
+ − 704
Pbool = Free (Pbool, boolT),
+ − 705
fvar = fvar,
+ − 706
domT = domT,
+ − 707
ranT = ranT},
+ − 708
ctxt')
+ − 709
end
+ − 710
+ − 711
fun inst_RC thy fvar f (rcfix, rcassm, rcarg) =
+ − 712
let
+ − 713
fun inst_term t = subst_bound(f, abstract_over (fvar, t))
+ − 714
in
+ − 715
(rcfix, map (Thm.assume o cterm_of thy o inst_term o prop_of) rcassm, inst_term rcarg)
+ − 716
end
+ − 717
+ − 718
+ − 719
+ − 720
(**********************************************************
+ − 721
* PROVING THE RULES
+ − 722
**********************************************************)
+ − 723
+ − 724
fun mk_psimps thy globals R clauses valthms f_iff graph_is_function =
+ − 725
let
+ − 726
val Globals {domT, z, ...} = globals
+ − 727
+ − 728
fun mk_psimp (ClauseInfo {qglr = (oqs, _, _, _), cdata = ClauseContext {cqs, lhs, ags, ...}, ...}) valthm =
+ − 729
let
+ − 730
val lhs_acc = cterm_of thy (HOLogic.mk_Trueprop (mk_acc domT R $ lhs)) (* "acc R lhs" *)
+ − 731
val z_smaller = cterm_of thy (HOLogic.mk_Trueprop (R $ z $ lhs)) (* "R z lhs" *)
+ − 732
in
+ − 733
((Thm.assume z_smaller) RS ((Thm.assume lhs_acc) RS acc_downward))
+ − 734
|> (fn it => it COMP graph_is_function)
+ − 735
|> Thm.implies_intr z_smaller
+ − 736
|> Thm.forall_intr (cterm_of thy z)
+ − 737
|> (fn it => it COMP valthm)
+ − 738
|> Thm.implies_intr lhs_acc
3227
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 739
|> asm_simplify (put_simpset HOL_basic_ss (Proof_Context.init_global thy) addsimps [f_iff])
2665
+ − 740
|> fold_rev (Thm.implies_intr o cprop_of) ags
+ − 741
|> fold_rev forall_intr_rename (map fst oqs ~~ cqs)
+ − 742
end
+ − 743
in
+ − 744
map2 mk_psimp clauses valthms
+ − 745
end
+ − 746
+ − 747
+ − 748
(** Induction rule **)
+ − 749
+ − 750
+ − 751
val acc_subset_induct = @{thm predicate1I} RS @{thm accp_subset_induct}
+ − 752
+ − 753
+ − 754
fun mk_partial_induct_rule thy globals R complete_thm clauses =
+ − 755
let
+ − 756
val Globals {domT, x, z, a, P, D, ...} = globals
+ − 757
val acc_R = mk_acc domT R
+ − 758
+ − 759
val x_D = Thm.assume (cterm_of thy (HOLogic.mk_Trueprop (D $ x)))
+ − 760
val a_D = cterm_of thy (HOLogic.mk_Trueprop (D $ a))
+ − 761
+ − 762
val D_subset = cterm_of thy (Logic.all x
+ − 763
(Logic.mk_implies (HOLogic.mk_Trueprop (D $ x), HOLogic.mk_Trueprop (acc_R $ x))))
+ − 764
+ − 765
val D_dcl = (* "!!x z. [| x: D; (z,x):R |] ==> z:D" *)
+ − 766
Logic.all x (Logic.all z (Logic.mk_implies (HOLogic.mk_Trueprop (D $ x),
+ − 767
Logic.mk_implies (HOLogic.mk_Trueprop (R $ z $ x),
+ − 768
HOLogic.mk_Trueprop (D $ z)))))
+ − 769
|> cterm_of thy
+ − 770
+ − 771
(* Inductive Hypothesis: !!z. (z,x):R ==> P z *)
3108
+ − 772
val ihyp = Logic.all_const domT $ Abs ("z", domT,
2665
+ − 773
Logic.mk_implies (HOLogic.mk_Trueprop (R $ Bound 0 $ x),
+ − 774
HOLogic.mk_Trueprop (P $ Bound 0)))
+ − 775
|> cterm_of thy
+ − 776
+ − 777
val aihyp = Thm.assume ihyp
+ − 778
+ − 779
fun prove_case clause =
+ − 780
let
+ − 781
val ClauseInfo {cdata = ClauseContext {ctxt, qs, cqs, ags, gs, lhs, case_hyp, ...},
+ − 782
RCs, qglr = (oqs, _, _, _), ...} = clause
+ − 783
+ − 784
val case_hyp_conv = K (case_hyp RS eq_reflection)
+ − 785
local open Conv in
+ − 786
val lhs_D = fconv_rule (arg_conv (arg_conv (case_hyp_conv))) x_D
+ − 787
val sih =
+ − 788
fconv_rule (Conv.binder_conv
+ − 789
(K (arg1_conv (arg_conv (arg_conv case_hyp_conv)))) ctxt) aihyp
+ − 790
end
+ − 791
+ − 792
fun mk_Prec (RCInfo {llRI, RIvs, CCas, rcarg, ...}) = sih
+ − 793
|> Thm.forall_elim (cterm_of thy rcarg)
+ − 794
|> Thm.elim_implies llRI
+ − 795
|> fold_rev (Thm.implies_intr o cprop_of) CCas
+ − 796
|> fold_rev (Thm.forall_intr o cterm_of thy o Free) RIvs
+ − 797
+ − 798
val P_recs = map mk_Prec RCs (* [P rec1, P rec2, ... ] *)
+ − 799
+ − 800
val step = HOLogic.mk_Trueprop (P $ lhs)
+ − 801
|> fold_rev (curry Logic.mk_implies o prop_of) P_recs
+ − 802
|> fold_rev (curry Logic.mk_implies) gs
+ − 803
|> curry Logic.mk_implies (HOLogic.mk_Trueprop (D $ lhs))
+ − 804
|> fold_rev mk_forall_rename (map fst oqs ~~ qs)
+ − 805
|> cterm_of thy
+ − 806
+ − 807
val P_lhs = Thm.assume step
+ − 808
|> fold Thm.forall_elim cqs
+ − 809
|> Thm.elim_implies lhs_D
+ − 810
|> fold Thm.elim_implies ags
+ − 811
|> fold Thm.elim_implies P_recs
+ − 812
+ − 813
val res = cterm_of thy (HOLogic.mk_Trueprop (P $ x))
+ − 814
|> Conv.arg_conv (Conv.arg_conv case_hyp_conv)
+ − 815
|> Thm.symmetric (* P lhs == P x *)
+ − 816
|> (fn eql => Thm.equal_elim eql P_lhs) (* "P x" *)
+ − 817
|> Thm.implies_intr (cprop_of case_hyp)
+ − 818
|> fold_rev (Thm.implies_intr o cprop_of) ags
+ − 819
|> fold_rev Thm.forall_intr cqs
+ − 820
in
+ − 821
(res, step)
+ − 822
end
+ − 823
+ − 824
val (cases, steps) = split_list (map prove_case clauses)
+ − 825
+ − 826
val istep = complete_thm
+ − 827
|> Thm.forall_elim_vars 0
+ − 828
|> fold (curry op COMP) cases (* P x *)
+ − 829
|> Thm.implies_intr ihyp
+ − 830
|> Thm.implies_intr (cprop_of x_D)
+ − 831
|> Thm.forall_intr (cterm_of thy x)
+ − 832
+ − 833
val subset_induct_rule =
+ − 834
acc_subset_induct
+ − 835
|> (curry op COMP) (Thm.assume D_subset)
+ − 836
|> (curry op COMP) (Thm.assume D_dcl)
+ − 837
|> (curry op COMP) (Thm.assume a_D)
+ − 838
|> (curry op COMP) istep
+ − 839
|> fold_rev Thm.implies_intr steps
+ − 840
|> Thm.implies_intr a_D
+ − 841
|> Thm.implies_intr D_dcl
+ − 842
|> Thm.implies_intr D_subset
+ − 843
+ − 844
val simple_induct_rule =
+ − 845
subset_induct_rule
+ − 846
|> Thm.forall_intr (cterm_of thy D)
+ − 847
|> Thm.forall_elim (cterm_of thy acc_R)
+ − 848
|> assume_tac 1 |> Seq.hd
+ − 849
|> (curry op COMP) (acc_downward
+ − 850
|> (instantiate' [SOME (ctyp_of thy domT)]
+ − 851
(map (SOME o cterm_of thy) [R, x, z]))
+ − 852
|> Thm.forall_intr (cterm_of thy z)
+ − 853
|> Thm.forall_intr (cterm_of thy x))
+ − 854
|> Thm.forall_intr (cterm_of thy a)
+ − 855
|> Thm.forall_intr (cterm_of thy P)
+ − 856
in
+ − 857
simple_induct_rule
+ − 858
end
+ − 859
+ − 860
+ − 861
(* FIXME: broken by design *)
+ − 862
fun mk_domain_intro ctxt (Globals {domT, ...}) R R_cases clause =
+ − 863
let
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>
diff
changeset
+ − 864
val thy = Proof_Context.theory_of ctxt
2665
+ − 865
val ClauseInfo {cdata = ClauseContext {gs, lhs, cqs, ...},
+ − 866
qglr = (oqs, _, _, _), ...} = clause
+ − 867
val goal = HOLogic.mk_Trueprop (mk_acc domT R $ lhs)
+ − 868
|> fold_rev (curry Logic.mk_implies) gs
+ − 869
|> cterm_of thy
+ − 870
in
+ − 871
Goal.init goal
+ − 872
|> (SINGLE (resolve_tac [accI] 1)) |> the
+ − 873
|> (SINGLE (eresolve_tac [Thm.forall_elim_vars 0 R_cases] 1)) |> the
2788
+ − 874
|> (SINGLE (auto_tac ctxt)) |> the
2665
+ − 875
|> Goal.conclude
+ − 876
|> fold_rev forall_intr_rename (map fst oqs ~~ cqs)
+ − 877
end
+ − 878
+ − 879
+ − 880
+ − 881
(** Termination rule **)
+ − 882
+ − 883
val wf_induct_rule = @{thm Wellfounded.wfP_induct_rule}
3229
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 884
val wf_in_rel = @{thm Fun_Def.wf_in_rel}
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 885
val in_rel_def = @{thm Fun_Def.in_rel_def}
2665
+ − 886
+ − 887
fun mk_nest_term_case thy globals R' ihyp clause =
+ − 888
let
+ − 889
val Globals {z, ...} = globals
+ − 890
val ClauseInfo {cdata = ClauseContext {qs, cqs, ags, lhs, case_hyp, ...}, tree,
+ − 891
qglr=(oqs, _, _, _), ...} = clause
+ − 892
3218
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 893
val ih_case =
3227
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 894
full_simplify (put_simpset HOL_basic_ss (Proof_Context.init_global thy) addsimps [case_hyp])
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 895
ihyp
2665
+ − 896
+ − 897
fun step (fixes, assumes) (_ $ arg) u (sub,(hyps,thms)) =
+ − 898
let
+ − 899
val used = (u @ sub)
+ − 900
|> map (fn (ctx,thm) => Function_Ctx_Tree.export_thm thy ctx thm)
+ − 901
+ − 902
val hyp = HOLogic.mk_Trueprop (R' $ arg $ lhs)
+ − 903
|> fold_rev (curry Logic.mk_implies o prop_of) used (* additional hyps *)
+ − 904
|> Function_Ctx_Tree.export_term (fixes, assumes)
+ − 905
|> fold_rev (curry Logic.mk_implies o prop_of) ags
+ − 906
|> fold_rev mk_forall_rename (map fst oqs ~~ qs)
+ − 907
|> cterm_of thy
+ − 908
+ − 909
val thm = Thm.assume hyp
+ − 910
|> fold Thm.forall_elim cqs
+ − 911
|> fold Thm.elim_implies ags
+ − 912
|> Function_Ctx_Tree.import_thm thy (fixes, assumes)
+ − 913
|> fold Thm.elim_implies used (* "(arg, lhs) : R'" *)
+ − 914
+ − 915
val z_eq_arg = HOLogic.mk_Trueprop (mk_eq (z, arg))
+ − 916
|> cterm_of thy |> Thm.assume
+ − 917
+ − 918
val acc = thm COMP ih_case
+ − 919
val z_acc_local = acc
+ − 920
|> Conv.fconv_rule
+ − 921
(Conv.arg_conv (Conv.arg_conv (K (Thm.symmetric (z_eq_arg RS eq_reflection)))))
+ − 922
+ − 923
val ethm = z_acc_local
+ − 924
|> Function_Ctx_Tree.export_thm thy (fixes,
+ − 925
z_eq_arg :: case_hyp :: ags @ assumes)
+ − 926
|> fold_rev forall_intr_rename (map fst oqs ~~ cqs)
+ − 927
+ − 928
val sub' = sub @ [(([],[]), acc)]
+ − 929
in
+ − 930
(sub', (hyp :: hyps, ethm :: thms))
+ − 931
end
+ − 932
| step _ _ _ _ = raise Match
+ − 933
in
+ − 934
Function_Ctx_Tree.traverse_tree step tree
+ − 935
end
+ − 936
+ − 937
+ − 938
fun mk_nest_term_rule thy globals R R_cases clauses =
+ − 939
let
+ − 940
val Globals { domT, x, z, ... } = globals
+ − 941
val acc_R = mk_acc domT R
+ − 942
+ − 943
val R' = Free ("R", fastype_of R)
+ − 944
+ − 945
val Rrel = Free ("R", HOLogic.mk_setT (HOLogic.mk_prodT (domT, domT)))
3229
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 946
val inrel_R = Const (@{const_name Fun_Def.in_rel},
2665
+ − 947
HOLogic.mk_setT (HOLogic.mk_prodT (domT, domT)) --> fastype_of R) $ Rrel
+ − 948
+ − 949
val wfR' = HOLogic.mk_Trueprop (Const (@{const_name Wellfounded.wfP},
+ − 950
(domT --> domT --> boolT) --> boolT) $ R')
+ − 951
|> cterm_of thy (* "wf R'" *)
+ − 952
+ − 953
(* Inductive Hypothesis: !!z. (z,x):R' ==> z : acc R *)
3108
+ − 954
val ihyp = Logic.all_const domT $ Abs ("z", domT,
2665
+ − 955
Logic.mk_implies (HOLogic.mk_Trueprop (R' $ Bound 0 $ x),
+ − 956
HOLogic.mk_Trueprop (acc_R $ Bound 0)))
+ − 957
|> cterm_of thy
+ − 958
+ − 959
val ihyp_a = Thm.assume ihyp |> Thm.forall_elim_vars 0
+ − 960
+ − 961
val R_z_x = cterm_of thy (HOLogic.mk_Trueprop (R $ z $ x))
+ − 962
+ − 963
val (hyps, cases) = fold (mk_nest_term_case thy globals R' ihyp_a) clauses ([], [])
+ − 964
in
+ − 965
R_cases
+ − 966
|> Thm.forall_elim (cterm_of thy z)
+ − 967
|> Thm.forall_elim (cterm_of thy x)
+ − 968
|> Thm.forall_elim (cterm_of thy (acc_R $ z))
+ − 969
|> curry op COMP (Thm.assume R_z_x)
+ − 970
|> fold_rev (curry op COMP) cases
+ − 971
|> Thm.implies_intr R_z_x
+ − 972
|> Thm.forall_intr (cterm_of thy z)
+ − 973
|> (fn it => it COMP accI)
+ − 974
|> Thm.implies_intr ihyp
+ − 975
|> Thm.forall_intr (cterm_of thy x)
3220
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 976
|> (fn it => Drule.compose (it, 2, wf_induct_rule))
2665
+ − 977
|> curry op RS (Thm.assume wfR')
+ − 978
|> forall_intr_vars
+ − 979
|> (fn it => it COMP allI)
+ − 980
|> fold Thm.implies_intr hyps
+ − 981
|> Thm.implies_intr wfR'
+ − 982
|> Thm.forall_intr (cterm_of thy R')
+ − 983
|> Thm.forall_elim (cterm_of thy (inrel_R))
+ − 984
|> curry op RS wf_in_rel
3227
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 985
|> full_simplify (put_simpset HOL_basic_ss (Proof_Context.init_global thy) addsimps [in_rel_def])
2665
+ − 986
|> Thm.forall_intr (cterm_of thy Rrel)
+ − 987
end
+ − 988
+ − 989
+ − 990
+ − 991
(* nominal *)
+ − 992
fun prepare_nominal_function config defname [((fname, fT), mixfix)] abstract_qglrs lthy =
+ − 993
let
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>
diff
changeset
+ − 994
val NominalFunctionConfig {domintros, default=default_opt, inv=invariant_opt,...} = config
2665
+ − 995
+ − 996
val default_str = the_default "%x. undefined" default_opt (*FIXME dynamic scoping*)
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>
diff
changeset
+ − 997
val invariant_str = the_default "%x y. True" invariant_opt
2665
+ − 998
val fvar = Free (fname, fT)
+ − 999
val domT = domain_type fT
+ − 1000
val ranT = range_type fT
+ − 1001
+ − 1002
val default = Syntax.parse_term lthy default_str
+ − 1003
|> Type.constraint fT |> Syntax.check_term lthy
+ − 1004
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>
diff
changeset
+ − 1005
val invariant_trm = Syntax.parse_term lthy invariant_str
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>
diff
changeset
+ − 1006
|> Type.constraint ([domT, ranT] ---> @{typ bool}) |> Syntax.check_term lthy
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>
diff
changeset
+ − 1007
2665
+ − 1008
val (globals, ctxt') = fix_globals domT ranT fvar lthy
+ − 1009
+ − 1010
val Globals { x, h, ... } = globals
+ − 1011
+ − 1012
val clauses = map (mk_clause_context x ctxt') abstract_qglrs
+ − 1013
+ − 1014
val n = length abstract_qglrs
+ − 1015
+ − 1016
fun build_tree (ClauseContext { ctxt, rhs, ...}) =
+ − 1017
Function_Ctx_Tree.mk_tree (fname, fT) h ctxt rhs
+ − 1018
+ − 1019
val trees = map build_tree clauses
+ − 1020
val RCss = map find_calls trees
+ − 1021
2707
+ − 1022
val ((G, GIntro_thms, G_elim, G_induct), lthy) =
2665
+ − 1023
PROFILE "def_graph" (define_graph (graph_name defname) fvar domT ranT clauses RCss) lthy
+ − 1024
+ − 1025
val ((f, (_, f_defthm)), lthy) =
+ − 1026
PROFILE "def_fun" (define_function (defname ^ "_sumC_def") (fname, mixfix) domT ranT G default) lthy
+ − 1027
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>
diff
changeset
+ − 1028
val RCss = map (map (inst_RC (Proof_Context.theory_of lthy) fvar f)) RCss
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>
diff
changeset
+ − 1029
val trees = map (Function_Ctx_Tree.inst_tree (Proof_Context.theory_of lthy) fvar f) trees
2665
+ − 1030
+ − 1031
val ((R, RIntro_thmss, R_elim), lthy) =
+ − 1032
PROFILE "def_rel" (define_recursion_relation (rel_name defname) domT abstract_qglrs clauses RCss) lthy
+ − 1033
+ − 1034
val (_, lthy) =
+ − 1035
Local_Theory.abbrev Syntax.mode_default ((Binding.name (dom_name defname), NoSyn), mk_acc domT R) lthy
+ − 1036
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>
diff
changeset
+ − 1037
val newthy = Proof_Context.theory_of lthy
2665
+ − 1038
val clauses = map (transfer_clause_ctx newthy) clauses
+ − 1039
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>
diff
changeset
+ − 1040
val cert = cterm_of (Proof_Context.theory_of lthy)
2665
+ − 1041
+ − 1042
val xclauses = PROFILE "xclauses"
+ − 1043
(map7 (mk_clause_info globals G f) (1 upto n) clauses abstract_qglrs trees
+ − 1044
RCss GIntro_thms) RIntro_thmss
+ − 1045
+ − 1046
val complete =
+ − 1047
mk_completeness globals clauses abstract_qglrs |> cert |> Thm.assume
+ − 1048
+ − 1049
val compat =
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>
diff
changeset
+ − 1050
mk_compat_proof_obligations domT ranT fvar f RCss invariant_trm abstract_qglrs
2665
+ − 1051
|> map (cert #> Thm.assume)
+ − 1052
2707
+ − 1053
val G_eqvt = mk_eqvt G |> cert |> Thm.assume
+ − 1054
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>
diff
changeset
+ − 1055
val invariant = mk_invariant globals G invariant_trm |> cert |> Thm.assume
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>
diff
changeset
+ − 1056
2665
+ − 1057
val compat_store = store_compat_thms n compat
+ − 1058
+ − 1059
val (goalstate, values) = PROFILE "prove_stuff"
+ − 1060
(prove_stuff lthy globals G f R xclauses complete compat
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>
diff
changeset
+ − 1061
compat_store G_elim G_eqvt invariant) f_defthm
2665
+ − 1062
+ − 1063
fun mk_partial_rules provedgoal =
+ − 1064
let
+ − 1065
val newthy = theory_of_thm provedgoal (*FIXME*)
+ − 1066
2974
+ − 1067
val ((graph_is_function, complete_thm), graph_is_eqvt) =
2665
+ − 1068
provedgoal
+ − 1069
|> Conjunction.elim
2974
+ − 1070
|>> fst o Conjunction.elim
+ − 1071
|>> Conjunction.elim
+ − 1072
|>> apfst (Thm.forall_elim_vars 0)
2665
+ − 1073
+ − 1074
val f_iff = graph_is_function RS (f_defthm RS ex1_implies_iff)
2974
+ − 1075
val f_eqvt = graph_is_function RS (graph_is_eqvt RS (f_defthm RS @{thm fundef_ex1_eqvt}))
2973
d1038e67923a
added a flag (eqvt) to termination proofs arising fron nominal_primrecs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1076
2665
+ − 1077
val psimps = PROFILE "Proving simplification rules"
+ − 1078
(mk_psimps newthy globals R xclauses values f_iff) graph_is_function
+ − 1079
+ − 1080
val simple_pinduct = PROFILE "Proving partial induction rule"
+ − 1081
(mk_partial_induct_rule newthy globals R complete_thm) xclauses
+ − 1082
+ − 1083
val total_intro = PROFILE "Proving nested termination rule"
+ − 1084
(mk_nest_term_rule newthy globals R R_elim) xclauses
+ − 1085
+ − 1086
val dom_intros =
+ − 1087
if domintros then SOME (PROFILE "Proving domain introduction rules"
+ − 1088
(map (mk_domain_intro lthy globals R R_elim)) xclauses)
+ − 1089
else NONE
+ − 1090
+ − 1091
in
2973
d1038e67923a
added a flag (eqvt) to termination proofs arising fron nominal_primrecs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1092
NominalFunctionResult {fs=[f], G=G, R=R, cases=complete_thm,
2665
+ − 1093
psimps=psimps, simple_pinducts=[simple_pinduct],
2973
d1038e67923a
added a flag (eqvt) to termination proofs arising fron nominal_primrecs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1094
termination=total_intro, domintros=dom_intros, eqvts=[f_eqvt]}
2665
+ − 1095
end
+ − 1096
in
3197
25d11b449e92
definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1097
((f, G, GIntro_thms, G_induct, goalstate, mk_partial_rules), lthy)
2665
+ − 1098
end
+ − 1099
+ − 1100
+ − 1101
end